annotate src/zorn.agda @ 1096:55ab5de1ae02

recovery
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Fri, 23 Dec 2022 12:54:05 +0900
parents 63c1167b2343
children 40345abc0949
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
rev   line source
478
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 477
diff changeset
1 {-# OPTIONS --allow-unsolved-metas #-}
508
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 507
diff changeset
2 open import Level hiding ( suc ; zero )
431
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
3 open import Ordinals
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
4 open import Relation.Binary
552
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 551
diff changeset
5 open import Relation.Binary.Core
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 551
diff changeset
6 open import Relation.Binary.PropositionalEquality
1096
55ab5de1ae02 recovery
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1091
diff changeset
7 import OD
552
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 551
diff changeset
8 module zorn {n : Level } (O : Ordinals {n}) (_<_ : (x y : OD.HOD O ) → Set n ) (PO : IsStrictPartialOrder _≡_ _<_ ) where
431
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
9
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
10 --
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
11 -- Zorn-lemma : { A : HOD }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
12 -- → o∅ o< & A
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
13 -- → ( ( B : HOD) → (B⊆A : B ⊆ A) → IsTotalOrderSet B → SUP A B ) -- SUP condition
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
14 -- → Maximal A
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
15 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
16
1091
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
17 open import zf -- hiding ( _⊆_ )
477
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
18 open import logic
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
19 -- open import partfunc {n} O
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
20
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
21 open import Relation.Nullary
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
22 open import Data.Empty
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
23 import BAlgbra
431
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
24
555
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 554
diff changeset
25 open import Data.Nat hiding ( _<_ ; _≤_ )
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
26 open import Data.Nat.Properties
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
27 open import nat
555
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 554
diff changeset
28
431
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
29
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
30 open inOrdinal O
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
31 open OD O
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
32 open OD.OD
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
33 open ODAxiom odAxiom
477
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
34 import OrdUtil
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
35 import ODUtil
431
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
36 open Ordinals.Ordinals O
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
37 open Ordinals.IsOrdinals isOrdinal
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
38 open Ordinals.IsNext isNext
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
39 open OrdUtil O
477
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
40 open ODUtil O
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
41
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
42
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
43 import ODC
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
44
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
45 open _∧_
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
46 open _∨_
24b4b854b310 separate zorn lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 476
diff changeset
47 open Bool
431
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
48
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
49 open HOD
a5f8084b8368 reorganiztion for apkg
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
50
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
51 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
52 -- Partial Order on HOD ( possibly limited in A )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
53 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
54
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
55 _<<_ : (x y : Ordinal ) → Set n
570
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
56 x << y = * x < * y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
57
1031
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
58 _≤_ : (x y : Ordinal ) → Set n -- we can't use * x ≡ * y, it is Set (Level.suc n). Level (suc n) troubles Chain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
59 x ≤ y = (x ≡ y ) ∨ ( * x < * y )
765
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 764
diff changeset
60
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
61 POO : IsStrictPartialOrder _≡_ _<<_
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
62 POO = record { isEquivalence = record { refl = refl ; sym = sym ; trans = trans }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
63 ; trans = IsStrictPartialOrder.trans PO
570
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 569
diff changeset
64 ; irrefl = λ x=y x<y → IsStrictPartialOrder.irrefl PO (cong (*) x=y) x<y
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
65 ; <-resp-≈ = record { fst = λ {x} {y} {y1} y=y1 xy1 → subst (λ k → x << k ) y=y1 xy1 ; snd = λ {x} {x1} {y} x=x1 x1y → subst (λ k → k << x ) x=x1 x1y } }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
66
1031
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
67 ≤-ftrans : {x y z : Ordinal } → x ≤ y → y ≤ z → x ≤ z
554
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 553
diff changeset
68 ≤-ftrans {x} {y} {z} (case1 refl ) (case1 refl ) = case1 refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 553
diff changeset
69 ≤-ftrans {x} {y} {z} (case1 refl ) (case2 y<z) = case2 y<z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 553
diff changeset
70 ≤-ftrans {x} {_} {z} (case2 x<y ) (case1 refl ) = case2 x<y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 553
diff changeset
71 ≤-ftrans {x} {y} {z} (case2 x<y) (case2 y<z) = case2 ( IsStrictPartialOrder.trans PO x<y y<z )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 553
diff changeset
72
1031
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
73 ftrans≤-< : {x y z : Ordinal } → x ≤ y → y << z → x << z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
74 ftrans≤-< {x} {y} {z} (case1 eq) y<z = subst (λ k → k < * z) (sym (cong (*) eq)) y<z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
75 ftrans≤-< {x} {y} {z} (case2 lt) y<z = IsStrictPartialOrder.trans PO lt y<z
951
86a2bfb7222e supf mc = mc
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 950
diff changeset
76
1031
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
77 ftrans<-≤ : {x y z : Ordinal } → x << y → y ≤ z → x << z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
78 ftrans<-≤ {x} {y} {z} x<y (case1 eq) = subst (λ k → * x < k ) ((cong (*) eq)) x<y
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
79 ftrans<-≤ {x} {y} {z} x<y (case2 lt) = IsStrictPartialOrder.trans PO x<y lt
779
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 778
diff changeset
80
556
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
81 <-irr : {a b : HOD} → (a ≡ b ) ∨ (a < b ) → b < a → ⊥
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
82 <-irr {a} {b} (case1 a=b) b<a = IsStrictPartialOrder.irrefl PO (sym a=b) b<a
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
83 <-irr {a} {b} (case2 a<b) b<a = IsStrictPartialOrder.irrefl PO refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
84 (IsStrictPartialOrder.trans PO b<a a<b)
490
00c71d1dc316 IsPartialOrder
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 489
diff changeset
85
1032
382680c3e9be minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1031
diff changeset
86 <<-irr : {a b : Ordinal } → a ≤ b → b << a → ⊥
382680c3e9be minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1031
diff changeset
87 <<-irr {a} {b} (case1 a=b) b<a = IsStrictPartialOrder.irrefl PO (cong (*) (sym a=b)) b<a
382680c3e9be minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1031
diff changeset
88 <<-irr {a} {b} (case2 a<b) b<a = IsStrictPartialOrder.irrefl PO refl (IsStrictPartialOrder.trans PO b<a a<b)
382680c3e9be minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1031
diff changeset
89
561
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
90 ptrans = IsStrictPartialOrder.trans PO
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 560
diff changeset
91
492
e28b1da1b58d Partial Order
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 491
diff changeset
92 open _==_
1091
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
93 -- open _⊆_ -- we use different definition
492
e28b1da1b58d Partial Order
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 491
diff changeset
94
1091
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
95 -- We cannot prove this without Well foundedness of A
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
96 --
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
97 -- <-TransFinite : {A x : HOD} → {P : HOD → Set n} → x ∈ A
1091
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
98 -- → ({y : HOD} → A ∋ y → y < x → P y ) → P x
879
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
99 -- <-TransFinite = ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
100
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
101 --
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
102 -- Closure of ≤-monotonic function f has total order
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
103 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
104
1031
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
105 ≤-monotonic-f : (A : HOD) → ( Ordinal → Ordinal ) → Set n
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
106 ≤-monotonic-f A f = (x : Ordinal ) → odef A x → ( x ≤ (f x) ) ∧ odef A (f x )
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
107
992
4aaecae58da5 ... x < & A ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 991
diff changeset
108 <-monotonic-f : (A : HOD) → ( Ordinal → Ordinal ) → Set n
4aaecae58da5 ... x < & A ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 991
diff changeset
109 <-monotonic-f A f = (x : Ordinal ) → odef A x → ( * x < * (f x) ) ∧ odef A (f x )
4aaecae58da5 ... x < & A ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 991
diff changeset
110
551
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
111 data FClosure (A : HOD) (f : Ordinal → Ordinal ) (s : Ordinal) : Ordinal → Set n where
783
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 782
diff changeset
112 init : {s1 : Ordinal } → odef A s → s ≡ s1 → FClosure A f s s1
555
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 554
diff changeset
113 fsuc : (x : Ordinal) ( p : FClosure A f s x ) → FClosure A f s (f x)
554
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 553
diff changeset
114
556
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
115 A∋fc : {A : HOD} (s : Ordinal) {y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f) → (fcy : FClosure A f s y ) → odef A y
783
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 782
diff changeset
116 A∋fc {A} s f mf (init as refl ) = as
556
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
117 A∋fc {A} s f mf (fsuc y fcy) = proj2 (mf y ( A∋fc {A} s f mf fcy ) )
555
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 554
diff changeset
118
714
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 713
diff changeset
119 A∋fcs : {A : HOD} (s : Ordinal) {y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f) → (fcy : FClosure A f s y ) → odef A s
783
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 782
diff changeset
120 A∋fcs {A} s f mf (init as refl) = as
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
121 A∋fcs {A} s f mf (fsuc y fcy) = A∋fcs {A} s f mf fcy
714
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 713
diff changeset
122
1031
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
123 s≤fc : {A : HOD} (s : Ordinal ) {y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f) → (fcy : FClosure A f s y ) → s ≤ y
783
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 782
diff changeset
124 s≤fc {A} s {.s} f mf (init x refl ) = case1 refl
556
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
125 s≤fc {A} s {.(f x)} f mf (fsuc x fcy) with proj1 (mf x (A∋fc s f mf fcy ) )
1031
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
126 ... | case1 x=fx = subst₂ (λ j k → j ≤ k ) refl x=fx (s≤fc s f mf fcy)
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
127 ... | case2 x<fx with s≤fc {A} s f mf fcy
1031
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
128 ... | case1 s≡x = case2 ( subst₂ (λ j k → j < k ) (sym (cong (*) s≡x )) refl x<fx )
556
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 555
diff changeset
129 ... | case2 s<x = case2 ( IsStrictPartialOrder.trans PO s<x x<fx )
555
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 554
diff changeset
130
800
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
131 fcn : {A : HOD} (s : Ordinal) { x : Ordinal} {f : Ordinal → Ordinal} → (mf : ≤-monotonic-f A f) → FClosure A f s x → ℕ
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
132 fcn s mf (init as refl) = zero
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
133 fcn {A} s {x} {f} mf (fsuc y p) with proj1 (mf y (A∋fc s f mf p))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
134 ... | case1 eq = fcn s mf p
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
135 ... | case2 y<fy = suc (fcn s mf p )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
136
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
137 fcn-inject : {A : HOD} (s : Ordinal) { x y : Ordinal} {f : Ordinal → Ordinal} → (mf : ≤-monotonic-f A f)
800
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
138 → (cx : FClosure A f s x ) (cy : FClosure A f s y ) → fcn s mf cx ≡ fcn s mf cy → * x ≡ * y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
139 fcn-inject {A} s {x} {y} {f} mf cx cy eq = fc00 (fcn s mf cx) (fcn s mf cy) eq cx cy refl refl where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
140 fc06 : {y : Ordinal } { as : odef A s } (eq : s ≡ y ) { j : ℕ } → ¬ suc j ≡ fcn {A} s {y} {f} mf (init as eq )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
141 fc06 {x} {y} refl {j} not = fc08 not where
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
142 fc08 : {j : ℕ} → ¬ suc j ≡ 0
800
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
143 fc08 ()
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
144 fc07 : {x : Ordinal } (cx : FClosure A f s x ) → 0 ≡ fcn s mf cx → * s ≡ * x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
145 fc07 {x} (init as refl) eq = refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
146 fc07 {.(f x)} (fsuc x cx) eq with proj1 (mf x (A∋fc s f mf cx ) )
1031
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
147 ... | case1 x=fx = subst (λ k → * s ≡ k ) (cong (*) x=fx) ( fc07 cx eq )
800
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
148 -- ... | case2 x<fx = ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
149 fc00 : (i j : ℕ ) → i ≡ j → {x y : Ordinal } → (cx : FClosure A f s x ) (cy : FClosure A f s y ) → i ≡ fcn s mf cx → j ≡ fcn s mf cy → * x ≡ * y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
150 fc00 (suc i) (suc j) x cx (init x₃ x₄) x₁ x₂ = ⊥-elim ( fc06 x₄ x₂ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
151 fc00 (suc i) (suc j) x (init x₃ x₄) (fsuc x₅ cy) x₁ x₂ = ⊥-elim ( fc06 x₄ x₁ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
152 fc00 zero zero refl (init _ refl) (init x₁ refl) i=x i=y = refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
153 fc00 zero zero refl (init as refl) (fsuc y cy) i=x i=y with proj1 (mf y (A∋fc s f mf cy ) )
1031
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
154 ... | case1 y=fy = subst (λ k → * s ≡ * k ) y=fy (fc07 cy i=y) -- ( fc00 zero zero refl (init as refl) cy i=x i=y )
800
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
155 fc00 zero zero refl (fsuc x cx) (init as refl) i=x i=y with proj1 (mf x (A∋fc s f mf cx ) )
1031
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
156 ... | case1 x=fx = subst (λ k → * k ≡ * s ) x=fx (sym (fc07 cx i=x)) -- ( fc00 zero zero refl cx (init as refl) i=x i=y )
800
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
157 fc00 zero zero refl (fsuc x cx) (fsuc y cy) i=x i=y with proj1 (mf x (A∋fc s f mf cx ) ) | proj1 (mf y (A∋fc s f mf cy ) )
1031
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
158 ... | case1 x=fx | case1 y=fy = subst₂ (λ j k → * j ≡ * k ) x=fx y=fy ( fc00 zero zero refl cx cy i=x i=y )
800
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
159 fc00 (suc i) (suc j) i=j {.(f x)} {.(f y)} (fsuc x cx) (fsuc y cy) i=x j=y with proj1 (mf x (A∋fc s f mf cx ) ) | proj1 (mf y (A∋fc s f mf cy ) )
1031
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
160 ... | case1 x=fx | case1 y=fy = subst₂ (λ j k → * j ≡ * k ) x=fx y=fy ( fc00 (suc i) (suc j) i=j cx cy i=x j=y )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
161 ... | case1 x=fx | case2 y<fy = subst (λ k → * k ≡ * (f y)) x=fx (fc02 x cx i=x) where
800
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
162 fc02 : (x1 : Ordinal) → (cx1 : FClosure A f s x1 ) → suc i ≡ fcn s mf cx1 → * x1 ≡ * (f y)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
163 fc02 x1 (init x₁ x₂) x = ⊥-elim (fc06 x₂ x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
164 fc02 .(f x1) (fsuc x1 cx1) i=x1 with proj1 (mf x1 (A∋fc s f mf cx1 ) )
1031
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
165 ... | case1 eq = trans (sym (cong (*) eq )) ( fc02 x1 cx1 i=x1 ) -- derefence while f x ≡ x
800
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
166 ... | case2 lt = subst₂ (λ j k → * (f j) ≡ * (f k )) &iso &iso ( cong (λ k → * ( f (& k ))) fc04) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
167 fc04 : * x1 ≡ * y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
168 fc04 = fc00 i j (cong pred i=j) cx1 cy (cong pred i=x1) (cong pred j=y)
1031
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
169 ... | case2 x<fx | case1 y=fy = subst (λ k → * (f x) ≡ * k ) y=fy (fc03 y cy j=y) where
800
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
170 fc03 : (y1 : Ordinal) → (cy1 : FClosure A f s y1 ) → suc j ≡ fcn s mf cy1 → * (f x) ≡ * y1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
171 fc03 y1 (init x₁ x₂) x = ⊥-elim (fc06 x₂ x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
172 fc03 .(f y1) (fsuc y1 cy1) j=y1 with proj1 (mf y1 (A∋fc s f mf cy1 ) )
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
173 ... | case1 eq = trans ( fc03 y1 cy1 j=y1 ) (cong (*) eq)
800
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
174 ... | case2 lt = subst₂ (λ j k → * (f j) ≡ * (f k )) &iso &iso ( cong (λ k → * ( f (& k ))) fc05) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
175 fc05 : * x ≡ * y1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
176 fc05 = fc00 i j (cong pred i=j) cx cy1 (cong pred i=x) (cong pred j=y1)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
177 ... | case2 x₁ | case2 x₂ = subst₂ (λ j k → * (f j) ≡ * (f k) ) &iso &iso (cong (λ k → * (f (& k))) (fc00 i j (cong pred i=j) cx cy (cong pred i=x) (cong pred j=y)))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
178
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
179
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
180 fcn-< : {A : HOD} (s : Ordinal ) { x y : Ordinal} {f : Ordinal → Ordinal} → (mf : ≤-monotonic-f A f)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
181 → (cx : FClosure A f s x ) (cy : FClosure A f s y ) → fcn s mf cx Data.Nat.< fcn s mf cy → * x < * y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
182 fcn-< {A} s {x} {y} {f} mf cx cy x<y = fc01 (fcn s mf cy) cx cy refl x<y where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
183 fc06 : {y : Ordinal } { as : odef A s } (eq : s ≡ y ) { j : ℕ } → ¬ suc j ≡ fcn {A} s {y} {f} mf (init as eq )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
184 fc06 {x} {y} refl {j} not = fc08 not where
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
185 fc08 : {j : ℕ} → ¬ suc j ≡ 0
800
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
186 fc08 ()
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
187 fc01 : (i : ℕ ) → {y : Ordinal } → (cx : FClosure A f s x ) (cy : FClosure A f s y ) → (i ≡ fcn s mf cy ) → fcn s mf cx Data.Nat.< i → * x < * y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
188 fc01 (suc i) cx (init x₁ x₂) x (s≤s x₃) = ⊥-elim (fc06 x₂ x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
189 fc01 (suc i) {y} cx (fsuc y1 cy) i=y (s≤s x<i) with proj1 (mf y1 (A∋fc s f mf cy ) )
1031
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
190 ... | case1 y=fy = subst (λ k → * x < k ) (cong (*) y=fy) ( fc01 (suc i) {y1} cx cy i=y (s≤s x<i) )
800
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
191 ... | case2 y<fy with <-cmp (fcn s mf cx ) i
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
192 ... | tri> ¬a ¬b c = ⊥-elim ( nat-≤> x<i c )
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
193 ... | tri≈ ¬a b ¬c = subst (λ k → k < * (f y1) ) (fcn-inject s mf cy cx (sym (trans b (cong pred i=y) ))) y<fy
800
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
194 ... | tri< a ¬b ¬c = IsStrictPartialOrder.trans PO fc02 y<fy where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
195 fc03 : suc i ≡ suc (fcn s mf cy) → i ≡ fcn s mf cy
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
196 fc03 eq = cong pred eq
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
197 fc02 : * x < * y1
800
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
198 fc02 = fc01 i cx cy (fc03 i=y ) a
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
199
557
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 556
diff changeset
200
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
201 fcn-cmp : {A : HOD} (s : Ordinal) { x y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f)
554
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 553
diff changeset
202 → (cx : FClosure A f s x) → (cy : FClosure A f s y ) → Tri (* x < * y) (* x ≡ * y) (* y < * x )
800
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
203 fcn-cmp {A} s {x} {y} f mf cx cy with <-cmp ( fcn s mf cx ) (fcn s mf cy )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
204 ... | tri< a ¬b ¬c = tri< fc11 (λ eq → <-irr (case1 (sym eq)) fc11) (λ lt → <-irr (case2 fc11) lt) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
205 fc11 : * x < * y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
206 fc11 = fcn-< {A} s {x} {y} {f} mf cx cy a
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
207 ... | tri≈ ¬a b ¬c = tri≈ (λ lt → <-irr (case1 (sym fc10)) lt) fc10 (λ lt → <-irr (case1 fc10) lt) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
208 fc10 : * x ≡ * y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
209 fc10 = fcn-inject {A} s {x} {y} {f} mf cx cy b
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
210 ... | tri> ¬a ¬b c = tri> (λ lt → <-irr (case2 fc12) lt) (λ eq → <-irr (case1 eq) fc12) fc12 where
800
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
211 fc12 : * y < * x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 799
diff changeset
212 fc12 = fcn-< {A} s {y} {x} {f} mf cy cx c
600
71a1ed72cd21 not yet ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 599
diff changeset
213
563
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 562
diff changeset
214
729
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 728
diff changeset
215
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
216 -- open import Relation.Binary.Properties.Poset as Poset
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
217
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
218 IsTotalOrderSet : ( A : HOD ) → Set (Level.suc n)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
219 IsTotalOrderSet A = {a b : HOD} → odef A (& a) → odef A (& b) → Tri (a < b) (a ≡ b) (b < a )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
220
1091
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
221 ⊆-IsTotalOrderSet : { A B : HOD } → B ⊆ A → IsTotalOrderSet A → IsTotalOrderSet B
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
222 ⊆-IsTotalOrderSet {A} {B} B⊆A T ax ay = T (B⊆A ax) (B⊆A ay)
554
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 553
diff changeset
223
958
33891adf80ea IsMinSup contains not HasPrev
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 957
diff changeset
224 record HasPrev (A B : HOD) ( f : Ordinal → Ordinal ) (x : Ordinal ) : Set n where
533
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 532
diff changeset
225 field
836
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 835
diff changeset
226 ax : odef A x
534
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 533
diff changeset
227 y : Ordinal
541
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 540
diff changeset
228 ay : odef B y
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
229 x=fy : x ≡ f y
529
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 528
diff changeset
230
1031
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
231 record IsSUP (A B : HOD) (x : Ordinal ) : Set n where
962
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 961
diff changeset
232 field
1031
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
233 ax : odef A x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
234 x≤sup : {y : Ordinal} → odef B y → (y ≡ x ) ∨ (y << x ) -- B is Total, use positive
568
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
235
656
db9477c80dce data Chain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 655
diff changeset
236 record SUP ( A B : HOD ) : Set (Level.suc n) where
db9477c80dce data Chain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 655
diff changeset
237 field
db9477c80dce data Chain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 655
diff changeset
238 sup : HOD
1031
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
239 isSUP : IsSUP A B (& sup)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
240 ax = IsSUP.ax isSUP
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
241 x≤sup = IsSUP.x≤sup isSUP
656
db9477c80dce data Chain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 655
diff changeset
242
690
33f90b483211 Chain with chainf
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 689
diff changeset
243 --
1091
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
244 -- Our Proof strategy of the Zorn Lemma
878
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 877
diff changeset
245 --
990
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 989
diff changeset
246 -- f (f ( ... (supf y))) f (f ( ... (supf z1)))
878
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 877
diff changeset
247 -- / | / |
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 877
diff changeset
248 -- / | / |
990
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 989
diff changeset
249 -- supf y < supf z1 < supf z2
878
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 877
diff changeset
250 -- o< o<
990
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 989
diff changeset
251 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 989
diff changeset
252 -- if sup z1 ≡ sup z2, the chain is stopped at sup z1, then f (sup z1) ≡ sup z1
1039
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1038
diff changeset
253 -- this means sup z1 is the Maximal, so f is <-monotonic if we have no Maximal.
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1038
diff changeset
254 --
990
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 989
diff changeset
255
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
256 fc-stop : ( A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f) { a b : Ordinal }
990
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 989
diff changeset
257 → (aa : odef A a ) →( {y : Ordinal} → FClosure A f a y → (y ≡ b ) ∨ (y << b )) → a ≡ b → f a ≡ a
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 989
diff changeset
258 fc-stop A f mf {a} {b} aa x≤sup a=b with x≤sup (fsuc a (init aa refl ))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 989
diff changeset
259 ... | case1 eq = trans eq (sym a=b)
1031
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
260 ... | case2 lt = ⊥-elim (<-irr (case1 (cong (λ k → * (f k) ) (sym a=b))) (ftrans<-≤ lt fc00 ) ) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
261 fc00 : b ≤ (f b)
990
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 989
diff changeset
262 fc00 = proj1 (mf _ (subst (λ k → odef A k) a=b aa ))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 989
diff changeset
263
694
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 693
diff changeset
264 ∈∧P→o< : {A : HOD } {y : Ordinal} → {P : Set n} → odef A y ∧ P → y o< & A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 693
diff changeset
265 ∈∧P→o< {A } {y} p = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) (proj1 p )))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 693
diff changeset
266
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
267 -- Union of supf z and FClosure A f y
662
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 661
diff changeset
268
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
269 data UChain { A : HOD } { f : Ordinal → Ordinal } {supf : Ordinal → Ordinal} {y : Ordinal } (ay : odef A y )
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
270 (x : Ordinal) : (z : Ordinal) → Set n where
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
271 ch-init : {z : Ordinal } (fc : FClosure A f y z) → UChain ay x z
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
272 ch-is-sup : (u : Ordinal) {z : Ordinal } (u<x : u o< x) (supu=u : supf u ≡ u) ( fc : FClosure A f (supf u) z ) → UChain ay x z
1030
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1029
diff changeset
273
1060
a09f5e728f92 IChain?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1059
diff changeset
274 UnionCF : ( A : HOD ) ( f : Ordinal → Ordinal ) {y : Ordinal } (ay : odef A y ) ( supf : Ordinal → Ordinal ) ( x : Ordinal ) → HOD
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
275 UnionCF A f ay supf x
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
276 = record { od = record { def = λ z → odef A z ∧ UChain {A} {f} {supf} ay x z } ;
1028
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1027
diff changeset
277 odmax = & A ; <odmax = λ {y} sy → ∈∧P→o< sy }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1027
diff changeset
278
1061
fc37082d2a8a another IChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1060
diff changeset
279 -- Union of chain lower than x
1060
a09f5e728f92 IChain?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1059
diff changeset
280
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
281 data IChain {A : HOD} { f : Ordinal → Ordinal } {y : Ordinal } (ay : odef A y )
1063
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1062
diff changeset
282 {x : Ordinal } (supfz : {z : Ordinal } → z o< x → Ordinal) : (z : Ordinal ) → Set n where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1062
diff changeset
283 ic-init : {z : Ordinal } (fc : FClosure A f y z) → IChain ay supfz z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1062
diff changeset
284 ic-isup : {z : Ordinal} (i : Ordinal) (i<x : i o< x) (s<x : supfz i<x o≤ i ) (fc : FClosure A f (supfz i<x) z) → IChain ay supfz z
1060
a09f5e728f92 IChain?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1059
diff changeset
285
1063
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1062
diff changeset
286 UnionIC : ( A : HOD ) ( f : Ordinal → Ordinal ) { x : Ordinal } {y : Ordinal } (ay : odef A y ) (supfz : {z : Ordinal } → z o< x → Ordinal) → HOD
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
287 UnionIC A f ay supfz
1063
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1062
diff changeset
288 = record { od = record { def = λ z → odef A z ∧ IChain {A} {f} ay supfz z } ;
1060
a09f5e728f92 IChain?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1059
diff changeset
289 odmax = & A ; <odmax = λ {y} sy → ∈∧P→o< sy }
1028
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1027
diff changeset
290
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
291 supf-inject0 : {x y : Ordinal } {supf : Ordinal → Ordinal } → (supf-mono : {x y : Ordinal } → x o≤ y → supf x o≤ supf y )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
292 → supf x o< supf y → x o< y
842
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 841
diff changeset
293 supf-inject0 {x} {y} {supf} supf-mono sx<sy with trio< x y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 841
diff changeset
294 ... | tri< a ¬b ¬c = a
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 841
diff changeset
295 ... | tri≈ ¬a refl ¬c = ⊥-elim ( o<¬≡ (cong supf refl) sx<sy )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 841
diff changeset
296 ... | tri> ¬a ¬b y<x with osuc-≡< (supf-mono (o<→≤ y<x) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 841
diff changeset
297 ... | case1 eq = ⊥-elim ( o<¬≡ (sym eq) sx<sy )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 841
diff changeset
298 ... | case2 lt = ⊥-elim ( o<> sx<sy lt )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 841
diff changeset
299
1031
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
300 record IsMinSUP ( A B : HOD ) (sup : Ordinal) : Set n where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
301 field
1032
382680c3e9be minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1031
diff changeset
302 as : odef A sup
1031
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
303 x≤sup : {x : Ordinal } → odef B x → (x ≡ sup ) ∨ (x << sup )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
304 minsup : { sup1 : Ordinal } → odef A sup1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
305 → ( {x : Ordinal } → odef B x → (x ≡ sup1 ) ∨ (x << sup1 )) → sup o≤ sup1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
306
879
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
307 record MinSUP ( A B : HOD ) : Set n where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
308 field
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
309 sup : Ordinal
1031
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
310 isMinSUP : IsMinSUP A B sup
1032
382680c3e9be minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1031
diff changeset
311 as = IsMinSUP.as isMinSUP
1031
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
312 x≤sup = IsMinSUP.x≤sup isMinSUP
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
313 minsup = IsMinSUP.minsup isMinSUP
879
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
314
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
315 z09 : {b : Ordinal } { A : HOD } → odef A b → b o< & A
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
316 z09 {b} {A} ab = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) ab))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
317
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
318 chain-mono : {A : HOD} ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y) (supf : Ordinal → Ordinal )
919
213f12f27003 supf u o< supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 918
diff changeset
319 (supf-mono : {x y : Ordinal } → x o≤ y → supf x o≤ supf y ) {a b c : Ordinal} → a o≤ b
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
320 → odef (UnionCF A f ay supf a) c → odef (UnionCF A f ay supf b) c
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
321 chain-mono f mf ay supf supf-mono {a} {b} {c} a≤b ⟪ ua , ch-init fc ⟫ = ⟪ ua , ch-init fc ⟫
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
322 chain-mono f mf ay supf supf-mono {a} {b} {c} a≤b ⟪ ua , ch-is-sup u u<x supu=u fc ⟫ = ⟪ ua , ch-is-sup u (ordtrans<-≤ u<x a≤b) supu=u fc ⟫
908
d917831fb607 supf (supf x) ≡ supf x is bad
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 907
diff changeset
323
1038
dfbac4b59bae mf< everywhere
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1037
diff changeset
324 record ZChain ( A : HOD ) ( f : Ordinal → Ordinal ) (mf< : <-monotonic-f A f)
783
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 782
diff changeset
325 {y : Ordinal} (ay : odef A y) ( z : Ordinal ) : Set (Level.suc n) where
655
b602e3f070df UChain rewrite
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 654
diff changeset
326 field
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
327 supf : Ordinal → Ordinal
1028
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1027
diff changeset
328
1090
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
329 supf-mono : {a b : Ordinal } → a o≤ b → supf a o≤ supf b
1046
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1045
diff changeset
330 cfcs : {a b w : Ordinal } → a o< b → b o≤ z → supf a o< b → FClosure A f (supf a) w → odef (UnionCF A f ay supf b) w
868
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 867
diff changeset
331 asupf : {x : Ordinal } → odef A (supf x)
1059
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1058
diff changeset
332 zo≤sz : {x : Ordinal } → x o≤ z → x o≤ supf x
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
333 is-minsup : {x : Ordinal } → (x≤z : x o≤ z) → IsMinSUP A (UnionCF A f ay supf x) (supf x)
994
a15f1cddf4c6 u ≤ x again?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 993
diff changeset
334
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
335 chain : HOD
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
336 chain = UnionCF A f ay supf z
1091
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
337 chain⊆A : chain ⊆ A
861
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 860
diff changeset
338 chain⊆A = λ lt → proj1 lt
934
ebcad8e5ae55 resync zorn.agda
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 930
diff changeset
339
1058
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1057
diff changeset
340 chain∋init : {x : Ordinal } → odef (UnionCF A f ay supf x) y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1057
diff changeset
341 chain∋init {x} = ⟪ ay , ch-init (init ay refl) ⟫
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
342
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
343 mf : ≤-monotonic-f A f
1040
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1039
diff changeset
344 mf x ax = ⟪ case2 mf00 , proj2 (mf< x ax ) ⟫ where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1039
diff changeset
345 mf00 : * x < * (f x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1039
diff changeset
346 mf00 = proj1 ( mf< x ax )
1038
dfbac4b59bae mf< everywhere
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1037
diff changeset
347
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
348 f-next : {a z : Ordinal} → odef (UnionCF A f ay supf z) a → odef (UnionCF A f ay supf z) (f a)
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
349 f-next {a} ⟪ ua , ch-init fc ⟫ = ⟪ proj2 ( mf _ ua) , ch-init (fsuc _ fc) ⟫
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
350 f-next {a} ⟪ ua , ch-is-sup u su<x su=u fc ⟫ = ⟪ proj2 ( mf _ ua) , ch-is-sup u su<x su=u (fsuc _ fc) ⟫
861
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 860
diff changeset
351
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
352 supf-inject : {x y : Ordinal } → supf x o< supf y → x o< y
825
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 824
diff changeset
353 supf-inject {x} {y} sx<sy with trio< x y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 824
diff changeset
354 ... | tri< a ¬b ¬c = a
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 824
diff changeset
355 ... | tri≈ ¬a refl ¬c = ⊥-elim ( o<¬≡ (cong supf refl) sx<sy )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 824
diff changeset
356 ... | tri> ¬a ¬b y<x with osuc-≡< (supf-mono (o<→≤ y<x) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 824
diff changeset
357 ... | case1 eq = ⊥-elim ( o<¬≡ (sym eq) sx<sy )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 824
diff changeset
358 ... | case2 lt = ⊥-elim ( o<> sx<sy lt )
798
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 797
diff changeset
359
1038
dfbac4b59bae mf< everywhere
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1037
diff changeset
360 csupf : {b : Ordinal } → supf b o< supf z → supf b o< z → odef (UnionCF A f ay supf z) (supf b) -- supf z is not an element of this chain
dfbac4b59bae mf< everywhere
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1037
diff changeset
361 csupf {b} sb<sz sb<z = cfcs (supf-inject sb<sz) o≤-refl sb<z (init asupf refl)
994
a15f1cddf4c6 u ≤ x again?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 993
diff changeset
362
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
363 minsup : {x : Ordinal } → x o≤ z → MinSUP A (UnionCF A f ay supf x)
1032
382680c3e9be minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1031
diff changeset
364 minsup {x} x≤z = record { sup = supf x ; isMinSUP = is-minsup x≤z }
382680c3e9be minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1031
diff changeset
365
382680c3e9be minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1031
diff changeset
366 supf-is-minsup : {x : Ordinal } → (x≤z : x o≤ z) → supf x ≡ MinSUP.sup (minsup x≤z)
382680c3e9be minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1031
diff changeset
367 supf-is-minsup _ = refl
965
1c1c6a6ed4fa removing ch-init is no good because of initialization
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 964
diff changeset
368
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
369 -- different from order because y o< supf
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
370 fcy<sup : {u w : Ordinal } → u o≤ z → FClosure A f y w → (w ≡ supf u ) ∨ ( w << supf u )
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
371 fcy<sup {u} {w} u≤z fc with MinSUP.x≤sup (minsup u≤z) ⟪ subst (λ k → odef A k ) (sym &iso) (A∋fc {A} y f mf fc)
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
372 , ch-init (subst (λ k → FClosure A f y k) (sym &iso) fc ) ⟫
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
373 ... | case1 eq = case1 (subst (λ k → k ≡ supf u ) &iso (trans eq (sym (supf-is-minsup u≤z ) ) ))
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
374 ... | case2 lt = case2 (subst₂ (λ j k → j << k ) &iso (sym (supf-is-minsup u≤z )) lt )
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
375
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
376 initial : {x : Ordinal } → x o≤ z → odef (UnionCF A f ay supf x) x → y ≤ x
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
377 initial {x} x≤z ⟪ aa , ch-init fc ⟫ = s≤fc y f mf fc
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
378 initial {x} x≤z ⟪ aa , ch-is-sup u u<x is-sup fc ⟫ = ≤-ftrans (fcy<sup (ordtrans u<x x≤z) (init ay refl)) (s≤fc _ f mf fc)
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
379
1090
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
380 sup=u : {b : Ordinal} → (ab : odef A b) → b o≤ z
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
381 → IsSUP A (UnionCF A f ay supf b) b → supf b ≡ b
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
382 sup=u {b} ab b≤z is-sup = z50 where
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
383 z48 : supf b o≤ b
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
384 z48 = IsMinSUP.minsup (is-minsup b≤z) ab (λ ux → IsSUP.x≤sup is-sup ux )
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
385 z50 : supf b ≡ b
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
386 z50 with trio< (supf b) b
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
387 ... | tri< sb<b ¬b ¬c = ⊥-elim ( o≤> z47 sb<b ) where
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
388 z47 : b o≤ supf b
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
389 z47 = zo≤sz b≤z
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
390 ... | tri≈ ¬a b ¬c = b
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
391 ... | tri> ¬a ¬b b<sb = ⊥-elim ( o≤> z48 b<sb )
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
392
1091
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
393 supfeq : {a b : Ordinal } → a o≤ z → b o≤ z → UnionCF A f ay supf a ≡ UnionCF A f ay supf b → supf a ≡ supf b
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
394 supfeq {a} {b} a≤z b≤z ua=ub with trio< (supf a) (supf b)
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
395 ... | tri< sa<sb ¬b ¬c = ⊥-elim ( o≤> (
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
396 IsMinSUP.minsup (is-minsup b≤z) asupf (λ {z} uzb → IsMinSUP.x≤sup (is-minsup a≤z) (subst (λ k → odef k z) (sym ua=ub) uzb)) ) sa<sb )
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
397 ... | tri≈ ¬a b ¬c = b
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
398 ... | tri> ¬a ¬b sb<sa = ⊥-elim ( o≤> (
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
399 IsMinSUP.minsup (is-minsup a≤z) asupf (λ {z} uza → IsMinSUP.x≤sup (is-minsup b≤z) (subst (λ k → odef k z) ua=ub uza)) ) sb<sa )
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
400
1090
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
401 union-max : {a b : Ordinal } → b o≤ supf a → b o≤ z → supf a o< supf b → UnionCF A f ay supf a ≡ UnionCF A f ay supf b
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
402 union-max {a} {b} b≤sa b≤z sa<sb = ==→o≡ record { eq→ = z47 ; eq← = z48 } where
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
403 z47 : {w : Ordinal } → odef (UnionCF A f ay supf a ) w → odef ( UnionCF A f ay supf b ) w
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
404 z47 {w} ⟪ aw , ch-init fc ⟫ = ⟪ aw , ch-init fc ⟫
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
405 z47 {w} ⟪ aw , ch-is-sup u u<a su=u fc ⟫ = ⟪ aw , ch-is-sup u u<b su=u fc ⟫ where
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
406 u<b : u o< b
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
407 u<b = ordtrans u<a (supf-inject sa<sb )
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
408 z48 : {w : Ordinal } → odef (UnionCF A f ay supf b ) w → odef ( UnionCF A f ay supf a ) w
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
409 z48 {w} ⟪ aw , ch-init fc ⟫ = ⟪ aw , ch-init fc ⟫
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
410 z48 {w} ⟪ aw , ch-is-sup u u<b su=u fc ⟫ = ⟪ aw , ch-is-sup u u<a su=u fc ⟫ where
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
411 u<a : u o< a
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
412 u<a = supf-inject ( osucprev (begin
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
413 osuc (supf u) ≡⟨ cong osuc su=u ⟩
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
414 osuc u ≤⟨ osucc u<b ⟩
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
415 b ≤⟨ b≤sa ⟩
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
416 supf a ∎ )) where open o≤-Reasoning O
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
417
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
418 x≤supfx→¬sa<sa : {a b : Ordinal } → b o≤ z → b o≤ supf a → ¬ (supf a o< supf b )
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
419 x≤supfx→¬sa<sa {a} {b} b≤z b≤sa sa<sb = ⊥-elim ( o<¬≡ z27 sa<sb ) where -- x o≤ supf a ∧ supf a o< supf b → ⊥, because it defines the same UnionCF
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
420 z27 : supf a ≡ supf b
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
421 z27 = supfeq (ordtrans (supf-inject sa<sb) b≤z) b≤z ( union-max b≤sa b≤z sa<sb)
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
422
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
423 order : {a b w : Ordinal } → b o≤ z → supf a o< supf b → FClosure A f (supf a) w → w ≤ supf b
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
424 order {a} {b} {w} b≤z sa<sb fc = IsMinSUP.x≤sup (is-minsup b≤z) (cfcs (supf-inject sa<sb) b≤z sa<b fc) where
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
425 sa<b : supf a o< b
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
426 sa<b with x<y∨y≤x (supf a) b
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
427 ... | case1 lt = lt
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
428 ... | case2 b≤sa = ⊥-elim (x≤supfx→¬sa<sa b≤z b≤sa sa<sb)
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
429
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
430 supf-idem : {b : Ordinal } → b o≤ z → supf b o≤ z → supf (supf b) ≡ supf b
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
431 supf-idem {b} b≤z sfb≤x = z52 where
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
432 z54 : {w : Ordinal} → odef (UnionCF A f ay supf (supf b)) w → (w ≡ supf b) ∨ (w << supf b)
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
433 z54 {w} ⟪ aw , ch-init fc ⟫ = fcy<sup b≤z fc
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
434 z54 {w} ⟪ aw , ch-is-sup u u<x su=u fc ⟫ = order b≤z (subst (λ k → k o< supf b) (sym su=u) u<x) fc where
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
435 u<b : u o< b
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
436 u<b = supf-inject (subst (λ k → k o< supf b ) (sym (su=u)) u<x )
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
437 z52 : supf (supf b) ≡ supf b
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
438 z52 = sup=u asupf sfb≤x record { ax = asupf ; x≤sup = z54 }
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
439
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
440 supf-mono< : {a b : Ordinal } → b o≤ z → supf a o< supf b → supf a << supf b
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
441 supf-mono< {a} {b} b≤z sa<sb with order {a} {b} b≤z sa<sb (init asupf refl)
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
442 ... | case2 lt = lt
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
443 ... | case1 eq = ⊥-elim ( o<¬≡ eq sa<sb )
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
444
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
445 f-total : IsTotalOrderSet chain
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
446 f-total {a} {b} ⟪ uaa , ch-is-sup ua sua<x sua=ua fca ⟫ ⟪ uab , ch-is-sup ub sub<x sub=ub fcb ⟫ =
1035
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1034
diff changeset
447 subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso fc-total where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1034
diff changeset
448 fc-total : Tri (* (& a) < * (& b)) (* (& a) ≡ * (& b)) (* (& b) < * (& a) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1034
diff changeset
449 fc-total with trio< ua ub
1046
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1045
diff changeset
450 ... | tri< a₁ ¬b ¬c with ≤-ftrans (order (o<→≤ sub<x) (subst₂ (λ j k → j o< k) (sym sua=ua) (sym sub=ub) a₁) fca ) (s≤fc (supf ub) f mf fcb )
1035
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1034
diff changeset
451 ... | case1 eq1 = tri≈ (λ lt → ⊥-elim (<-irr (case1 (sym ct00)) lt)) ct00 (λ lt → ⊥-elim (<-irr (case1 ct00) lt)) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1034
diff changeset
452 ct00 : * (& a) ≡ * (& b)
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
453 ct00 = cong (*) eq1
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
454 ... | case2 a<b = tri< a<b (λ eq → <-irr (case1 (sym eq)) a<b ) (λ lt → <-irr (case2 a<b ) lt)
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
455 fc-total | tri≈ _ refl _ = fcn-cmp _ f mf fca fcb
1046
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1045
diff changeset
456 fc-total | tri> ¬a ¬b c with ≤-ftrans (order (o<→≤ sua<x) (subst₂ (λ j k → j o< k) (sym sub=ub) (sym sua=ua) c) fcb ) (s≤fc (supf ua) f mf fca )
1035
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1034
diff changeset
457 ... | case1 eq1 = tri≈ (λ lt → ⊥-elim (<-irr (case1 (sym ct00)) lt)) ct00 (λ lt → ⊥-elim (<-irr (case1 ct00) lt)) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1034
diff changeset
458 ct00 : * (& a) ≡ * (& b)
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
459 ct00 = cong (*) (sym eq1)
1035
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1034
diff changeset
460 ... | case2 b<a = tri> (λ lt → <-irr (case2 b<a ) lt) (λ eq → <-irr (case1 eq) b<a ) b<a
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
461 f-total {a} {b} ⟪ uaa , ch-init fca ⟫ ⟪ uab , ch-is-sup ub sub<x sub=ub fcb ⟫ = ft00 where
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
462 ft01 : (& a) ≤ (& b) → Tri ( a < b) ( a ≡ b) ( b < a )
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
463 ft01 (case1 eq) = tri≈ (λ lt → ⊥-elim (<-irr (case1 (sym a=b)) lt)) a=b (λ lt → ⊥-elim (<-irr (case1 a=b) lt)) where
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
464 a=b : a ≡ b
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
465 a=b = subst₂ (λ j k → j ≡ k ) *iso *iso (cong (*) eq)
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
466 ft01 (case2 lt) = tri< a<b (λ eq → <-irr (case1 (sym eq)) a<b ) (λ lt → <-irr (case2 a<b ) lt) where
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
467 a<b : a < b
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
468 a<b = subst₂ (λ j k → j < k ) *iso *iso lt
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
469 ft00 : Tri ( a < b) ( a ≡ b) ( b < a )
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
470 ft00 = ft01 (≤-ftrans (fcy<sup (o<→≤ sub<x) fca) (s≤fc {A} _ f mf fcb))
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
471 f-total {a} {b} ⟪ uaa , ch-is-sup ua sua<x sua=ua fca ⟫ ⟪ uab , ch-init fcb ⟫ = ft00 where
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
472 ft01 : (& b) ≤ (& a) → Tri ( a < b) ( a ≡ b) ( b < a )
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
473 ft01 (case1 eq) = tri≈ (λ lt → ⊥-elim (<-irr (case1 (sym a=b)) lt)) a=b (λ lt → ⊥-elim (<-irr (case1 a=b) lt)) where
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
474 a=b : a ≡ b
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
475 a=b = subst₂ (λ j k → j ≡ k ) *iso *iso (cong (*) (sym eq))
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
476 ft01 (case2 lt) = tri> (λ lt → <-irr (case2 b<a ) lt) (λ eq → <-irr (case1 eq) b<a ) b<a where
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
477 b<a : b < a
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
478 b<a = subst₂ (λ j k → j < k ) *iso *iso lt
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
479 ft00 : Tri ( a < b) ( a ≡ b) ( b < a )
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
480 ft00 = ft01 (≤-ftrans (fcy<sup (o<→≤ sua<x) fcb) (s≤fc {A} _ f mf fca))
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
481 f-total {a} {b} ⟪ uaa , ch-init fca ⟫ ⟪ uab , ch-init fcb ⟫ =
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
482 subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso (fcn-cmp y f mf fca fcb )
825
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 824
diff changeset
483
1038
dfbac4b59bae mf< everywhere
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1037
diff changeset
484 record ZChain1 ( A : HOD ) ( f : Ordinal → Ordinal ) (mf< : <-monotonic-f A f)
dfbac4b59bae mf< everywhere
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1037
diff changeset
485 {y : Ordinal} (ay : odef A y) (zc : ZChain A f mf< ay (& A)) ( z : Ordinal ) : Set (Level.suc n) where
1027
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1026
diff changeset
486 supf = ZChain.supf zc
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1026
diff changeset
487 field
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
488 is-max : {a b : Ordinal } → (ca : odef (UnionCF A f ay supf z) a ) → b o< z → (ab : odef A b)
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
489 → HasPrev A (UnionCF A f ay supf z) f b ∨ IsSUP A (UnionCF A f ay supf b) b
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
490 → * a < * b → odef ((UnionCF A f ay supf z)) b
1027
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1026
diff changeset
491
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1026
diff changeset
492 record Maximal ( A : HOD ) : Set (Level.suc n) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1026
diff changeset
493 field
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1026
diff changeset
494 maximal : HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1026
diff changeset
495 as : A ∋ maximal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1026
diff changeset
496 ¬maximal<x : {x : HOD} → A ∋ x → ¬ maximal < x -- A is Partial, use negative
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1026
diff changeset
497
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1026
diff changeset
498 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1026
diff changeset
499 -- supf in TransFinite indution may differ each other, but it is the same because of the minimul sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1026
diff changeset
500 --
1038
dfbac4b59bae mf< everywhere
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1037
diff changeset
501 supf-unique : ( A : HOD ) ( f : Ordinal → Ordinal ) (mf< : <-monotonic-f A f)
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
502 {y xa xb : Ordinal} → (ay : odef A y) → (xa o≤ xb ) → (za : ZChain A f mf< ay xa ) (zb : ZChain A f mf< ay xb )
1007
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1006
diff changeset
503 → {z : Ordinal } → z o≤ xa → ZChain.supf za z ≡ ZChain.supf zb z
1038
dfbac4b59bae mf< everywhere
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1037
diff changeset
504 supf-unique A f mf< {y} {xa} {xb} ay xa≤xb za zb {z} z≤xa = TransFinite0 {λ z → z o≤ xa → ZChain.supf za z ≡ ZChain.supf zb z } ind z z≤xa where
1007
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1006
diff changeset
505 supfa = ZChain.supf za
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1006
diff changeset
506 supfb = ZChain.supf zb
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1006
diff changeset
507 ind : (x : Ordinal) → ((w : Ordinal) → w o< x → w o≤ xa → supfa w ≡ supfb w) → x o≤ xa → supfa x ≡ supfb x
1008
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1007
diff changeset
508 ind x prev x≤xa = sxa=sxb where
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
509 ma = ZChain.minsup za x≤xa
1008
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1007
diff changeset
510 mb = ZChain.minsup zb (OrdTrans x≤xa xa≤xb )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1007
diff changeset
511 spa = MinSUP.sup ma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1007
diff changeset
512 spb = MinSUP.sup mb
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1007
diff changeset
513 sax=spa : supfa x ≡ spa
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
514 sax=spa = ZChain.supf-is-minsup za x≤xa
1008
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1007
diff changeset
515 sbx=spb : supfb x ≡ spb
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1007
diff changeset
516 sbx=spb = ZChain.supf-is-minsup zb (OrdTrans x≤xa xa≤xb )
1007
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1006
diff changeset
517 sxa=sxb : supfa x ≡ supfb x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1006
diff changeset
518 sxa=sxb with trio< (supfa x) (supfb x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1006
diff changeset
519 ... | tri≈ ¬a b ¬c = b
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1006
diff changeset
520 ... | tri< a ¬b ¬c = ⊥-elim ( o≤> (
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1006
diff changeset
521 begin
1008
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1007
diff changeset
522 supfb x ≡⟨ sbx=spb ⟩
1032
382680c3e9be minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1031
diff changeset
523 spb ≤⟨ MinSUP.minsup mb (MinSUP.as ma) (λ {z} uzb → MinSUP.x≤sup ma (z53 uzb)) ⟩
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
524 spa ≡⟨ sym sax=spa ⟩
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
525 supfa x ∎ ) a ) where
1008
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1007
diff changeset
526 open o≤-Reasoning O
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
527 z53 : {z : Ordinal } → odef (UnionCF A f ay (ZChain.supf zb) x) z → odef (UnionCF A f ay (ZChain.supf za) x) z
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
528 z53 ⟪ as , ch-init fc ⟫ = ⟪ as , ch-init fc ⟫
1034
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
529 z53 {z} ⟪ as , ch-is-sup u u<x su=u fc ⟫ = ⟪ as , ch-is-sup u u<x (trans ua=ub su=u) z55 ⟫ where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
530 ua=ub : supfa u ≡ supfb u
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
531 ua=ub = prev u u<x (ordtrans u<x x≤xa )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
532 z55 : FClosure A f (ZChain.supf za u) z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
533 z55 = subst (λ k → FClosure A f k z ) (sym ua=ub) fc
1007
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1006
diff changeset
534 ... | tri> ¬a ¬b c = ⊥-elim ( o≤> (
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1006
diff changeset
535 begin
1008
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1007
diff changeset
536 supfa x ≡⟨ sax=spa ⟩
1032
382680c3e9be minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1031
diff changeset
537 spa ≤⟨ MinSUP.minsup ma (MinSUP.as mb) (λ uza → MinSUP.x≤sup mb (z53 uza)) ⟩
1008
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1007
diff changeset
538 spb ≡⟨ sym sbx=spb ⟩
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
539 supfb x ∎ ) c ) where
1009
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1008
diff changeset
540 open o≤-Reasoning O
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
541 z53 : {z : Ordinal } → odef (UnionCF A f ay (ZChain.supf za) x) z → odef (UnionCF A f ay (ZChain.supf zb) x) z
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
542 z53 ⟪ as , ch-init fc ⟫ = ⟪ as , ch-init fc ⟫
1034
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
543 z53 {z} ⟪ as , ch-is-sup u u<x su=u fc ⟫ = ⟪ as , ch-is-sup u u<x (trans ub=ua su=u) z55 ⟫ where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
544 ub=ua : supfb u ≡ supfa u
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
545 ub=ua = sym ( prev u u<x (ordtrans u<x x≤xa ))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
546 z55 : FClosure A f (ZChain.supf zb u) z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
547 z55 = subst (λ k → FClosure A f k z ) (sym ub=ua) fc
1028
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1027
diff changeset
548
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
549 Zorn-lemma : { A : HOD }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
550 → o∅ o< & A
1091
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
551 → ( ( B : HOD) → (B⊆A : B ⊆ A) → IsTotalOrderSet B → SUP A B ) -- SUP condition
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
552 → Maximal A
552
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 551
diff changeset
553 Zorn-lemma {A} 0<A supP = zorn00 where
571
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
554 <-irr0 : {a b : HOD} → A ∋ a → A ∋ b → (a ≡ b ) ∨ (a < b ) → b < a → ⊥
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 570
diff changeset
555 <-irr0 {a} {b} A∋a A∋b = <-irr
788
c164f4f7cfd1 u<x in UChain again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 787
diff changeset
556 z07 : {y : Ordinal} {A : HOD } → {P : Set n} → odef A y ∧ P → y o< & A
c164f4f7cfd1 u<x in UChain again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 787
diff changeset
557 z07 {y} {A} p = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) (proj1 p )))
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
558 s : HOD
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
559 s = ODC.minimal O A (λ eq → ¬x<0 ( subst (λ k → o∅ o< k ) (=od∅→≡o∅ eq) 0<A ))
568
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
560 as : A ∋ * ( & s )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
561 as = subst (λ k → odef A (& k) ) (sym *iso) ( ODC.x∋minimal O A (λ eq → ¬x<0 ( subst (λ k → o∅ o< k ) (=od∅→≡o∅ eq) 0<A )) )
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
562 as0 : odef A (& s )
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
563 as0 = subst (λ k → odef A k ) &iso as
547
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 546
diff changeset
564 s<A : & s o< & A
568
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 567
diff changeset
565 s<A = c<→o< (subst (λ k → odef A (& k) ) *iso as )
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
566 HasMaximal : HOD
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
567 HasMaximal = record { od = record { def = λ x → odef A x ∧ ( (m : Ordinal) → odef A m → ¬ (* x < * m)) } ; odmax = & A ; <odmax = z07 }
537
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
568 no-maximum : HasMaximal =h= od∅ → (x : Ordinal) → odef A x ∧ ((m : Ordinal) → odef A m → odef A x ∧ (¬ (* x < * m) )) → ⊥
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
569 no-maximum nomx x P = ¬x<0 (eq→ nomx {x} ⟪ proj1 P , (λ m ma p → proj2 ( proj2 P m ma ) p ) ⟫ )
532
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
570 Gtx : { x : HOD} → A ∋ x → HOD
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
571 Gtx {x} ax = record { od = record { def = λ y → odef A y ∧ (x < (* y)) } ; odmax = & A ; <odmax = z07 }
537
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
572 z08 : ¬ Maximal A → HasMaximal =h= od∅
804
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 803
diff changeset
573 z08 nmx = record { eq→ = λ {x} lt → ⊥-elim ( nmx record {maximal = * x ; as = subst (λ k → odef A k) (sym &iso) (proj1 lt)
537
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
574 ; ¬maximal<x = λ {y} ay → subst (λ k → ¬ (* x < k)) *iso (proj2 lt (& y) ay) } ) ; eq← = λ {y} lt → ⊥-elim ( ¬x<0 lt )}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
575 x-is-maximal : ¬ Maximal A → {x : Ordinal} → (ax : odef A x) → & (Gtx (subst (λ k → odef A k ) (sym &iso) ax)) ≡ o∅ → (m : Ordinal) → odef A m → odef A x ∧ (¬ (* x < * m))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
576 x-is-maximal nmx {x} ax nogt m am = ⟪ subst (λ k → odef A k) &iso (subst (λ k → odef A k ) (sym &iso) ax) , ¬x<m ⟫ where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
577 ¬x<m : ¬ (* x < * m)
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
578 ¬x<m x<m = ∅< {Gtx (subst (λ k → odef A k ) (sym &iso) ax)} {* m} ⟪ subst (λ k → odef A k) (sym &iso) am , subst (λ k → * x < k ) (cong (*) (sym &iso)) x<m ⟫ (≡o∅→=od∅ nogt)
543
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 542
diff changeset
579
1027
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1026
diff changeset
580 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1026
diff changeset
581 -- we have minsup using LEM, this is similar to the proof of the axiom of choice
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1026
diff changeset
582 --
1091
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
583 minsupP : ( B : HOD) → (B⊆A : B ⊆ A) → IsTotalOrderSet B → MinSUP A B
879
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
584 minsupP B B⊆A total = m02 where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
585 xsup : (sup : Ordinal ) → Set n
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
586 xsup sup = {w : Ordinal } → odef B w → (w ≡ sup ) ∨ (w << sup )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
587 ∀-imply-or : {A : Ordinal → Set n } {B : Set n }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
588 → ((x : Ordinal ) → A x ∨ B) → ((x : Ordinal ) → A x) ∨ B
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
589 ∀-imply-or {A} {B} ∀AB with ODC.p∨¬p O ((x : Ordinal ) → A x) -- LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
590 ∀-imply-or {A} {B} ∀AB | case1 t = case1 t
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
591 ∀-imply-or {A} {B} ∀AB | case2 x = case2 (lemma (λ not → x not )) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
592 lemma : ¬ ((x : Ordinal ) → A x) → B
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
593 lemma not with ODC.p∨¬p O B
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
594 lemma not | case1 b = b
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
595 lemma not | case2 ¬b = ⊥-elim (not (λ x → dont-orb (∀AB x) ¬b ))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
596 m00 : (x : Ordinal ) → ( ( z : Ordinal) → z o< x → ¬ (odef A z ∧ xsup z) ) ∨ MinSUP A B
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
597 m00 x = TransFinite0 ind x where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
598 ind : (x : Ordinal) → ((z : Ordinal) → z o< x → ( ( w : Ordinal) → w o< z → ¬ (odef A w ∧ xsup w )) ∨ MinSUP A B)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
599 → ( ( w : Ordinal) → w o< x → ¬ (odef A w ∧ xsup w) ) ∨ MinSUP A B
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
600 ind x prev = ∀-imply-or m01 where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
601 m01 : (z : Ordinal) → (z o< x → ¬ (odef A z ∧ xsup z)) ∨ MinSUP A B
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
602 m01 z with trio< z x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
603 ... | tri≈ ¬a b ¬c = case1 ( λ lt → ⊥-elim ( ¬a lt ) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
604 ... | tri> ¬a ¬b c = case1 ( λ lt → ⊥-elim ( ¬a lt ) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
605 ... | tri< a ¬b ¬c with prev z a
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
606 ... | case2 mins = case2 mins
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
607 ... | case1 not with ODC.p∨¬p O (odef A z ∧ xsup z)
1032
382680c3e9be minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1031
diff changeset
608 ... | case1 mins = case2 record { sup = z ; isMinSUP = record { as = proj1 mins ; x≤sup = proj2 mins ; minsup = m04 } } where
879
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
609 m04 : {sup1 : Ordinal} → odef A sup1 → ({w : Ordinal} → odef B w → (w ≡ sup1) ∨ (w << sup1)) → z o≤ sup1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
610 m04 {s} as lt with trio< z s
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
611 ... | tri< a ¬b ¬c = o<→≤ a
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
612 ... | tri≈ ¬a b ¬c = o≤-refl0 b
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
613 ... | tri> ¬a ¬b s<z = ⊥-elim ( not s s<z ⟪ as , lt ⟫ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
614 ... | case2 notz = case1 (λ _ → notz )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
615 m03 : ¬ ((z : Ordinal) → z o< & A → ¬ odef A z ∧ xsup z)
1031
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
616 m03 not = ⊥-elim ( not s1 (z09 (SUP.ax S)) ⟪ SUP.ax S , m05 ⟫ ) where
879
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
617 S : SUP A B
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
618 S = supP B B⊆A total
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
619 s1 = & (SUP.sup S)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
620 m05 : {w : Ordinal } → odef B w → (w ≡ s1 ) ∨ (w << s1 )
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
621 m05 {w} bw with SUP.x≤sup S bw
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
622 ... | case1 eq = case1 ( subst₂ (λ j k → j ≡ k ) &iso refl (trans &iso eq))
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
623 ... | case2 lt = case2 lt
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
624 m02 : MinSUP A B
879
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
625 m02 = dont-or (m00 (& A)) m03
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 878
diff changeset
626
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
627 -- Uncountable ascending chain by axiom of choice
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
628 cf : ¬ Maximal A → Ordinal → Ordinal
532
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
629 cf nmx x with ODC.∋-p O A (* x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
630 ... | no _ = o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
631 ... | yes ax with is-o∅ (& ( Gtx ax ))
538
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
632 ... | yes nogt = -- no larger element, so it is maximal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
633 ⊥-elim (no-maximum (z08 nmx) x ⟪ subst (λ k → odef A k) &iso ax , x-is-maximal nmx (subst (λ k → odef A k ) &iso ax) nogt ⟫ )
532
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
634 ... | no not = & (ODC.minimal O (Gtx ax) (λ eq → not (=od∅→≡o∅ eq)))
537
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
635 is-cf : (nmx : ¬ Maximal A ) → {x : Ordinal} → odef A x → odef A (cf nmx x) ∧ ( * x < * (cf nmx x) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
636 is-cf nmx {x} ax with ODC.∋-p O A (* x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
637 ... | no not = ⊥-elim ( not (subst (λ k → odef A k ) (sym &iso) ax ))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
638 ... | yes ax with is-o∅ (& ( Gtx ax ))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
639 ... | yes nogt = ⊥-elim (no-maximum (z08 nmx) x ⟪ subst (λ k → odef A k) &iso ax , x-is-maximal nmx (subst (λ k → odef A k ) &iso ax) nogt ⟫ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
640 ... | no not = ODC.x∋minimal O (Gtx ax) (λ eq → not (=od∅→≡o∅ eq))
606
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 605
diff changeset
641
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 605
diff changeset
642 ---
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 605
diff changeset
643 --- infintie ascention sequence of f
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 605
diff changeset
644 ---
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
645 cf-is-<-monotonic : (nmx : ¬ Maximal A ) → (x : Ordinal) → odef A x → ( * x < * (cf nmx x) ) ∧ odef A (cf nmx x )
537
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 536
diff changeset
646 cf-is-<-monotonic nmx x ax = ⟪ proj2 (is-cf nmx ax ) , proj1 (is-cf nmx ax ) ⟫
530
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 529
diff changeset
647 cf-is-≤-monotonic : (nmx : ¬ Maximal A ) → ≤-monotonic-f A ( cf nmx )
532
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
648 cf-is-≤-monotonic nmx x ax = ⟪ case2 (proj1 ( cf-is-<-monotonic nmx x ax )) , proj2 ( cf-is-<-monotonic nmx x ax ) ⟫
543
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 542
diff changeset
649
803
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 802
diff changeset
650 --
953
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 952
diff changeset
651 -- maximality of chain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 952
diff changeset
652 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 952
diff changeset
653 -- supf is fixed for z ≡ & A , we can prove order and is-max
1016
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1015
diff changeset
654 -- we have supf-unique now, it is provable in the first Tranfinte induction
803
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 802
diff changeset
655
992
4aaecae58da5 ... x < & A ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 991
diff changeset
656 SZ1 : ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f) (mf< : <-monotonic-f A f)
1038
dfbac4b59bae mf< everywhere
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1037
diff changeset
657 {init : Ordinal} (ay : odef A init) (zc : ZChain A f mf< ay (& A)) (x : Ordinal) → x o≤ & A → ZChain1 A f mf< ay zc x
993
e11c244d7eac SZ1 done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 992
diff changeset
658 SZ1 f mf mf< {y} ay zc x x≤A = zc1 x x≤A where
900
ac4daa43ef2a roll back to u<x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 899
diff changeset
659 chain-mono1 : {a b c : Ordinal} → a o≤ b
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
660 → odef (UnionCF A f ay (ZChain.supf zc) a) c → odef (UnionCF A f ay (ZChain.supf zc) b) c
919
213f12f27003 supf u o< supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 918
diff changeset
661 chain-mono1 {a} {b} {c} a≤b = chain-mono f mf ay (ZChain.supf zc) (ZChain.supf-mono zc) a≤b
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
662 is-max-hp : (x : Ordinal) {a : Ordinal} {b : Ordinal} → odef (UnionCF A f ay (ZChain.supf zc) x) a → (ab : odef A b)
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
663 → HasPrev A (UnionCF A f ay (ZChain.supf zc) x) f b
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
664 → * a < * b → odef (UnionCF A f ay (ZChain.supf zc) x) b
920
a2f8d14012aa fixpoint?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 919
diff changeset
665 is-max-hp x {a} {b} ua ab has-prev a<b with HasPrev.ay has-prev
1034
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
666 ... | ⟪ ab0 , ch-init fc ⟫ = ⟪ ab , ch-init ( subst (λ k → FClosure A f y k) (sym (HasPrev.x=fy has-prev)) (fsuc _ fc )) ⟫
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
667 ... | ⟪ ab0 , ch-is-sup u u<x su=u fc ⟫ = ⟪ ab , subst (λ k → UChain ay x k )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
668 (sym (HasPrev.x=fy has-prev)) ( ch-is-sup u u<x su=u (fsuc _ fc)) ⟫
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
669
869
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 868
diff changeset
670 supf = ZChain.supf zc
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 868
diff changeset
671
1038
dfbac4b59bae mf< everywhere
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1037
diff changeset
672 zc1 : (x : Ordinal ) → x o≤ & A → ZChain1 A f mf< ay zc x
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
673 zc1 x x≤A with Oprev-p x
1024
ab72526316bd supf-< and ZChain1.order is removed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1023
diff changeset
674 ... | yes op = record { is-max = is-max } where
988
9a85233384f7 is-max and supf b = supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 987
diff changeset
675 px = Oprev.oprev op
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
676 is-max : {a : Ordinal} {b : Ordinal} → odef (UnionCF A f ay supf x) a →
988
9a85233384f7 is-max and supf b = supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 987
diff changeset
677 b o< x → (ab : odef A b) →
1034
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
678 HasPrev A (UnionCF A f ay supf x) f b ∨ IsSUP A (UnionCF A f ay supf b) b →
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
679 * a < * b → odef (UnionCF A f ay supf x) b
988
9a85233384f7 is-max and supf b = supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 987
diff changeset
680 is-max {a} {b} ua b<x ab P a<b with ODC.or-exclude O P
9a85233384f7 is-max and supf b = supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 987
diff changeset
681 is-max {a} {b} ua b<x ab P a<b | case1 has-prev = is-max-hp x {a} {b} ua ab has-prev a<b
989
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 988
diff changeset
682 is-max {a} {b} ua b<x ab P a<b | case2 is-sup with osuc-≡< (ZChain.supf-mono zc (o<→≤ b<x))
1024
ab72526316bd supf-< and ZChain1.order is removed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1023
diff changeset
683 ... | case2 sb<sx = m10 where
988
9a85233384f7 is-max and supf b = supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 987
diff changeset
684 b<A : b o< & A
9a85233384f7 is-max and supf b = supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 987
diff changeset
685 b<A = z09 ab
9a85233384f7 is-max and supf b = supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 987
diff changeset
686 m05 : ZChain.supf zc b ≡ b
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
687 m05 = ZChain.sup=u zc ab (o<→≤ (z09 ab) ) record { ax = ab ; x≤sup = λ {z} uz → IsSUP.x≤sup (proj2 is-sup) uz }
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
688 m10 : odef (UnionCF A f ay supf x) b
1039
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1038
diff changeset
689 m10 = ZChain.cfcs zc b<x x≤A (subst (λ k → k o< x) (sym m05) b<x) (init (ZChain.asupf zc) m05)
992
4aaecae58da5 ... x < & A ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 991
diff changeset
690 ... | case1 sb=sx = ⊥-elim (<-irr (case1 (cong (*) m10)) (proj1 (mf< (supf b) (ZChain.asupf zc)))) where
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
691 m17 : MinSUP A (UnionCF A f ay supf x) -- supf z o< supf ( supf x )
992
4aaecae58da5 ... x < & A ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 991
diff changeset
692 m17 = ZChain.minsup zc x≤A
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
693 m18 : supf x ≡ MinSUP.sup m17
992
4aaecae58da5 ... x < & A ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 991
diff changeset
694 m18 = ZChain.supf-is-minsup zc x≤A
990
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 989
diff changeset
695 m10 : f (supf b) ≡ supf b
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 989
diff changeset
696 m10 = fc-stop A f mf (ZChain.asupf zc) m11 sb=sx where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 989
diff changeset
697 m11 : {z : Ordinal} → FClosure A f (supf b) z → (z ≡ ZChain.supf zc x) ∨ (z << ZChain.supf zc x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 989
diff changeset
698 m11 {z} fc = subst (λ k → (z ≡ k) ∨ (z << k)) (sym m18) ( MinSUP.x≤sup m17 m13 ) where
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
699 m04 : ¬ HasPrev A (UnionCF A f ay supf b) f b
1019
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1018
diff changeset
700 m04 nhp = proj1 is-sup ( record { ax = HasPrev.ax nhp ; y = HasPrev.y nhp ; ay =
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1018
diff changeset
701 chain-mono1 (o<→≤ b<x) (HasPrev.ay nhp) ; x=fy = HasPrev.x=fy nhp } )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1018
diff changeset
702 m05 : ZChain.supf zc b ≡ b
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
703 m05 = ZChain.sup=u zc ab (o<→≤ (z09 ab) ) record { ax = ab ; x≤sup = λ {z} uz → IsSUP.x≤sup (proj2 is-sup) uz }
1019
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1018
diff changeset
704 m14 : ZChain.supf zc b o< x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1018
diff changeset
705 m14 = subst (λ k → k o< x ) (sym m05) b<x
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
706 m13 : odef (UnionCF A f ay supf x) z
1039
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1038
diff changeset
707 m13 = ZChain.cfcs zc b<x x≤A m14 fc
989
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 988
diff changeset
708
1024
ab72526316bd supf-< and ZChain1.order is removed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1023
diff changeset
709 ... | no lim = record { is-max = is-max } where
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
710 is-max : {a : Ordinal} {b : Ordinal} → odef (UnionCF A f ay supf x) a →
988
9a85233384f7 is-max and supf b = supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 987
diff changeset
711 b o< x → (ab : odef A b) →
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
712 HasPrev A (UnionCF A f ay supf x) f b ∨ IsSUP A (UnionCF A f ay supf b) b →
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
713 * a < * b → odef (UnionCF A f ay supf x) b
988
9a85233384f7 is-max and supf b = supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 987
diff changeset
714 is-max {a} {b} ua b<x ab P a<b with ODC.or-exclude O P
9a85233384f7 is-max and supf b = supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 987
diff changeset
715 is-max {a} {b} ua b<x ab P a<b | case1 has-prev = is-max-hp x {a} {b} ua ab has-prev a<b
1058
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1057
diff changeset
716 is-max {a} {b} ua b<x ab P a<b | case2 is-sup with IsSUP.x≤sup (proj2 is-sup) (ZChain.chain∋init zc )
1034
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
717 ... | case1 b=y = ⟪ subst (λ k → odef A k ) b=y ay , ch-init (subst (λ k → FClosure A f y k ) b=y (init ay refl )) ⟫
990
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 989
diff changeset
718 ... | case2 y<b with osuc-≡< (ZChain.supf-mono zc (o<→≤ b<x))
1024
ab72526316bd supf-< and ZChain1.order is removed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1023
diff changeset
719 ... | case2 sb<sx = m10 where
988
9a85233384f7 is-max and supf b = supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 987
diff changeset
720 m09 : b o< & A
9a85233384f7 is-max and supf b = supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 987
diff changeset
721 m09 = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) ab))
9a85233384f7 is-max and supf b = supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 987
diff changeset
722 m05 : ZChain.supf zc b ≡ b
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
723 m05 = ZChain.sup=u zc ab (o<→≤ m09) record { ax = ab ; x≤sup = λ lt → IsSUP.x≤sup (proj2 is-sup) lt }
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
724 m10 : odef (UnionCF A f ay supf x) b
1039
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1038
diff changeset
725 m10 = ZChain.cfcs zc b<x x≤A (subst (λ k → k o< x) (sym m05) b<x) (init (ZChain.asupf zc) m05)
992
4aaecae58da5 ... x < & A ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 991
diff changeset
726 ... | case1 sb=sx = ⊥-elim (<-irr (case1 (cong (*) m10)) (proj1 (mf< (supf b) (ZChain.asupf zc)))) where
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
727 m17 : MinSUP A (UnionCF A f ay supf x) -- supf z o< supf ( supf x )
992
4aaecae58da5 ... x < & A ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 991
diff changeset
728 m17 = ZChain.minsup zc x≤A
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
729 m18 : supf x ≡ MinSUP.sup m17
992
4aaecae58da5 ... x < & A ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 991
diff changeset
730 m18 = ZChain.supf-is-minsup zc x≤A
990
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 989
diff changeset
731 m10 : f (supf b) ≡ supf b
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 989
diff changeset
732 m10 = fc-stop A f mf (ZChain.asupf zc) m11 sb=sx where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 989
diff changeset
733 m11 : {z : Ordinal} → FClosure A f (supf b) z → (z ≡ ZChain.supf zc x) ∨ (z << ZChain.supf zc x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 989
diff changeset
734 m11 {z} fc = subst (λ k → (z ≡ k) ∨ (z << k)) (sym m18) ( MinSUP.x≤sup m17 m13 ) where
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
735 m05 = ZChain.sup=u zc ab (o<→≤ (z09 ab) ) record { ax = ab ; x≤sup = λ {z} uz → IsSUP.x≤sup (proj2 is-sup) uz }
1019
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1018
diff changeset
736 m14 : ZChain.supf zc b o< x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1018
diff changeset
737 m14 = subst (λ k → k o< x ) (sym m05) b<x
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
738 m13 : odef (UnionCF A f ay supf x) z
1039
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1038
diff changeset
739 m13 = ZChain.cfcs zc b<x x≤A m14 fc
727
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 726
diff changeset
740
757
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 756
diff changeset
741 uchain : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y) → HOD
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
742 uchain f mf {y} ay = record { od = record { def = λ x → FClosure A f y x } ; odmax = & A ; <odmax =
757
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 756
diff changeset
743 λ {z} cz → subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) (A∋fc y f mf cz ))) }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 756
diff changeset
744
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
745 utotal : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y)
757
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 756
diff changeset
746 → IsTotalOrderSet (uchain f mf ay)
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
747 utotal f mf {y} ay {a} {b} ca cb = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso uz01 where
757
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 756
diff changeset
748 uz01 : Tri (* (& a) < * (& b)) (* (& a) ≡ * (& b)) (* (& b) < * (& a) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 756
diff changeset
749 uz01 = fcn-cmp y f mf ca cb
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 756
diff changeset
750
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
751 ysup : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y)
928
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 927
diff changeset
752 → MinSUP A (uchain f mf ay)
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
753 ysup f mf {y} ay = minsupP (uchain f mf ay) (λ lt → A∋fc y f mf lt) (utotal f mf ay)
757
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 756
diff changeset
754
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
755 --
547
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 546
diff changeset
756 -- create all ZChains under o< x
560
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 559
diff changeset
757 --
608
6655f03984f9 mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 607
diff changeset
758
1038
dfbac4b59bae mf< everywhere
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1037
diff changeset
759 ind : ( f : Ordinal → Ordinal ) → (mf< : <-monotonic-f A f ) {y : Ordinal} (ay : odef A y) → (x : Ordinal)
dfbac4b59bae mf< everywhere
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1037
diff changeset
760 → ((z : Ordinal) → z o< x → ZChain A f mf< ay z) → ZChain A f mf< ay x
dfbac4b59bae mf< everywhere
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1037
diff changeset
761 ind f mf< {y} ay x prev with Oprev-p x
1089
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
762 ... | yes op = zc41 sup1 where
682
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 681
diff changeset
763 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 681
diff changeset
764 -- we have previous ordinal to use induction
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 681
diff changeset
765 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 681
diff changeset
766 px = Oprev.oprev op
1038
dfbac4b59bae mf< everywhere
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1037
diff changeset
767 zc : ZChain A f mf< ay (Oprev.oprev op)
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
768 zc = prev px (subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc )
682
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 681
diff changeset
769 px<x : px o< x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 681
diff changeset
770 px<x = subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc
918
4c33f8383d7d supf px o< px is in csupf
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 911
diff changeset
771 opx=x : osuc px ≡ x
4c33f8383d7d supf px o< px is in csupf
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 911
diff changeset
772 opx=x = Oprev.oprev=x op
4c33f8383d7d supf px o< px is in csupf
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 911
diff changeset
773
709
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 708
diff changeset
774 zc-b<x : (b : Ordinal ) → b o< x → b o< osuc px
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
775 zc-b<x b lt = subst (λ k → b o< k ) (sym (Oprev.oprev=x op)) lt
697
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 696
diff changeset
776
754
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 750
diff changeset
777 supf0 = ZChain.supf zc
869
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 868
diff changeset
778 pchain : HOD
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
779 pchain = UnionCF A f ay supf0 px
835
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 834
diff changeset
780
857
266e0b9027cd supf-max
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 856
diff changeset
781 supf-mono = ZChain.supf-mono zc
844
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 843
diff changeset
782
861
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 860
diff changeset
783 zc04 : {b : Ordinal} → b o≤ x → (b o≤ px ) ∨ (b ≡ x )
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
784 zc04 {b} b≤x with trio< b px
861
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 860
diff changeset
785 ... | tri< a ¬b ¬c = case1 (o<→≤ a)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 860
diff changeset
786 ... | tri≈ ¬a b ¬c = case1 (o≤-refl0 b)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 860
diff changeset
787 ... | tri> ¬a ¬b px<b with osuc-≡< b≤x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 860
diff changeset
788 ... | case1 eq = case2 eq
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
789 ... | case2 b<x = ⊥-elim ( ¬p<x<op ⟪ px<b , subst (λ k → b o< k ) (sym (Oprev.oprev=x op)) b<x ⟫ )
840
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 839
diff changeset
790
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
791 mf : ≤-monotonic-f A f
1039
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1038
diff changeset
792 mf x ax = ⟪ case2 mf00 , proj2 (mf< x ax ) ⟫ where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1038
diff changeset
793 mf00 : * x < * (f x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1038
diff changeset
794 mf00 = proj1 ( mf< x ax )
1038
dfbac4b59bae mf< everywhere
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1037
diff changeset
795
954
e43a5cc72287 IsSUP is now min sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 953
diff changeset
796 --
e43a5cc72287 IsSUP is now min sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 953
diff changeset
797 -- find the next value of supf
e43a5cc72287 IsSUP is now min sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 953
diff changeset
798 --
e43a5cc72287 IsSUP is now min sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 953
diff changeset
799
e43a5cc72287 IsSUP is now min sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 953
diff changeset
800 pchainpx : HOD
1034
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
801 pchainpx = record { od = record { def = λ z → (odef A z ∧ UChain ay px z )
1043
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1042
diff changeset
802 ∨ (FClosure A f (supf0 px) z ∧ (supf0 px o< x)) } ; odmax = & A ; <odmax = zc00 } where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1042
diff changeset
803 zc00 : {z : Ordinal } → (odef A z ∧ UChain ay px z ) ∨ (FClosure A f (supf0 px) z ∧ (supf0 px o< x) )→ z o< & A
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
804 zc00 {z} (case1 lt) = z07 lt
1043
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1042
diff changeset
805 zc00 {z} (case2 fc) = z09 ( A∋fc (supf0 px) f mf (proj1 fc) )
954
e43a5cc72287 IsSUP is now min sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 953
diff changeset
806
1043
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1042
diff changeset
807 zc02 : { a b : Ordinal } → odef A a ∧ UChain ay px a → FClosure A f (supf0 px) b ∧ ( supf0 px o< x) → a ≤ b
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1042
diff changeset
808 zc02 {a} {b} ca fb = zc05 (proj1 fb) where
1031
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
809 zc05 : {b : Ordinal } → FClosure A f (supf0 px) b → a ≤ b
954
e43a5cc72287 IsSUP is now min sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 953
diff changeset
810 zc05 (fsuc b1 fb ) with proj1 ( mf b1 (A∋fc (supf0 px) f mf fb ))
1034
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
811 ... | case1 eq = subst (λ k → a ≤ k ) eq (zc05 fb)
1031
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
812 ... | case2 lt = ≤-ftrans (zc05 fb) (case2 lt)
1034
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
813 zc05 (init b1 refl) = MinSUP.x≤sup (ZChain.minsup zc o≤-refl) ca
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
814
954
e43a5cc72287 IsSUP is now min sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 953
diff changeset
815 ptotal : IsTotalOrderSet pchainpx
1034
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
816 ptotal (case1 a) (case1 b) = ZChain.f-total zc a b
954
e43a5cc72287 IsSUP is now min sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 953
diff changeset
817 ptotal {a0} {b0} (case1 a) (case2 b) with zc02 a b
e43a5cc72287 IsSUP is now min sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 953
diff changeset
818 ... | case1 eq = tri≈ (<-irr (case1 (sym eq1))) eq1 (<-irr (case1 eq1)) where
e43a5cc72287 IsSUP is now min sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 953
diff changeset
819 eq1 : a0 ≡ b0
e43a5cc72287 IsSUP is now min sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 953
diff changeset
820 eq1 = subst₂ (λ j k → j ≡ k ) *iso *iso (cong (*) eq )
e43a5cc72287 IsSUP is now min sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 953
diff changeset
821 ... | case2 lt = tri< lt1 (λ eq → <-irr (case1 (sym eq)) lt1) (<-irr (case2 lt1)) where
e43a5cc72287 IsSUP is now min sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 953
diff changeset
822 lt1 : a0 < b0
e43a5cc72287 IsSUP is now min sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 953
diff changeset
823 lt1 = subst₂ (λ j k → j < k ) *iso *iso lt
e43a5cc72287 IsSUP is now min sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 953
diff changeset
824 ptotal {b0} {a0} (case2 b) (case1 a) with zc02 a b
e43a5cc72287 IsSUP is now min sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 953
diff changeset
825 ... | case1 eq = tri≈ (<-irr (case1 eq1)) (sym eq1) (<-irr (case1 (sym eq1))) where
e43a5cc72287 IsSUP is now min sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 953
diff changeset
826 eq1 : a0 ≡ b0
e43a5cc72287 IsSUP is now min sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 953
diff changeset
827 eq1 = subst₂ (λ j k → j ≡ k ) *iso *iso (cong (*) eq )
e43a5cc72287 IsSUP is now min sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 953
diff changeset
828 ... | case2 lt = tri> (<-irr (case2 lt1)) (λ eq → <-irr (case1 eq) lt1) lt1 where
e43a5cc72287 IsSUP is now min sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 953
diff changeset
829 lt1 : a0 < b0
e43a5cc72287 IsSUP is now min sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 953
diff changeset
830 lt1 = subst₂ (λ j k → j < k ) *iso *iso lt
1043
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1042
diff changeset
831 ptotal (case2 a) (case2 b) = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso (fcn-cmp (supf0 px) f mf (proj1 a) (proj1 b))
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
832
1091
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
833 pcha : pchainpx ⊆ A
954
e43a5cc72287 IsSUP is now min sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 953
diff changeset
834 pcha (case1 lt) = proj1 lt
1043
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1042
diff changeset
835 pcha (case2 fc) = A∋fc _ f mf (proj1 fc)
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
836
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
837 sup1 : MinSUP A pchainpx
954
e43a5cc72287 IsSUP is now min sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 953
diff changeset
838 sup1 = minsupP pchainpx pcha ptotal
e43a5cc72287 IsSUP is now min sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 953
diff changeset
839
e43a5cc72287 IsSUP is now min sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 953
diff changeset
840 --
e43a5cc72287 IsSUP is now min sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 953
diff changeset
841 -- supf0 px o≤ sp1
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
842 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
843
1089
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
844 zc41 : MinSUP A pchainpx → ZChain A f mf< ay x
1090
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
845 zc41 sup1 = record { supf = supf1 ; asupf = asupf1 ; zo≤sz = zo≤sz ; is-minsup = is-minsup ; cfcs = cfcs ; supf-mono = supf1-mono } where
1034
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
846
1089
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
847 sp1 = MinSUP.sup sup1
883
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 882
diff changeset
848
871
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 870
diff changeset
849 supf1 : Ordinal → Ordinal
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
850 supf1 z with trio< z px
871
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 870
diff changeset
851 ... | tri< a ¬b ¬c = supf0 z
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
852 ... | tri≈ ¬a b ¬c = supf0 z
901
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 900
diff changeset
853 ... | tri> ¬a ¬b c = sp1
871
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 870
diff changeset
854
886
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 885
diff changeset
855 sf1=sf0 : {z : Ordinal } → z o≤ px → supf1 z ≡ supf0 z
901
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 900
diff changeset
856 sf1=sf0 {z} z≤px with trio< z px
874
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 873
diff changeset
857 ... | tri< a ¬b ¬c = refl
901
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 900
diff changeset
858 ... | tri≈ ¬a b ¬c = refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 900
diff changeset
859 ... | tri> ¬a ¬b c = ⊥-elim ( o≤> z≤px c )
883
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 882
diff changeset
860
901
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 900
diff changeset
861 sf1=sp1 : {z : Ordinal } → px o< z → supf1 z ≡ sp1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 900
diff changeset
862 sf1=sp1 {z} px<z with trio< z px
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 900
diff changeset
863 ... | tri< a ¬b ¬c = ⊥-elim ( o<> px<z a )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 900
diff changeset
864 ... | tri≈ ¬a b ¬c = ⊥-elim ( o<¬≡ (sym b) px<z )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 900
diff changeset
865 ... | tri> ¬a ¬b c = refl
873
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 872
diff changeset
866
968
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 967
diff changeset
867 sf=eq : { z : Ordinal } → z o< x → supf0 z ≡ supf1 z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 967
diff changeset
868 sf=eq {z} z<x = sym (sf1=sf0 (subst (λ k → z o< k ) (sym (Oprev.oprev=x op)) z<x ))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 967
diff changeset
869
903
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 902
diff changeset
870 asupf1 : {z : Ordinal } → odef A (supf1 z)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 902
diff changeset
871 asupf1 {z} with trio< z px
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
872 ... | tri< a ¬b ¬c = ZChain.asupf zc
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
873 ... | tri≈ ¬a b ¬c = ZChain.asupf zc
1032
382680c3e9be minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1031
diff changeset
874 ... | tri> ¬a ¬b c = MinSUP.as sup1
903
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 902
diff changeset
875
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
876 supf1-mono : {a b : Ordinal } → a o≤ b → supf1 a o≤ supf1 b
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
877 supf1-mono {a} {b} a≤b with trio< b px
901
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 900
diff changeset
878 ... | tri< a ¬b ¬c = subst₂ (λ j k → j o≤ k ) (sym (sf1=sf0 (o<→≤ (ordtrans≤-< a≤b a)))) refl ( supf-mono a≤b )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 900
diff changeset
879 ... | tri≈ ¬a b ¬c = subst₂ (λ j k → j o≤ k ) (sym (sf1=sf0 (subst (λ k → a o≤ k) b a≤b))) refl ( supf-mono a≤b )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 900
diff changeset
880 supf1-mono {a} {b} a≤b | tri> ¬a ¬b c with trio< a px
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 900
diff changeset
881 ... | tri< a<px ¬b ¬c = zc19 where
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
882 zc21 : MinSUP A (UnionCF A f ay supf0 a)
901
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 900
diff changeset
883 zc21 = ZChain.minsup zc (o<→≤ a<px)
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
884 zc24 : {x₁ : Ordinal} → odef (UnionCF A f ay supf0 a) x₁ → (x₁ ≡ sp1) ∨ (x₁ << sp1)
950
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 949
diff changeset
885 zc24 {x₁} ux = MinSUP.x≤sup sup1 (case1 (chain-mono f mf ay supf0 (ZChain.supf-mono zc) (o<→≤ a<px) ux ) )
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
886 zc19 : supf0 a o≤ sp1
1032
382680c3e9be minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1031
diff changeset
887 zc19 = subst (λ k → k o≤ sp1) (sym (ZChain.supf-is-minsup zc (o<→≤ a<px))) ( MinSUP.minsup zc21 (MinSUP.as sup1) zc24 )
901
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 900
diff changeset
888 ... | tri≈ ¬a b ¬c = zc18 where
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
889 zc21 : MinSUP A (UnionCF A f ay supf0 a)
901
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 900
diff changeset
890 zc21 = ZChain.minsup zc (o≤-refl0 b)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 900
diff changeset
891 zc20 : MinSUP.sup zc21 ≡ supf0 a
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
892 zc20 = sym (ZChain.supf-is-minsup zc (o≤-refl0 b))
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
893 zc24 : {x₁ : Ordinal} → odef (UnionCF A f ay supf0 a) x₁ → (x₁ ≡ sp1) ∨ (x₁ << sp1)
950
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 949
diff changeset
894 zc24 {x₁} ux = MinSUP.x≤sup sup1 (case1 (chain-mono f mf ay supf0 (ZChain.supf-mono zc) (o≤-refl0 b) ux ) )
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
895 zc18 : supf0 a o≤ sp1
1032
382680c3e9be minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1031
diff changeset
896 zc18 = subst (λ k → k o≤ sp1) zc20( MinSUP.minsup zc21 (MinSUP.as sup1) zc24 )
901
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 900
diff changeset
897 ... | tri> ¬a ¬b c = o≤-refl
885
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 884
diff changeset
898
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
899 fcup : {u z : Ordinal } → FClosure A f (supf1 u) z → u o≤ px → FClosure A f (supf0 u) z
903
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 902
diff changeset
900 fcup {u} {z} fc u≤px = subst (λ k → FClosure A f k z ) (sf1=sf0 u≤px) fc
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
901 fcpu : {u z : Ordinal } → FClosure A f (supf0 u) z → u o≤ px → FClosure A f (supf1 u) z
903
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 902
diff changeset
902 fcpu {u} {z} fc u≤px = subst (λ k → FClosure A f k z ) (sym (sf1=sf0 u≤px)) fc
967
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 966
diff changeset
903
999
3ffbdd53d1ea fcs<sup requires <-monotonicity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 998
diff changeset
904 -- this is a kind of maximality, so we cannot prove this without <-monotonicity
3ffbdd53d1ea fcs<sup requires <-monotonicity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 998
diff changeset
905 --
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
906 cfcs : {a b w : Ordinal }
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
907 → a o< b → b o≤ x → supf1 a o< b → FClosure A f (supf1 a) w → odef (UnionCF A f ay supf1 b) w
1090
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
908 cfcs {a} {b} {w} a<b b≤x sa<b fc with x<y∨y≤x (supf0 a) px
1012
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1011
diff changeset
909 ... | case2 px≤sa = z50 where
1023
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1022
diff changeset
910 a<x : a o< x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1022
diff changeset
911 a<x = ordtrans<-≤ a<b b≤x
1012
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1011
diff changeset
912 a≤px : a o≤ px
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1011
diff changeset
913 a≤px = subst (λ k → a o< k) (sym (Oprev.oprev=x op)) (ordtrans<-≤ a<b b≤x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1011
diff changeset
914 -- supf0 a ≡ px we cannot use previous cfcs, it is in the chain because
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1011
diff changeset
915 -- supf0 a ≡ supf0 (supf0 a) ≡ supf0 px o< x
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
916 z50 : odef (UnionCF A f ay supf1 b) w
1012
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1011
diff changeset
917 z50 with osuc-≡< px≤sa
1034
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
918 ... | case1 px=sa = ⟪ A∋fc {A} _ f mf fc , cp ⟫ where
1023
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1022
diff changeset
919 sa≤px : supf0 a o≤ px
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1022
diff changeset
920 sa≤px = subst₂ (λ j k → j o< k) px=sa (sym (Oprev.oprev=x op)) px<x
1026
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1025
diff changeset
921 spx=sa : supf0 px ≡ supf0 a
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1025
diff changeset
922 spx=sa = begin
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
923 supf0 px ≡⟨ cong supf0 px=sa ⟩
1039
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1038
diff changeset
924 supf0 (supf0 a) ≡⟨ ZChain.supf-idem zc a≤px sa≤px ⟩
1026
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1025
diff changeset
925 supf0 a ∎ where open ≡-Reasoning
1020
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1019
diff changeset
926 z51 : supf0 px o< b
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
927 z51 = subst (λ k → k o< b ) (sym ( begin supf0 px ≡⟨ spx=sa ⟩
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
928 supf0 a ≡⟨ sym (sf1=sf0 a≤px) ⟩
1025
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
929 supf1 a ∎ )) sa<b where open ≡-Reasoning
1020
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1019
diff changeset
930 z52 : supf1 a ≡ supf1 (supf0 px)
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
931 z52 = begin supf1 a ≡⟨ sf1=sf0 a≤px ⟩
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
932 supf0 a ≡⟨ sym (ZChain.supf-idem zc a≤px sa≤px ) ⟩
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
933 supf0 (supf0 a) ≡⟨ sym (sf1=sf0 sa≤px) ⟩
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
934 supf1 (supf0 a) ≡⟨ cong supf1 (sym spx=sa) ⟩
1025
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
935 supf1 (supf0 px) ∎ where open ≡-Reasoning
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
936 z53 : supf1 (supf0 px) ≡ supf0 px
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
937 z53 = begin
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
938 supf1 (supf0 px) ≡⟨ cong supf1 spx=sa ⟩
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
939 supf1 (supf0 a) ≡⟨ sf1=sf0 sa≤px ⟩
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
940 supf0 (supf0 a) ≡⟨ sym ( cong supf0 px=sa ) ⟩
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
941 supf0 px ∎ where open ≡-Reasoning
1034
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
942 cp : UChain ay b w
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
943 cp = ch-is-sup (supf0 px) z51 z53 (subst (λ k → FClosure A f k w) z52 fc)
1020
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1019
diff changeset
944 ... | case2 px<sa = ⊥-elim ( ¬p<x<op ⟪ px<sa , subst₂ (λ j k → j o< k ) (sf1=sf0 a≤px) (sym (Oprev.oprev=x op)) z53 ⟫ ) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1019
diff changeset
945 z53 : supf1 a o< x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1019
diff changeset
946 z53 = ordtrans<-≤ sa<b b≤x
1012
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1011
diff changeset
947 ... | case1 sa<px with trio< a px
996
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 995
diff changeset
948 ... | tri< a<px ¬b ¬c = z50 where
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
949 z50 : odef (UnionCF A f ay supf1 b) w
997
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 996
diff changeset
950 z50 with osuc-≡< b≤x
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
951 ... | case2 lt with ZChain.cfcs zc a<b (subst (λ k → b o< k) (sym (Oprev.oprev=x op)) lt) sa<b fc
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
952 ... | ⟪ az , ch-init fc ⟫ = ⟪ az , ch-init fc ⟫
1034
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
953 ... | ⟪ az , ch-is-sup u u<b su=u fc ⟫ = ⟪ az , ch-is-sup u u<b (trans (sym (sf=eq u<x)) su=u) (fcpu fc u≤px ) ⟫ where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
954 u≤px : u o≤ px
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
955 u≤px = subst (λ k → u o< k) (sym (Oprev.oprev=x op)) (ordtrans<-≤ u<b b≤x )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
956 u<x : u o< x
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
957 u<x = ordtrans<-≤ u<b b≤x
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
958 z50 | case1 eq with ZChain.cfcs zc a<px o≤-refl sa<px fc
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
959 ... | ⟪ az , ch-init fc ⟫ = ⟪ az , ch-init fc ⟫
1034
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
960 ... | ⟪ az , ch-is-sup u u<px su=u fc ⟫ = ⟪ az , ch-is-sup u u<b (trans (sym (sf=eq u<x)) su=u) (fcpu fc (o<→≤ u<px)) ⟫ where -- u o< px → u o< b ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
961 u<b : u o< b
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
962 u<b = subst (λ k → u o< k ) (trans (Oprev.oprev=x op) (sym eq) ) (ordtrans u<px <-osuc )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
963 u<x : u o< x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
964 u<x = subst (λ k → u o< k ) (Oprev.oprev=x op) ( ordtrans u<px <-osuc )
1000
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 999
diff changeset
965 ... | tri≈ ¬a a=px ¬c = csupf1 where
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
966 -- a ≡ px , b ≡ x, sp o≤ x
995
04f4baee7b68 UChain is now u o< x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 994
diff changeset
967 px<b : px o< b
04f4baee7b68 UChain is now u o< x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 994
diff changeset
968 px<b = subst₂ (λ j k → j o< k) a=px refl a<b
04f4baee7b68 UChain is now u o< x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 994
diff changeset
969 b=x : b ≡ x
04f4baee7b68 UChain is now u o< x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 994
diff changeset
970 b=x with trio< b x
996
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 995
diff changeset
971 ... | tri< a ¬b ¬c = ⊥-elim ( ¬p<x<op ⟪ px<b , subst (λ k → b o< k ) (sym (Oprev.oprev=x op)) a ⟫ ) -- px o< b o< x
995
04f4baee7b68 UChain is now u o< x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 994
diff changeset
972 ... | tri≈ ¬a b ¬c = b
996
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 995
diff changeset
973 ... | tri> ¬a ¬b c = ⊥-elim ( o≤> b≤x c ) -- x o< b
997
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 996
diff changeset
974 z51 : FClosure A f (supf1 px) w
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 996
diff changeset
975 z51 = subst (λ k → FClosure A f k w) (sym (trans (cong supf1 (sym a=px)) (sf1=sf0 (o≤-refl0 a=px) ))) fc
1001
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1000
diff changeset
976 z53 : odef A w
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1000
diff changeset
977 z53 = A∋fc {A} _ f mf fc
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
978 csupf1 : odef (UnionCF A f ay supf1 b) w
1090
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
979 csupf1 with x<y∨y≤x px (supf0 px)
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
980 ... | case2 spx≤px = ⟪ z53 , ch-is-sup (supf0 px) z54 z52 fc1 ⟫ where
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
981 z54 : supf0 px o< b
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
982 z54 = subst (λ k → supf0 px o< k ) (trans (Oprev.oprev=x op) (sym b=x) ) spx≤px
1003
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1002
diff changeset
983 z52 : supf1 (supf0 px) ≡ supf0 px
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
984 z52 = trans (sf1=sf0 spx≤px ) ( ZChain.supf-idem zc o≤-refl spx≤px )
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
985 fc1 : FClosure A f (supf1 (supf0 px)) w
1004
5c62c97adac9 first cfcs done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1003
diff changeset
986 fc1 = subst (λ k → FClosure A f k w ) (trans (cong supf0 a=px) (sym z52) ) fc
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
987 ... | case1 px<spx = ⊥-elim (¬p<x<op ⟪ px<spx , z54 ⟫ ) where -- supf1 px o≤ spuf1 x → supf1 px ≡ x o< x
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
988 z54 : supf0 px o≤ px
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
989 z54 = subst₂ (λ j k → supf0 j o< k ) a=px (trans b=x (sym (Oprev.oprev=x op))) sa<b
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
990
996
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 995
diff changeset
991 ... | tri> ¬a ¬b c = ⊥-elim ( ¬p<x<op ⟪ c , subst (λ k → a o< k ) (sym (Oprev.oprev=x op)) ( ordtrans<-≤ a<b b≤x) ⟫ ) -- px o< a o< b o≤ x
994
a15f1cddf4c6 u ≤ x again?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 993
diff changeset
992
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
993 zc11 : {z : Ordinal} → odef (UnionCF A f ay supf1 x) z → odef pchainpx z
1034
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
994 zc11 {z} ⟪ az , ch-init fc ⟫ = case1 ⟪ az , ch-init fc ⟫
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
995 zc11 {z} ⟪ az , ch-is-sup u u<x su=u fc ⟫ = zc21 fc where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
996 zc21 : {z1 : Ordinal } → FClosure A f (supf1 u) z1 → odef pchainpx z1
903
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 902
diff changeset
997 zc21 {z1} (fsuc z2 fc ) with zc21 fc
1034
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
998 ... | case1 ⟪ ua1 , ch-init fc₁ ⟫ = case1 ⟪ proj2 ( mf _ ua1) , ch-init (fsuc _ fc₁) ⟫
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1033
diff changeset
999 ... | case1 ⟪ ua1 , ch-is-sup u u<x su=u fc₁ ⟫ = case1 ⟪ proj2 ( mf _ ua1) , ch-is-sup u u<x su=u (fsuc _ fc₁) ⟫
1043
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1042
diff changeset
1000 ... | case2 fc = case2 ⟪ fsuc _ (proj1 fc) , proj2 fc ⟫
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
1001 zc21 (init asp refl ) with trio< (supf0 u) (supf0 px)
1036
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1002 ... | tri< a ¬b ¬c = case1 ⟪ asp , ch-is-sup u u<px (trans (sym (sf1=sf0 (o<→≤ u<px))) su=u )(init asp0 (sym (sf1=sf0 (o<→≤ u<px))) ) ⟫ where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1003 u<px : u o< px
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1004 u<px = ZChain.supf-inject zc a
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1005 asp0 : odef A (supf0 u)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1006 asp0 = ZChain.asupf zc
1043
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1042
diff changeset
1007 ... | tri≈ ¬a b ¬c = case2 ⟪ (init (subst (λ k → odef A k) b (ZChain.asupf zc) )
1044
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1043
diff changeset
1008 (sym (trans (sf1=sf0 (zc-b<x _ u<x)) b ))) , spx<x ⟫ where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1043
diff changeset
1009 spx<x : supf0 px o< x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1043
diff changeset
1010 spx<x = osucprev ( begin
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1043
diff changeset
1011 osuc (supf0 px) ≡⟨ cong osuc (sym b) ⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1043
diff changeset
1012 osuc (supf0 u) ≡⟨ cong osuc (sym (sf1=sf0 (zc-b<x _ u<x) )) ⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1043
diff changeset
1013 osuc (supf1 u) ≡⟨ cong osuc su=u ⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1043
diff changeset
1014 osuc u ≤⟨ osucc u<x ⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1043
diff changeset
1015 x ∎ ) where open o≤-Reasoning O
1036
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1016 ... | tri> ¬a ¬b c = ⊥-elim ( ¬p<x<op ⟪ ZChain.supf-inject zc c , subst (λ k → u o< k ) (sym (Oprev.oprev=x op)) u<x ⟫ )
967
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 966
diff changeset
1017
1035
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1034
diff changeset
1018 is-minsup : {z : Ordinal} → z o≤ x → IsMinSUP A (UnionCF A f ay supf1 z) (supf1 z)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1034
diff changeset
1019 is-minsup {z} z≤x with osuc-≡< z≤x
1036
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1020 ... | case1 z=x = record { as = zc22 ; x≤sup = z23 ; minsup = z24 } where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1021 px<z : px o< z
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
1022 px<z = subst (λ k → px o< k) (sym z=x) px<x
1036
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1023 zc22 : odef A (supf1 z)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1024 zc22 = subst (λ k → odef A k ) (sym (sf1=sp1 px<z )) ( MinSUP.as sup1 )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1025 z23 : {w : Ordinal } → odef ( UnionCF A f ay supf1 z ) w → w ≤ supf1 z
1045
022d2ef3f20b is-minsup in px case done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1044
diff changeset
1026 z23 {w} uz = subst (λ k → w ≤ k ) (sym (sf1=sp1 px<z)) ( MinSUP.x≤sup sup1 (
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1027 zc11 (subst (λ k → odef (UnionCF A f ay supf1 k) w) z=x uz )))
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
1028 z24 : {s : Ordinal } → odef A s → ( {w : Ordinal } → odef ( UnionCF A f ay supf1 z ) w → w ≤ s )
1036
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1029 → supf1 z o≤ s
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1030 z24 {s} as sup = subst (λ k → k o≤ s ) (sym (sf1=sp1 px<z)) ( MinSUP.minsup sup1 as z25 ) where
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
1031 z25 : {w : Ordinal } → odef pchainpx w → w ≤ s
1043
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1042
diff changeset
1032 z25 {w} (case2 fc) = sup ⟪ A∋fc _ f mf (proj1 fc) , ch-is-sup (supf0 px) z28 z27 fc1 ⟫ where
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
1033 -- z=x , supf0 px o< x
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
1034 z28 : supf0 px o< z -- supf0 px ≡ supf1 px o≤ supf1 x ≡ sp1 o≤ x ≡ z
1044
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1043
diff changeset
1035 z28 = subst (λ k → supf0 px o< k) (sym z=x) (proj2 fc)
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
1036 z29 : supf0 px o≤ px
1044
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1043
diff changeset
1037 z29 = zc-b<x _ (proj2 fc)
1040
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1039
diff changeset
1038 z27 : supf1 (supf0 px) ≡ supf0 px
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1039
diff changeset
1039 z27 = trans (sf1=sf0 z29) ( ZChain.supf-idem zc o≤-refl z29 )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1039
diff changeset
1040 fc1 : FClosure A f (supf1 (supf0 px)) w
1043
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1042
diff changeset
1041 fc1 = subst (λ k → FClosure A f k w) (sym z27) (proj1 fc)
1036
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1042 z25 {w} (case1 ⟪ ua , ch-init fc ⟫) = sup ⟪ ua , ch-init fc ⟫
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1043 z25 {w} (case1 ⟪ ua , ch-is-sup u u<x su=u fc ⟫) = sup ⟪ ua , ch-is-sup u u<z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1044 (trans (sf1=sf0 u≤px) su=u) (fcpu fc u≤px) ⟫ where
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
1045 u≤px : u o< osuc px
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
1046 u≤px = ordtrans u<x <-osuc
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
1047 u<z : u o< z
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
1048 u<z = ordtrans u<x (subst (λ k → px o< k ) (sym z=x) px<x )
1036
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1049 ... | case2 z<x = record { as = zc22 ; x≤sup = z23 ; minsup = z24 } where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1050 z≤px = zc-b<x _ z<x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1051 m = ZChain.is-minsup zc z≤px
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1052 zc22 : odef A (supf1 z)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1053 zc22 = subst (λ k → odef A k ) (sym (sf1=sf0 z≤px)) ( IsMinSUP.as m )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1054 z23 : {w : Ordinal } → odef ( UnionCF A f ay supf1 z ) w → w ≤ supf1 z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1055 z23 {w} ⟪ ua , ch-init fc ⟫ = subst (λ k → w ≤ k ) (sym (sf1=sf0 z≤px)) ( ZChain.fcy<sup zc z≤px fc )
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
1056 z23 {w} ⟪ ua , ch-is-sup u u<x su=u fc ⟫ = subst (λ k → w ≤ k ) (sym (sf1=sf0 z≤px))
1036
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1057 (IsMinSUP.x≤sup m ⟪ ua , ch-is-sup u u<x (trans (sym (sf1=sf0 u≤px )) su=u) (fcup fc u≤px ) ⟫ ) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1058 u≤px : u o≤ px
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1059 u≤px = ordtrans u<x z≤px
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1060 z24 : {s : Ordinal } → odef A s → ( {w : Ordinal } → odef ( UnionCF A f ay supf1 z ) w → w ≤ s )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1061 → supf1 z o≤ s
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1062 z24 {s} as sup = subst (λ k → k o≤ s ) (sym (sf1=sf0 z≤px)) ( IsMinSUP.minsup m as z25 ) where
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
1063 z25 : {w : Ordinal } → odef ( UnionCF A f ay supf0 z ) w → w ≤ s
1036
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1064 z25 {w} ⟪ ua , ch-init fc ⟫ = sup ⟪ ua , ch-init fc ⟫
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
1065 z25 {w} ⟪ ua , ch-is-sup u u<x su=u fc ⟫ = sup ⟪ ua , ch-is-sup u u<x
1036
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1066 (trans (sf1=sf0 u≤px) su=u) (fcpu fc u≤px) ⟫ where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1067 u≤px : u o≤ px
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1035
diff changeset
1068 u≤px = ordtrans u<x z≤px
885
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 884
diff changeset
1069
1059
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1058
diff changeset
1070 zo≤sz : {z : Ordinal} → z o≤ x → z o≤ supf1 z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1058
diff changeset
1071 zo≤sz {z} z≤x with osuc-≡< z≤x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1058
diff changeset
1072 ... | case2 z<x = subst (λ k → z o≤ k) (sym (sf1=sf0 (zc-b<x _ z<x ))) (ZChain.zo≤sz zc (zc-b<x _ z<x ))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1058
diff changeset
1073 ... | case1 refl with osuc-≡< (supf1-mono (o<→≤ (px<x))) -- px o≤ supf1 px o≤ supf1 x ≡ sp1 → x o≤ sp1
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1074 ... | case2 lt = begin
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1075 x ≡⟨ sym (Oprev.oprev=x op) ⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1076 osuc px ≤⟨ osucc (ZChain.zo≤sz zc o≤-refl) ⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1077 osuc (supf0 px) ≡⟨ sym (cong osuc (sf1=sf0 o≤-refl )) ⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1078 osuc (supf1 px) ≤⟨ osucc lt ⟩
1059
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1058
diff changeset
1079 supf1 x ∎ where open o≤-Reasoning O
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1058
diff changeset
1080 ... | case1 spx=sx with osuc-≡< ( ZChain.zo≤sz zc o≤-refl )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1058
diff changeset
1081 ... | case2 lt = begin
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1082 x ≡⟨ sym (Oprev.oprev=x op) ⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1083 osuc px ≤⟨ osucc lt ⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1084 supf0 px ≡⟨ sym (sf1=sf0 o≤-refl) ⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1085 supf1 px ≤⟨ supf1-mono (o<→≤ px<x) ⟩
1059
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1058
diff changeset
1086 supf1 x ∎ where open o≤-Reasoning O
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1058
diff changeset
1087 ... | case1 px=spx = ⊥-elim ( <<-irr zc40 (proj1 ( mf< (supf0 px) (ZChain.asupf zc))) ) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1058
diff changeset
1088 zc37 : supf0 px ≡ px
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1058
diff changeset
1089 zc37 = sym px=spx
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1058
diff changeset
1090 zc39 : supf0 px ≡ sp1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1058
diff changeset
1091 zc39 = begin
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1058
diff changeset
1092 supf0 px ≡⟨ sym (sf1=sf0 o≤-refl) ⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1058
diff changeset
1093 supf1 px ≡⟨ spx=sx ⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1058
diff changeset
1094 supf1 x ≡⟨ sf1=sp1 px<x ⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1058
diff changeset
1095 sp1 ∎ where open ≡-Reasoning
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1096 zc40 : f (supf0 px) ≤ supf0 px
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1097 zc40 = subst (λ k → f (supf0 px) ≤ k ) (sym zc39)
1059
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1058
diff changeset
1098 ( MinSUP.x≤sup sup1 (case2 ⟪ fsuc _ (init (ZChain.asupf zc) refl) , subst (λ k → k o< x) (sym zc37) px<x ⟫ ))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1058
diff changeset
1099
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
1100 ... | no lim with trio< x o∅
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
1101 ... | tri< a ¬b ¬c = ⊥-elim ( ¬x<0 a )
1090
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
1102 ... | tri≈ ¬a x=0 ¬c = record { supf = λ _ → MinSUP.sup (ysup f mf ay) ; asupf = MinSUP.as (ysup f mf ay) ; supf-mono = λ _ → o≤-refl
1089
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1103 ; zo≤sz = zo≤sz ; is-minsup = is-minsup ; cfcs = λ a<b b≤0 → ⊥-elim ( ¬x<0 (subst (λ k → _ o< k ) x=0 (ordtrans<-≤ a<b b≤0))) } where
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1104
1081
7089a047e49f strange bug of agda
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1080
diff changeset
1105 mf : ≤-monotonic-f A f
7089a047e49f strange bug of agda
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1080
diff changeset
1106 mf x ax = ⟪ case2 mf00 , proj2 (mf< x ax ) ⟫ where
7089a047e49f strange bug of agda
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1080
diff changeset
1107 mf00 : * x < * (f x)
7089a047e49f strange bug of agda
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1080
diff changeset
1108 mf00 = proj1 ( mf< x ax )
7089a047e49f strange bug of agda
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1080
diff changeset
1109 ym = MinSUP.sup (ysup f mf ay)
1083
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1082
diff changeset
1110
1081
7089a047e49f strange bug of agda
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1080
diff changeset
1111 zo≤sz : {z : Ordinal} → z o≤ x → z o≤ MinSUP.sup (ysup f mf ay)
7089a047e49f strange bug of agda
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1080
diff changeset
1112 zo≤sz {z} z≤x with osuc-≡< z≤x
1083
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1082
diff changeset
1113 ... | case1 refl = subst (λ k → k o≤ _) (sym x=0) o∅≤z
1081
7089a047e49f strange bug of agda
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1080
diff changeset
1114 ... | case2 lt = ⊥-elim ( ¬x<0 (subst (λ k → z o< k ) x=0 lt ) )
1083
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1082
diff changeset
1115
1081
7089a047e49f strange bug of agda
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1080
diff changeset
1116 is-minsup : {z : Ordinal} → z o≤ x →
7089a047e49f strange bug of agda
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1080
diff changeset
1117 IsMinSUP A (UnionCF A f ay (λ _ → MinSUP.sup (ysup f mf ay)) z) (MinSUP.sup (ysup f mf ay))
7089a047e49f strange bug of agda
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1080
diff changeset
1118 is-minsup {z} z≤x with osuc-≡< z≤x
7089a047e49f strange bug of agda
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1080
diff changeset
1119 ... | case1 refl = record { as = MinSUP.as (ysup f mf ay) ; x≤sup = λ {w} uw → is00 uw ; minsup = λ {s} as sup → is01 as sup } where
7089a047e49f strange bug of agda
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1080
diff changeset
1120 is00 : {w : Ordinal } → odef (UnionCF A f ay (λ _ → MinSUP.sup (ysup f mf ay)) x ) w → w ≤ MinSUP.sup (ysup f mf ay)
7089a047e49f strange bug of agda
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1080
diff changeset
1121 is00 {w} ⟪ aw , ch-init fc ⟫ = MinSUP.x≤sup (ysup f mf ay) fc
7089a047e49f strange bug of agda
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1080
diff changeset
1122 is00 {w} ⟪ aw , ch-is-sup u u<z su=u fc ⟫ = ⊥-elim (¬x<0 (subst (λ k → u o< k ) x=0 u<z ))
7089a047e49f strange bug of agda
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1080
diff changeset
1123 is01 : { s : Ordinal } → odef A s → ( {w : Ordinal } → odef (UnionCF A f ay (λ _ → MinSUP.sup (ysup f mf ay)) x ) w → w ≤ s )
7089a047e49f strange bug of agda
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1080
diff changeset
1124 → ym o≤ s
7089a047e49f strange bug of agda
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1080
diff changeset
1125 is01 {s} as sup = MinSUP.minsup (ysup f mf ay) as is02 where
7089a047e49f strange bug of agda
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1080
diff changeset
1126 is02 : {w : Ordinal } → odef (uchain f mf ay) w → (w ≡ s) ∨ (w << s)
7089a047e49f strange bug of agda
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1080
diff changeset
1127 is02 fc = sup ⟪ A∋fc _ f mf fc , ch-init fc ⟫
7089a047e49f strange bug of agda
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1080
diff changeset
1128 ... | case2 lt = ⊥-elim ( ¬x<0 (subst (λ k → z o< k ) x=0 lt ) )
1089
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1129
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1130 ... | tri> ¬a ¬b 0<x = zc400 usup ssup where
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1131
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1132 mf : ≤-monotonic-f A f
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1133 mf x ax = ⟪ case2 mf00 , proj2 (mf< x ax ) ⟫ where
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1134 mf00 : * x < * (f x)
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1135 mf00 = proj1 ( mf< x ax )
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1136
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1137 pzc : {z : Ordinal} → z o< x → ZChain A f mf< ay z
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1138 pzc {z} z<x = prev z z<x
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1139
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1140 ysp = MinSUP.sup (ysup f mf ay)
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1141
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1142 supfz : {z : Ordinal } → z o< x → Ordinal
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1143 supfz {z} z<x = ZChain.supf (pzc (ob<x lim z<x)) z
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1144
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1145 pchainU : HOD
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1146 pchainU = UnionIC A f ay supfz
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1147
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1148 zeq : {xa xb z : Ordinal }
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1149 → (xa<x : xa o< x) → (xb<x : xb o< x) → xa o≤ xb → z o≤ xa
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1150 → ZChain.supf (pzc xa<x) z ≡ ZChain.supf (pzc xb<x) z
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1151 zeq {xa} {xb} {z} xa<x xb<x xa≤xb z≤xa = supf-unique A f mf< ay xa≤xb
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1152 (pzc xa<x) (pzc xb<x) z≤xa
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1153
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1154 iceq : {ix iy : Ordinal } → ix ≡ iy → {i<x : ix o< x } {i<y : iy o< x } → supfz i<x ≡ supfz i<y
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1155 iceq refl = cong supfz o<-irr
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1156
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1157 IChain-i : {z : Ordinal } → IChain ay supfz z → Ordinal
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1158 IChain-i (ic-init fc) = o∅
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1159 IChain-i (ic-isup ia ia<x sa<x fca) = ia
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1160
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1161 pic<x : {z : Ordinal } → (ic : IChain ay supfz z ) → osuc (IChain-i ic) o< x
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1162 pic<x {z} (ic-init fc) = ob<x lim 0<x -- 0<x ∧ lim x → osuc o∅ o< x
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1163 pic<x {z} (ic-isup ia ia<x sa<x fca) = ob<x lim ia<x
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1164
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1165 pchainU⊆chain : {z : Ordinal } → (pz : odef pchainU z) → odef (ZChain.chain (pzc (pic<x (proj2 pz)))) z
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1166 pchainU⊆chain {z} ⟪ aw , ic-init fc ⟫ = ⟪ aw , ch-init fc ⟫
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1167 pchainU⊆chain {z} ⟪ aw , (ic-isup ia ia<x sa<x fca) ⟫ = ZChain.cfcs (pzc (ob<x lim ia<x) ) <-osuc o≤-refl uz03 fca where
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1168 uz02 : FClosure A f (ZChain.supf (pzc (ob<x lim ia<x)) ia ) z
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1169 uz02 = fca
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1170 uz03 : ZChain.supf (pzc (ob<x lim ia<x)) ia o≤ ia
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1171 uz03 = sa<x
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1172
1091
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
1173 chain⊆pchainU : {z w : Ordinal } → (z<x : z o< x) → odef (UnionCF A f ay (ZChain.supf (pzc (ob<x lim z<x))) z) w → odef pchainU w
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
1174 chain⊆pchainU {z} {w} z<x ⟪ aw , ch-init fc ⟫ = ⟪ aw , ic-init fc ⟫
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
1175 chain⊆pchainU {z} {w} z<x ⟪ aw , ch-is-sup u u<oz su=u fc ⟫
1089
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1176 = ⟪ aw , ic-isup u u<x (o≤-refl0 su≡u) (subst (λ k → FClosure A f k w ) su=su fc) ⟫ where
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1177 u<x : u o< x
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1178 u<x = ordtrans u<oz z<x
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1179 su=su : ZChain.supf (pzc (ob<x lim z<x)) u ≡ supfz u<x
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1180 su=su = sym ( zeq _ _ (o<→≤ (osucc u<oz)) (o<→≤ <-osuc) )
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1181 su≡u : supfz u<x ≡ u
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1182 su≡u = begin
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1183 ZChain.supf (pzc (ob<x lim u<x )) u ≡⟨ sym su=su ⟩
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1184 ZChain.supf (pzc (ob<x lim z<x)) u ≡⟨ su=u ⟩
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1185 u ∎ where open ≡-Reasoning
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1186
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1187 IC⊆ : {a b : Ordinal } (ia : IChain ay supfz a ) (ib : IChain ay supfz b )
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1188 → IChain-i ia o< IChain-i ib → odef (ZChain.chain (pzc (pic<x ib))) a
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1189 IC⊆ {a} {b} (ic-init fc ) ib ia<ib = ⟪ A∋fc _ f mf fc , ch-init fc ⟫
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1190 IC⊆ {a} {b} (ic-isup i i<x sa<x fc ) (ic-init fcb ) ia<ib = ⊥-elim ( ¬x<0 ia<ib )
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1191 IC⊆ {a} {b} (ic-isup i i<x sa<x fc ) (ic-isup j j<x sb<x fcb ) ia<ib
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1192 = ZChain.cfcs (pzc (ob<x lim j<x) ) (o<→≤ ia<ib) o≤-refl (OrdTrans (ZChain.supf-mono (pzc (ob<x lim j<x)) (o<→≤ ia<ib)) sb<x)
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1193 (subst (λ k → FClosure A f k a) (zeq _ _ (osucc (o<→≤ ia<ib)) (o<→≤ <-osuc)) fc )
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1194
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1195 ptotalU : IsTotalOrderSet pchainU
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1196 ptotalU {a} {b} ia ib with trio< (IChain-i (proj2 ia)) (IChain-i (proj2 ib))
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1197 ... | tri< ia<ib ¬b ¬c = ZChain.f-total (pzc (pic<x (proj2 ib))) (IC⊆ (proj2 ia) (proj2 ib) ia<ib) (pchainU⊆chain ib)
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1198 ... | tri≈ ¬a ia=ib ¬c = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso ( pcmp (proj2 ia) (proj2 ib) ia=ib ) where
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1199 pcmp : (ia : IChain ay supfz (& a)) → (ib : IChain ay supfz (& b)) → IChain-i ia ≡ IChain-i ib
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1200 → Tri (* (& a) < * (& b)) (* (& a) ≡ * (& b)) (* (& b) < * (& a) )
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1201 pcmp (ic-init fca) (ic-init fcb) eq = fcn-cmp _ f mf fca fcb
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1202 pcmp (ic-init fca) (ic-isup i i<x s<x fcb) eq with ZChain.fcy<sup (pzc i<x) o≤-refl fca
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1203 ... | case1 eq1 = ct22 where
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1204 ct22 : Tri (* (& a) < * (& b)) (* (& a) ≡ * (& b)) (* (& b) < * (& a) )
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1205 ct22 with subst (λ k → k ≤ & b) (sym (zeq _ _ (o<→≤ <-osuc) o≤-refl )) (s≤fc _ f mf fcb )
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1206 ... | case1 eq2 = tri≈ (λ lt → ⊥-elim (<-irr (case1 (sym ct00)) lt)) ct00 (λ lt → ⊥-elim (<-irr (case1 ct00) lt)) where
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1207 ct00 : * (& a) ≡ * (& b)
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1208 ct00 = cong (*) (trans eq1 eq2)
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1209 ... | case2 lt = tri< fc11 (λ eq → <-irr (case1 (sym eq)) fc11) (λ lt → <-irr (case2 fc11) lt) where
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1210 fc11 : * (& a) < * (& b)
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1211 fc11 = subst (λ k → k < * (& b) ) (cong (*) (sym eq1)) lt
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1212 ... | case2 lt = tri< fc11 (λ eq → <-irr (case1 (sym eq)) fc11) (λ lt → <-irr (case2 fc11) lt) where
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1213 fc11 : * (& a) < * (& b)
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1214 fc11 = ftrans<-≤ lt (subst (λ k → k ≤ & b) (sym (zeq _ _ (o<→≤ <-osuc) o≤-refl )) (s≤fc _ f mf fcb ) )
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1215 pcmp (ic-isup i i<x s<x fca) (ic-init fcb) eq with ZChain.fcy<sup (pzc i<x) o≤-refl fcb
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1216 ... | case1 eq1 = ct22 where
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1217 ct22 : Tri (* (& a) < * (& b)) (* (& a) ≡ * (& b)) (* (& b) < * (& a) )
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1218 ct22 with subst (λ k → k ≤ & a) (sym (zeq _ _ (o<→≤ <-osuc) o≤-refl )) (s≤fc _ f mf fca )
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1219 ... | case1 eq2 = tri≈ (λ lt → ⊥-elim (<-irr (case1 (sym ct00)) lt)) ct00 (λ lt → ⊥-elim (<-irr (case1 ct00) lt)) where
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1220 ct00 : * (& a) ≡ * (& b)
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1221 ct00 = cong (*) (sym (trans eq1 eq2))
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1222 ... | case2 lt = tri> (λ lt → <-irr (case2 fc11) lt) (λ eq → <-irr (case1 eq) fc11) fc11 where
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1223 fc11 : * (& b) < * (& a)
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1224 fc11 = subst (λ k → k < * (& a) ) (cong (*) (sym eq1)) lt
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1225 ... | case2 lt = tri> (λ lt → <-irr (case2 fc12) lt) (λ eq → <-irr (case1 eq) fc12) fc12 where
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1226 fc12 : * (& b) < * (& a)
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1227 fc12 = ftrans<-≤ lt (subst (λ k → k ≤ & a) (sym (zeq _ _ (o<→≤ <-osuc) o≤-refl )) (s≤fc _ f mf fca ) )
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1228 pcmp (ic-isup i i<x s<x fca) (ic-isup i i<y s<y fcb) refl = fcn-cmp _ f mf fca (subst (λ k → FClosure A f k (& b)) pc01 fcb ) where
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1229 pc01 : supfz i<y ≡ supfz i<x
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1230 pc01 = cong supfz o<-irr
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1231 ... | tri> ¬a ¬b ib<ia = ZChain.f-total (pzc (pic<x (proj2 ia))) (pchainU⊆chain ia) (IC⊆ (proj2 ib) (proj2 ia) ib<ia)
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1232
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1233
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1234 usup : MinSUP A pchainU
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1235 usup = minsupP pchainU (λ ic → proj1 ic ) ptotalU
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1236 spu0 = MinSUP.sup usup
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1237
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1238
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1239 pchainS : HOD
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1240 pchainS = record { od = record { def = λ z → (odef A z ∧ IChain ay supfz z )
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1241 ∨ (FClosure A f spu0 z ∧ (spu0 o< x)) } ; odmax = & A ; <odmax = zc00 } where
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1242 zc00 : {z : Ordinal } → (odef A z ∧ IChain ay supfz z ) ∨ (FClosure A f spu0 z ∧ (spu0 o< x) )→ z o< & A
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1243 zc00 {z} (case1 lt) = z07 lt
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1244 zc00 {z} (case2 fc) = z09 ( A∋fc spu0 f mf (proj1 fc) )
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1245
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1246 zc02 : { a b : Ordinal } → odef A a ∧ IChain ay supfz a → FClosure A f spu0 b ∧ ( spu0 o< x) → a ≤ b
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1247 zc02 {a} {b} ca fb = zc05 (proj1 fb) where
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1248 zc05 : {b : Ordinal } → FClosure A f spu0 b → a ≤ b
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1249 zc05 (fsuc b1 fb ) with proj1 ( mf b1 (A∋fc spu0 f mf fb ))
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1250 ... | case1 eq = subst (λ k → a ≤ k ) eq (zc05 fb)
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1251 ... | case2 lt = ≤-ftrans (zc05 fb) (case2 lt)
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1252 zc05 (init b1 refl) = MinSUP.x≤sup usup ca
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1253
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1254 ptotalS : IsTotalOrderSet pchainS
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1255 ptotalS (case1 a) (case1 b) = ptotalU a b
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1256 ptotalS {a0} {b0} (case1 a) (case2 b) with zc02 a b
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1257 ... | case1 eq = tri≈ (<-irr (case1 (sym eq1))) eq1 (<-irr (case1 eq1)) where
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1258 eq1 : a0 ≡ b0
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1259 eq1 = subst₂ (λ j k → j ≡ k ) *iso *iso (cong (*) eq )
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1260 ... | case2 lt = tri< lt1 (λ eq → <-irr (case1 (sym eq)) lt1) (<-irr (case2 lt1)) where
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1261 lt1 : a0 < b0
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1262 lt1 = subst₂ (λ j k → j < k ) *iso *iso lt
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1263 ptotalS {b0} {a0} (case2 b) (case1 a) with zc02 a b
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1264 ... | case1 eq = tri≈ (<-irr (case1 eq1)) (sym eq1) (<-irr (case1 (sym eq1))) where
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1265 eq1 : a0 ≡ b0
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1266 eq1 = subst₂ (λ j k → j ≡ k ) *iso *iso (cong (*) eq )
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1267 ... | case2 lt = tri> (<-irr (case2 lt1)) (λ eq → <-irr (case1 eq) lt1) lt1 where
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1268 lt1 : a0 < b0
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1269 lt1 = subst₂ (λ j k → j < k ) *iso *iso lt
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1270 ptotalS (case2 a) (case2 b) = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso (fcn-cmp spu0 f mf (proj1 a) (proj1 b))
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1271
1091
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
1272 S⊆A : pchainS ⊆ A
1089
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1273 S⊆A (case1 lt) = proj1 lt
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1274 S⊆A (case2 fc) = A∋fc _ f mf (proj1 fc)
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1275
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1276 ssup : MinSUP A pchainS
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1277 ssup = minsupP pchainS S⊆A ptotalS
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1278
b627e3ea7266 try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1084
diff changeset
1279 zc400 : MinSUP A pchainU → MinSUP A pchainS → ZChain A f mf< ay x
1090
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
1280 zc400 usup ssup = record { supf = supf1 ; asupf = asupf ; zo≤sz = zo≤sz ; is-minsup = is-minsup ; cfcs = cfcs ; supf-mono = supf-mono } where
726
b2e2cd12e38f psupf-mono and is-max conflict
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 725
diff changeset
1281
1069
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1068
diff changeset
1282 spu = MinSUP.sup usup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1068
diff changeset
1283 sps = MinSUP.sup ssup
834
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 833
diff changeset
1284
794
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 793
diff changeset
1285 supf1 : Ordinal → Ordinal
835
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 834
diff changeset
1286 supf1 z with trio< z x
1069
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1068
diff changeset
1287 ... | tri< a ¬b ¬c = ZChain.supf (pzc (ob<x lim a)) z -- each sup o< x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1068
diff changeset
1288 ... | tri≈ ¬a b ¬c = spu -- sup of all sup o< x
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1289 ... | tri> ¬a ¬b c = sps -- sup of spu which o< x
1069
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1068
diff changeset
1290 -- if x o< spu, spu is not included in UnionCF x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1068
diff changeset
1291 -- the chain
755
b22327e78b03 u < osuc x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 754
diff changeset
1292
1010
f80d525e6a6b Recursive record IChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1009
diff changeset
1293 pchain : HOD
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
1294 pchain = UnionCF A f ay supf1 x
1010
f80d525e6a6b Recursive record IChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1009
diff changeset
1295
1069
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1068
diff changeset
1296 -- pchain ⊆ pchainU ⊆ pchianS
1009
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1008
diff changeset
1297
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
1298 sf1=sf : {z : Ordinal } → (a : z o< x ) → supf1 z ≡ ZChain.supf (pzc (ob<x lim a)) z
1009
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1008
diff changeset
1299 sf1=sf {z} z<x with trio< z x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1008
diff changeset
1300 ... | tri< a ¬b ¬c = cong ( λ k → ZChain.supf (pzc (ob<x lim k)) z) o<-irr
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1008
diff changeset
1301 ... | tri≈ ¬a b ¬c = ⊥-elim (¬a z<x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1008
diff changeset
1302 ... | tri> ¬a ¬b c = ⊥-elim (¬a z<x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1008
diff changeset
1303
1071
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1070
diff changeset
1304 sf1=spu : {z : Ordinal } → x ≡ z → supf1 z ≡ spu
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1070
diff changeset
1305 sf1=spu {z} eq with trio< z x
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1306 ... | tri< a ¬b ¬c = ⊥-elim (¬b (sym eq))
1009
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1008
diff changeset
1307 ... | tri≈ ¬a b ¬c = refl
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1308 ... | tri> ¬a ¬b c = ⊥-elim (¬b (sym eq))
1009
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1008
diff changeset
1309
1071
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1070
diff changeset
1310 sf1=sps : {z : Ordinal } → (a : x o< z ) → supf1 z ≡ sps
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1070
diff changeset
1311 sf1=sps {z} x<z with trio< z x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1070
diff changeset
1312 ... | tri< a ¬b ¬c = ⊥-elim (o<> x<z a)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1070
diff changeset
1313 ... | tri≈ ¬a b ¬c = ⊥-elim (¬c x<z )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1070
diff changeset
1314 ... | tri> ¬a ¬b c = refl
1069
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1068
diff changeset
1315
1057
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1056
diff changeset
1316 asupf : {z : Ordinal } → odef A (supf1 z)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1056
diff changeset
1317 asupf {z} with trio< z x
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1318 ... | tri< a ¬b ¬c = ZChain.asupf (pzc (ob<x lim a))
1057
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1056
diff changeset
1319 ... | tri≈ ¬a b ¬c = MinSUP.as usup
1069
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1068
diff changeset
1320 ... | tri> ¬a ¬b c = MinSUP.as ssup
1057
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1056
diff changeset
1321
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1056
diff changeset
1322 supf-mono : {z y : Ordinal } → z o≤ y → supf1 z o≤ supf1 y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1056
diff changeset
1323 supf-mono {z} {y} z≤y with trio< y x
1071
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1070
diff changeset
1324 ... | tri< y<x ¬b ¬c = zc01 where
1057
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1056
diff changeset
1325 open o≤-Reasoning O
1071
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1070
diff changeset
1326 zc01 : supf1 z o≤ ZChain.supf (pzc (ob<x lim y<x)) y
1072
4ce084a0dce2 supf-mono done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1071
diff changeset
1327 zc01 = begin
4ce084a0dce2 supf-mono done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1071
diff changeset
1328 supf1 z ≡⟨ sf1=sf (ordtrans≤-< z≤y y<x) ⟩
4ce084a0dce2 supf-mono done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1071
diff changeset
1329 ZChain.supf (pzc (ob<x lim (ordtrans≤-< z≤y y<x))) z ≡⟨ zeq _ _ (osucc z≤y) (o<→≤ <-osuc) ⟩
1070
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1069
diff changeset
1330 ZChain.supf (pzc (ob<x lim y<x)) z ≤⟨ ZChain.supf-mono (pzc (ob<x lim y<x)) z≤y ⟩
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1331 ZChain.supf (pzc (ob<x lim y<x)) y ∎
1072
4ce084a0dce2 supf-mono done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1071
diff changeset
1332 ... | tri≈ ¬a b ¬c = zc01 where -- supf1 z o≤ spu
4ce084a0dce2 supf-mono done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1071
diff changeset
1333 open o≤-Reasoning O
4ce084a0dce2 supf-mono done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1071
diff changeset
1334 zc01 : supf1 z o≤ spu
4ce084a0dce2 supf-mono done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1071
diff changeset
1335 zc01 with osuc-≡< (subst (λ k → z o≤ k) b z≤y)
4ce084a0dce2 supf-mono done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1071
diff changeset
1336 ... | case1 z=x = o≤-refl0 (sf1=spu (sym z=x))
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1337 ... | case2 z<x = subst (λ k → k o≤ spu ) (sym (sf1=sf z<x)) ( IsMinSUP.minsup (ZChain.is-minsup (pzc (ob<x lim z<x)) (o<→≤ <-osuc) )
1091
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
1338 (MinSUP.as usup) (λ uw → MinSUP.x≤sup usup (chain⊆pchainU z<x uw)) )
1072
4ce084a0dce2 supf-mono done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1071
diff changeset
1339 ... | tri> ¬a ¬b c = zc01 where -- supf1 z o≤ sps
4ce084a0dce2 supf-mono done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1071
diff changeset
1340 zc01 : supf1 z o≤ sps
4ce084a0dce2 supf-mono done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1071
diff changeset
1341 zc01 with trio< z x
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1342 ... | tri< z<x ¬b ¬c = IsMinSUP.minsup (ZChain.is-minsup (pzc (ob<x lim z<x)) (o<→≤ <-osuc) )
1091
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
1343 (MinSUP.as ssup) (λ uw → MinSUP.x≤sup ssup (case1 (chain⊆pchainU z<x uw)) )
1072
4ce084a0dce2 supf-mono done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1071
diff changeset
1344 ... | tri≈ ¬a z=x ¬c = MinSUP.minsup usup (MinSUP.as ssup) (λ uw → MinSUP.x≤sup ssup (case1 uw) )
4ce084a0dce2 supf-mono done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1071
diff changeset
1345 ... | tri> ¬a ¬b c = o≤-refl -- (sf1=sps c)
797
3a8493e6cd67 supf contraint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 796
diff changeset
1346
1035
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1034
diff changeset
1347 is-minsup : {z : Ordinal} → z o≤ x → IsMinSUP A (UnionCF A f ay supf1 z) (supf1 z)
1073
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1072
diff changeset
1348 is-minsup {z} z≤x with osuc-≡< z≤x
1075
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1074
diff changeset
1349 ... | case1 z=x = record { as = asupf ; x≤sup = zm00 ; minsup = zm01 } where
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1350 zm00 : {w : Ordinal } → odef (UnionCF A f ay supf1 z) w → w ≤ supf1 z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1351 zm00 {w} ⟪ az , ch-init fc ⟫ = subst (λ k → w ≤ k ) (sym (sf1=spu (sym z=x))) ( MinSUP.x≤sup usup ⟪ az , ic-init fc ⟫ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1352 zm00 {w} ⟪ az , ch-is-sup u u<b su=u fc ⟫ = subst (λ k → w ≤ k ) (sym (sf1=spu (sym z=x)))
1076
7e047b46c3b2 is-minsup done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1075
diff changeset
1353 ( MinSUP.x≤sup usup ⟪ az , ic-isup u u<x (o≤-refl0 zm05) (subst (λ k → FClosure A f k w) (sym zm06) fc) ⟫ ) where
7e047b46c3b2 is-minsup done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1075
diff changeset
1354 u<x : u o< x
7e047b46c3b2 is-minsup done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1075
diff changeset
1355 u<x = subst (λ k → u o< k) z=x u<b
7e047b46c3b2 is-minsup done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1075
diff changeset
1356 zm06 : supfz (subst (λ k → u o< k) z=x u<b) ≡ supf1 u
7e047b46c3b2 is-minsup done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1075
diff changeset
1357 zm06 = trans (zeq _ _ o≤-refl (o<→≤ <-osuc) ) (sym (sf1=sf u<x ))
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1358 zm05 : supfz (subst (λ k → u o< k) z=x u<b) ≡ u
1076
7e047b46c3b2 is-minsup done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1075
diff changeset
1359 zm05 = trans zm06 su=u
1075
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1074
diff changeset
1360 zm01 : { s : Ordinal } → odef A s → ( {x : Ordinal } → odef (UnionCF A f ay supf1 z) x → x ≤ s ) → supf1 z o≤ s
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1361 zm01 {s} as sup = subst (λ k → k o≤ s ) (sym (sf1=spu (sym z=x))) ( MinSUP.minsup usup as zm02 ) where
1075
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1074
diff changeset
1362 zm02 : {w : Ordinal } → odef pchainU w → w ≤ s
1076
7e047b46c3b2 is-minsup done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1075
diff changeset
1363 zm02 {w} uw with pchainU⊆chain uw
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1364 ... | ⟪ az , ch-init fc ⟫ = sup ⟪ az , ch-init fc ⟫
1076
7e047b46c3b2 is-minsup done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1075
diff changeset
1365 ... | ⟪ az , ch-is-sup u1 u<b su=u fc ⟫ = sup ⟪ az , ch-is-sup u1 (ordtrans u<b zm05) (trans zm03 su=u) zm04 ⟫ where
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1366 zm05 : osuc (IChain-i (proj2 uw)) o< z
1076
7e047b46c3b2 is-minsup done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1075
diff changeset
1367 zm05 = subst (λ k → osuc (IChain-i (proj2 uw)) o< k) (sym z=x) ( pic<x (proj2 uw) )
7e047b46c3b2 is-minsup done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1075
diff changeset
1368 u<x : u1 o< x
7e047b46c3b2 is-minsup done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1075
diff changeset
1369 u<x = subst (λ k → u1 o< k) z=x ( ordtrans u<b zm05 )
7e047b46c3b2 is-minsup done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1075
diff changeset
1370 zm03 : supf1 u1 ≡ ZChain.supf (prev (osuc (IChain-i (proj2 uw))) (pic<x (proj2 uw))) u1
7e047b46c3b2 is-minsup done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1075
diff changeset
1371 zm03 = trans (sf1=sf u<x) (zeq _ _ (osucc u<b) (o<→≤ <-osuc) )
7e047b46c3b2 is-minsup done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1075
diff changeset
1372 zm04 : FClosure A f (supf1 u1) w
7e047b46c3b2 is-minsup done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1075
diff changeset
1373 zm04 = subst (λ k → FClosure A f k w) (sym zm03) fc
1073
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1072
diff changeset
1374 ... | case2 z<x = record { as = asupf ; x≤sup = zm00 ; minsup = zm01 } where
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1375 supf0 = ZChain.supf (pzc (ob<x lim z<x))
1073
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1072
diff changeset
1376 msup : IsMinSUP A (UnionCF A f ay supf0 z) (supf0 z)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1072
diff changeset
1377 msup = ZChain.is-minsup (pzc (ob<x lim z<x)) (o<→≤ <-osuc)
1074
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1073
diff changeset
1378 s1=0 : {u : Ordinal } → u o< z → supf1 u ≡ supf0 u
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1379 s1=0 {u} u<z = trans (sf1=sf (ordtrans u<z z<x)) (zeq _ _ (o<→≤ (osucc u<z)) (o<→≤ <-osuc) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1380 zm00 : {w : Ordinal } → odef (UnionCF A f ay supf1 z) w → w ≤ supf1 z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1381 zm00 {w} ⟪ az , ch-init fc ⟫ = subst (λ k → w ≤ k ) (sym (sf1=sf z<x)) ( IsMinSUP.x≤sup msup ⟪ az , ch-init fc ⟫ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1382 zm00 {w} ⟪ az , ch-is-sup u u<b su=u fc ⟫ = subst (λ k → w ≤ k ) (sym (sf1=sf z<x))
1074
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1073
diff changeset
1383 ( IsMinSUP.x≤sup msup ⟪ az , ch-is-sup u u<b (trans (sym (s1=0 u<b)) su=u) (subst (λ k → FClosure A f k w) (s1=0 u<b) fc) ⟫ )
1073
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1072
diff changeset
1384 zm01 : { s : Ordinal } → odef A s → ( {x : Ordinal } → odef (UnionCF A f ay supf1 z) x → x ≤ s ) → supf1 z o≤ s
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1385 zm01 {s} as sup = subst (λ k → k o≤ s ) (sym (sf1=sf z<x)) ( IsMinSUP.minsup msup as zm02 ) where
1073
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1072
diff changeset
1386 zm02 : {w : Ordinal } → odef (UnionCF A f ay supf0 z) w → w ≤ s
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1072
diff changeset
1387 zm02 {w} ⟪ az , ch-init fc ⟫ = sup ⟪ az , ch-init fc ⟫
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1388 zm02 {w} ⟪ az , ch-is-sup u u<b su=u fc ⟫ = sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1389 ⟪ az , ch-is-sup u u<b (trans (s1=0 u<b) su=u) (subst (λ k → FClosure A f k w) (sym (s1=0 u<b)) fc) ⟫
1073
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1072
diff changeset
1390
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1072
diff changeset
1391
1039
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1038
diff changeset
1392 cfcs : {a b w : Ordinal } → a o< b → b o≤ x → supf1 a o< b → FClosure A f (supf1 a) w → odef (UnionCF A f ay supf1 b) w
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1038
diff changeset
1393 cfcs {a} {b} {w} a<b b≤x sa<b fc with osuc-≡< b≤x
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
1394 ... | case1 b=x with trio< a x
1016
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1015
diff changeset
1395 ... | tri< a<x ¬b ¬c = zc40 where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1015
diff changeset
1396 sa = ZChain.supf (pzc (ob<x lim a<x)) a
1020
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1019
diff changeset
1397 m = omax a sa -- x is limit ordinal, so we have sa o< m o< x
1016
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1015
diff changeset
1398 m<x : m o< x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1015
diff changeset
1399 m<x with trio< a sa | inspect (omax a) sa
1020
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1019
diff changeset
1400 ... | tri< a<sa ¬b ¬c | record { eq = eq } = ob<x lim (ordtrans<-≤ sa<b b≤x )
1016
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1015
diff changeset
1401 ... | tri≈ ¬a a=sa ¬c | record { eq = eq } = subst (λ k → k o< x) eq zc41 where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1015
diff changeset
1402 zc41 : omax a sa o< x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1015
diff changeset
1403 zc41 = osucprev ( begin
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1015
diff changeset
1404 osuc ( omax a sa ) ≡⟨ cong (λ k → osuc (omax a k)) (sym a=sa) ⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1015
diff changeset
1405 osuc ( omax a a ) ≡⟨ cong osuc (omxx _) ⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1015
diff changeset
1406 osuc ( osuc a ) ≤⟨ o<→≤ (ob<x lim (ob<x lim a<x)) ⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1015
diff changeset
1407 x ∎ ) where open o≤-Reasoning O
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1015
diff changeset
1408 ... | tri> ¬a ¬b c | record { eq = eq } = ob<x lim a<x
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
1409 sam = ZChain.supf (pzc (ob<x lim m<x)) a
1016
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1015
diff changeset
1410 zc42 : osuc a o≤ osuc m
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1015
diff changeset
1411 zc42 = osucc (o<→≤ ( omax-x _ _ ) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1015
diff changeset
1412 sam<m : sam o< m
1039
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1038
diff changeset
1413 sam<m = subst (λ k → k o< m ) (supf-unique A f mf< ay zc42 (pzc (ob<x lim a<x)) (pzc (ob<x lim m<x)) (o<→≤ <-osuc)) ( omax-y _ _ )
1016
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1015
diff changeset
1414 fcm : FClosure A f (ZChain.supf (pzc (ob<x lim m<x)) a) w
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1015
diff changeset
1415 fcm = subst (λ k → FClosure A f k w ) (zeq (ob<x lim a<x) (ob<x lim m<x) zc42 (o<→≤ <-osuc) ) fc
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
1416 zcm : odef (UnionCF A f ay (ZChain.supf (pzc (ob<x lim m<x))) (osuc (omax a sa))) w
1039
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1038
diff changeset
1417 zcm = ZChain.cfcs (pzc (ob<x lim m<x)) (o<→≤ (omax-x _ _)) o≤-refl (o<→≤ sam<m) fcm
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
1418 zc40 : odef (UnionCF A f ay supf1 b) w
1039
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1038
diff changeset
1419 zc40 with ZChain.cfcs (pzc (ob<x lim m<x)) (o<→≤ (omax-x _ _)) o≤-refl (o<→≤ sam<m) fcm
1058
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1057
diff changeset
1420 ... | ⟪ az , ch-init fc ⟫ = ⟪ az , ch-init fc ⟫
1070
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1069
diff changeset
1421 ... | ⟪ az , ch-is-sup u u<x su=u fc ⟫ = ⟪ az , ch-is-sup u u<b (trans zc45 su=u) zc44 ⟫ where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1069
diff changeset
1422 u<b : u o< b
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1069
diff changeset
1423 u<b = osucprev ( begin
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1069
diff changeset
1424 osuc u ≤⟨ osucc u<x ⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1069
diff changeset
1425 osuc m ≤⟨ osucc m<x ⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1069
diff changeset
1426 x ≡⟨ sym b=x ⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1069
diff changeset
1427 b ∎ ) where open o≤-Reasoning O
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1428 zc45 : supf1 u ≡ ZChain.supf (pzc (ob<x lim m<x)) u
1070
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1069
diff changeset
1429 zc45 = begin
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1069
diff changeset
1430 supf1 u ≡⟨ sf1=sf (subst (λ k → u o< k) b=x u<b ) ⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1069
diff changeset
1431 ZChain.supf (pzc (ob<x lim (subst (λ k → u o< k) b=x u<b ))) u ≡⟨ zeq _ _ (osucc u<x) (o<→≤ <-osuc) ⟩
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1432 ZChain.supf (pzc (ob<x lim m<x)) u ∎ where open ≡-Reasoning
1070
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1069
diff changeset
1433 zc44 : FClosure A f (supf1 u) w
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1069
diff changeset
1434 zc44 = subst (λ k → FClosure A f k w ) (sym zc45) fc
1016
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1015
diff changeset
1435 ... | tri≈ ¬a b ¬c = ⊥-elim ( ¬a (ordtrans<-≤ a<b b≤x))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1015
diff changeset
1436 ... | tri> ¬a ¬b c = ⊥-elim ( ¬a (ordtrans<-≤ a<b b≤x))
1039
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1038
diff changeset
1437 cfcs {a} {b} {w} a<b b≤x sa<b fc | case2 b<x = zc40 where
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
1438 supfb = ZChain.supf (pzc (ob<x lim b<x))
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
1439 sb=sa : {a : Ordinal } → a o< b → supf1 a ≡ ZChain.supf (pzc (ob<x lim b<x)) a
1020
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1019
diff changeset
1440 sb=sa {a} a<b = trans (sf1=sf (ordtrans<-≤ a<b b≤x)) (zeq _ _ (o<→≤ (osucc a<b)) (o<→≤ <-osuc) )
1017
ffdfd8d1303a trying cscf as odef (UnionCF A f mf ay supf z) w
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1016
diff changeset
1441 fcb : FClosure A f (supfb a) w
1042
912ca4abe806 pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1041
diff changeset
1442 fcb = subst (λ k → FClosure A f k w) (sb=sa a<b) fc
1020
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1019
diff changeset
1443 -- supfb a o< b assures it is in Union b
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
1444 zcb : odef (UnionCF A f ay supfb b) w
1039
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1038
diff changeset
1445 zcb = ZChain.cfcs (pzc (ob<x lim b<x)) a<b (o<→≤ <-osuc) (subst (λ k → k o< b) (sb=sa a<b) sa<b) fcb
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
1446 zc40 : odef (UnionCF A f ay supf1 b) w
1016
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1015
diff changeset
1447 zc40 with zcb
1058
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1057
diff changeset
1448 ... | ⟪ az , ch-init fc ⟫ = ⟪ az , ch-init fc ⟫
1070
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1069
diff changeset
1449 ... | ⟪ az , ch-is-sup u u<x su=u fc ⟫ = ⟪ az , ch-is-sup u u<x (trans zc45 su=u) zc44 ⟫ where
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1450 zc45 : supf1 u ≡ ZChain.supf (pzc (ob<x lim b<x)) u
1070
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1069
diff changeset
1451 zc45 = begin
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1069
diff changeset
1452 supf1 u ≡⟨ sf1=sf (ordtrans u<x b<x) ⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1069
diff changeset
1453 ZChain.supf (pzc (ob<x lim (ordtrans u<x b<x) )) u ≡⟨ zeq _ _ (o<→≤ (osucc u<x)) (o<→≤ <-osuc) ⟩
1078
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1077
diff changeset
1454 ZChain.supf (pzc (ob<x lim b<x )) u ∎ where open ≡-Reasoning
1070
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1069
diff changeset
1455 zc44 : FClosure A f (supf1 u) w
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1069
diff changeset
1456 zc44 = subst (λ k → FClosure A f k w ) (sym zc45) fc
1028
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1027
diff changeset
1457
1079
c2546206c891 order done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1078
diff changeset
1458 zo≤sz : {z : Ordinal} → z o≤ x → z o≤ supf1 z
1080
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1079
diff changeset
1459 zo≤sz {z} z≤x with osuc-≡< z≤x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1079
diff changeset
1460 ... | case2 z<x = subst (λ k → z o≤ k) (sym (trans (sf1=sf z<x) (sym (zeq _ _ (o<→≤ <-osuc) o≤-refl)))) ( ZChain.zo≤sz (pzc z<x) o≤-refl )
1083
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1082
diff changeset
1461 ... | case1 refl with x<y∨y≤x (supf1 spu) x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1082
diff changeset
1462 ... | case2 x≤ssp = z40 where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1082
diff changeset
1463 z40 : z o≤ supf1 z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1082
diff changeset
1464 z40 with x<y∨y≤x z spu
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1082
diff changeset
1465 ... | case1 z<spu = o<→≤ ( subst (λ k → z o< k ) (sym (sf1=spu refl)) z<spu )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1082
diff changeset
1466 ... | case2 spu≤z = begin -- x ≡ supf1 spu ≡ spu ≡ supf1 x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1082
diff changeset
1467 x ≤⟨ x≤ssp ⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1082
diff changeset
1468 supf1 spu ≤⟨ supf-mono spu≤z ⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1082
diff changeset
1469 supf1 x ∎ where open o≤-Reasoning O
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1082
diff changeset
1470 ... | case1 ssp<x = subst (λ k → x o≤ k) (sym (sf1=spu refl)) z47 where
1080
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1079
diff changeset
1471 z47 : x o≤ spu
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1079
diff changeset
1472 z47 with x<y∨y≤x spu x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1079
diff changeset
1473 ... | case2 lt = lt
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1079
diff changeset
1474 ... | case1 spu<x = ⊥-elim ( <<-irr (MinSUP.x≤sup usup z48) (proj1 ( mf< spu (MinSUP.as usup)))) where
1083
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1082
diff changeset
1475 z70 : odef (UnionCF A f ay supf1 z) (supf1 spu)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1082
diff changeset
1476 z70 = cfcs spu<x o≤-refl ssp<x (init asupf refl )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1082
diff changeset
1477 z73 : IsSUP A (UnionCF A f ay (ZChain.supf (pzc (ob<x lim spu<x))) spu) spu
1091
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
1478 z73 = record { ax = MinSUP.as usup ; x≤sup = λ uw → MinSUP.x≤sup usup (chain⊆pchainU spu<x uw ) }
1080
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1079
diff changeset
1479 z49 : supfz spu<x ≡ spu
1083
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1082
diff changeset
1480 z49 = begin
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1082
diff changeset
1481 supfz spu<x ≡⟨ ZChain.sup=u (pzc (ob<x lim spu<x)) (MinSUP.as usup) (o<→≤ <-osuc) z73 ⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1082
diff changeset
1482 spu ∎ where open ≡-Reasoning
1082
83d9000bf72f 0< is no good
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1081
diff changeset
1483 z50 : supfz spu<x o≤ spu
83d9000bf72f 0< is no good
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1081
diff changeset
1484 z50 = o≤-refl0 z49
1080
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1079
diff changeset
1485 z48 : odef pchainU (f spu)
1082
83d9000bf72f 0< is no good
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1081
diff changeset
1486 z48 = ⟪ proj2 (mf _ (MinSUP.as usup) ) , ic-isup _ (subst (λ k → k o< x) refl spu<x) z50
1080
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1079
diff changeset
1487 (fsuc _ (init (ZChain.asupf (pzc (ob<x lim spu<x))) z49)) ⟫
1079
c2546206c891 order done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1078
diff changeset
1488
1077
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1076
diff changeset
1489
921
c0cf2b383064 UnionZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 920
diff changeset
1490 ---
c0cf2b383064 UnionZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 920
diff changeset
1491 --- the maximum chain has fix point of any ≤-monotonic function
c0cf2b383064 UnionZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 920
diff changeset
1492 ---
c0cf2b383064 UnionZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 920
diff changeset
1493
1039
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1038
diff changeset
1494 SZ : ( f : Ordinal → Ordinal ) → (mf< : <-monotonic-f A f ) → {y : Ordinal} (ay : odef A y) → (x : Ordinal) → ZChain A f mf< ay x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1038
diff changeset
1495 SZ f mf< {y} ay x = TransFinite {λ z → ZChain A f mf< ay z } (λ x → ind f mf< ay x ) x
921
c0cf2b383064 UnionZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 920
diff changeset
1496
1039
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1038
diff changeset
1497 msp0 : ( f : Ordinal → Ordinal ) → (mf< : <-monotonic-f A f ) {x y : Ordinal} (ay : odef A y)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1038
diff changeset
1498 → (zc : ZChain A f mf< ay x )
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
1499 → MinSUP A (UnionCF A f ay (ZChain.supf zc) x)
1039
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1038
diff changeset
1500 msp0 f mf< {x} ay zc = minsupP (UnionCF A f ay (ZChain.supf zc) x) (ZChain.chain⊆A zc) (ZChain.f-total zc)
922
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 921
diff changeset
1501
1090
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
1502 -- f eventualy stop
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
1503 -- we can prove contradict here, it is here for a historical reason
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
1504 --
1039
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1038
diff changeset
1505 fixpoint : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) (mf< : <-monotonic-f A f ) (zc : ZChain A f mf< as0 (& A) )
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
1506 → (sp1 : MinSUP A (ZChain.chain zc))
959
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 958
diff changeset
1507 → f (MinSUP.sup sp1) ≡ MinSUP.sup sp1
992
4aaecae58da5 ... x < & A ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 991
diff changeset
1508 fixpoint f mf mf< zc sp1 = z14 where
924
a48dc906796c supf usp0 instead of supf (& A) ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 923
diff changeset
1509 chain = ZChain.chain zc
a48dc906796c supf usp0 instead of supf (& A) ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 923
diff changeset
1510 supf = ZChain.supf zc
935
ed711d7be191 mem exhaust fix on fixpoint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 934
diff changeset
1511 sp : Ordinal
959
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 958
diff changeset
1512 sp = MinSUP.sup sp1
935
ed711d7be191 mem exhaust fix on fixpoint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 934
diff changeset
1513 asp : odef A sp
1032
382680c3e9be minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1031
diff changeset
1514 asp = MinSUP.as sp1
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
1515 ay = as0
988
9a85233384f7 is-max and supf b = supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 987
diff changeset
1516 z10 : {a b : Ordinal } → (ca : odef chain a ) → b o< (& A) → (ab : odef A b )
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
1517 → HasPrev A chain f b ∨ IsSUP A (UnionCF A f ay (ZChain.supf zc) b) b
921
c0cf2b383064 UnionZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 920
diff changeset
1518 → * a < * b → odef chain b
993
e11c244d7eac SZ1 done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 992
diff changeset
1519 z10 = ZChain1.is-max (SZ1 f mf mf< as0 zc (& A) o≤-refl )
935
ed711d7be191 mem exhaust fix on fixpoint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 934
diff changeset
1520 z22 : sp o< & A
ed711d7be191 mem exhaust fix on fixpoint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 934
diff changeset
1521 z22 = z09 asp
ed711d7be191 mem exhaust fix on fixpoint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 934
diff changeset
1522 z12 : odef chain sp
ed711d7be191 mem exhaust fix on fixpoint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 934
diff changeset
1523 z12 with o≡? (& s) sp
1058
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1057
diff changeset
1524 ... | yes eq = subst (λ k → odef chain k) eq ( ZChain.chain∋init zc )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1057
diff changeset
1525 ... | no ne = ZChain1.is-max (SZ1 f mf mf< as0 zc (& A) o≤-refl) {& s} {sp} ( ZChain.chain∋init zc ) (z09 asp) asp (case2 z19 ) z13 where
935
ed711d7be191 mem exhaust fix on fixpoint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 934
diff changeset
1526 z13 : * (& s) < * sp
1058
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1057
diff changeset
1527 z13 with MinSUP.x≤sup sp1 ( ZChain.chain∋init zc )
960
b7370c39769e IsMinSUP< is wrong
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 959
diff changeset
1528 ... | case1 eq = ⊥-elim ( ne eq )
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
1529 ... | case2 lt = lt
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
1530 z19 : IsSUP A (UnionCF A f ay (ZChain.supf zc) sp) sp
1069
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1068
diff changeset
1531 z19 = record { ax = asp ; x≤sup = z20 } where
1033
2da8dcbb0825 ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1032
diff changeset
1532 z20 : {y : Ordinal} → odef (UnionCF A f ay (ZChain.supf zc) sp) y → (y ≡ sp) ∨ (y << sp)
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
1533 z20 {y} zy with MinSUP.x≤sup sp1
961
811135ad1904 supf sp = sp ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 960
diff changeset
1534 (subst (λ k → odef chain k ) (sym &iso) (chain-mono f mf as0 supf (ZChain.supf-mono zc) (o<→≤ z22) zy ))
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
1535 ... | case1 y=p = case1 (subst (λ k → k ≡ _ ) &iso y=p )
960
b7370c39769e IsMinSUP< is wrong
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 959
diff changeset
1536 ... | case2 y<p = case2 (subst (λ k → * k < _ ) &iso y<p )
935
ed711d7be191 mem exhaust fix on fixpoint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 934
diff changeset
1537 z14 : f sp ≡ sp
960
b7370c39769e IsMinSUP< is wrong
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 959
diff changeset
1538 z14 with ZChain.f-total zc (subst (λ k → odef chain k) (sym &iso) (ZChain.f-next zc z12 )) (subst (λ k → odef chain k) (sym &iso) z12 )
924
a48dc906796c supf usp0 instead of supf (& A) ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 923
diff changeset
1539 ... | tri< a ¬b ¬c = ⊥-elim z16 where
a48dc906796c supf usp0 instead of supf (& A) ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 923
diff changeset
1540 z16 : ⊥
1032
382680c3e9be minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1031
diff changeset
1541 z16 with proj1 (mf (( MinSUP.sup sp1)) ( MinSUP.as sp1 ))
1031
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1030
diff changeset
1542 ... | case1 eq = ⊥-elim (¬b (sym (cong (*) eq ) ))
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
1543 ... | case2 lt = ⊥-elim (¬c lt )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
1544 ... | tri≈ ¬a b ¬c = subst₂ (λ j k → j ≡ k ) &iso &iso ( cong (&) b )
924
a48dc906796c supf usp0 instead of supf (& A) ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 923
diff changeset
1545 ... | tri> ¬a ¬b c = ⊥-elim z17 where
959
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 958
diff changeset
1546 z15 : (f sp ≡ MinSUP.sup sp1) ∨ (* (f sp) < * (MinSUP.sup sp1) )
960
b7370c39769e IsMinSUP< is wrong
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 959
diff changeset
1547 z15 = MinSUP.x≤sup sp1 (ZChain.f-next zc z12 )
924
a48dc906796c supf usp0 instead of supf (& A) ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 923
diff changeset
1548 z17 : ⊥
a48dc906796c supf usp0 instead of supf (& A) ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 923
diff changeset
1549 z17 with z15
960
b7370c39769e IsMinSUP< is wrong
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 959
diff changeset
1550 ... | case1 eq = ¬b (cong (*) eq)
b7370c39769e IsMinSUP< is wrong
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 959
diff changeset
1551 ... | case2 lt = ¬a lt
924
a48dc906796c supf usp0 instead of supf (& A) ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 923
diff changeset
1552
921
c0cf2b383064 UnionZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 920
diff changeset
1553 -- ZChain contradicts ¬ Maximal
c0cf2b383064 UnionZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 920
diff changeset
1554 --
c0cf2b383064 UnionZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 920
diff changeset
1555 -- ZChain forces fix point on any ≤-monotonic function (fixpoint)
c0cf2b383064 UnionZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 920
diff changeset
1556 -- ¬ Maximal create cf which is a <-monotonic function by axiom of choice. This contradicts fix point of ZChain
c0cf2b383064 UnionZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 920
diff changeset
1557 --
924
a48dc906796c supf usp0 instead of supf (& A) ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 923
diff changeset
1558
1090
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
1559 ¬Maximal→¬cf-mono : (nmx : ¬ Maximal A ) → (zc : ZChain A (cf nmx) (cf-is-<-monotonic nmx) as0 (& A)) → ⊥
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
1560 ¬Maximal→¬cf-mono nmx zc = <-irr0 {* (cf nmx c)} {* c}
1032
382680c3e9be minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1031
diff changeset
1561 (subst (λ k → odef A k ) (sym &iso) (proj1 (is-cf nmx (MinSUP.as msp1 ))))
382680c3e9be minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1031
diff changeset
1562 (subst (λ k → odef A k) (sym &iso) (MinSUP.as msp1) )
992
4aaecae58da5 ... x < & A ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 991
diff changeset
1563 (case1 ( cong (*)( fixpoint (cf nmx) (cf-is-≤-monotonic nmx ) (cf-is-<-monotonic nmx ) zc msp1 ))) -- x ≡ f x ̄
1032
382680c3e9be minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1031
diff changeset
1564 (proj1 (cf-is-<-monotonic nmx c (MinSUP.as msp1 ))) where -- x < f x
937
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 936
diff changeset
1565
927
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 926
diff changeset
1566 supf = ZChain.supf zc
934
ebcad8e5ae55 resync zorn.agda
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 930
diff changeset
1567 msp1 : MinSUP A (ZChain.chain zc)
1039
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1038
diff changeset
1568 msp1 = msp0 (cf nmx) (cf-is-<-monotonic nmx) as0 zc
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
1569 c : Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
1570 c = MinSUP.sup msp1
934
ebcad8e5ae55 resync zorn.agda
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 930
diff changeset
1571
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
1572 zorn00 : Maximal A
1090
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
1573 zorn00 with is-o∅ ( & HasMaximal )
804
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 803
diff changeset
1574 ... | no not = record { maximal = ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) ; as = zorn01 ; ¬maximal<x = zorn02 } where
1090
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
1575 -- yes we have the maximal because of the axiom of choice
551
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
1576 zorn03 : odef HasMaximal ( & ( ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) ) )
606
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 605
diff changeset
1577 zorn03 = ODC.x∋minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) -- Axiom of choice
551
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
1578 zorn01 : A ∋ ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq))
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
1579 zorn01 = proj1 zorn03
551
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
1580 zorn02 : {x : HOD} → A ∋ x → ¬ (ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) < x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
1581 zorn02 {x} ax m<x = proj2 zorn03 (& x) ax (subst₂ (λ j k → j < k) (sym *iso) (sym *iso) m<x )
1090
2cf182f0f583 order removal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1089
diff changeset
1582 ... | yes ¬Maximal = ⊥-elim ( ¬Maximal→¬cf-mono nmx (SZ (cf nmx) (cf-is-<-monotonic nmx) as0 (& A) )) where
551
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
1583 -- if we have no maximal, make ZChain, which contradict SUP condition
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
1584 nmx : ¬ Maximal A
551
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
1585 nmx mx = ∅< {HasMaximal} zc5 ( ≡o∅→=od∅ ¬Maximal ) where
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
1586 zc5 : odef A (& (Maximal.maximal mx)) ∧ (( y : Ordinal ) → odef A y → ¬ (* (& (Maximal.maximal mx)) < * y))
804
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 803
diff changeset
1587 zc5 = ⟪ Maximal.as mx , (λ y ay mx<y → Maximal.¬maximal<x mx (subst (λ k → odef A k ) (sym &iso) ay) (subst (λ k → k < * y) *iso mx<y) ) ⟫
551
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 550
diff changeset
1588
516
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 515
diff changeset
1589 -- usage (see filter.agda )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 515
diff changeset
1590 --
1091
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
1591 -- import OD hiding ( _⊆_ )
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
1592 -- _⊆_ : ( A B : HOD ) → Set n
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
1593 -- _⊆_ A B = {x : Ordinal } → odef A x → odef B x
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
1594 --
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
1595 -- import zorn
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
1596 -- open zorn O _⊆_ -- Zorn on Set inclusion order
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
1597 --
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
1598 -- open import Relation.Binary.Structures
482
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 481
diff changeset
1599
966
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 965
diff changeset
1600 -- MaximumSubset : {L P : HOD}
497
2a8629b5cff9 other strategy
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 496
diff changeset
1601 -- → o∅ o< & L → o∅ o< & P → P ⊆ L
1091
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
1602 -- → IsPartialOrderSet P _⊆_
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
1603 -- → ( (B : HOD) → B ⊆ P → IsTotalOrderSet B _⊆_ → SUP P B _⊆_ )
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
1604 -- → Maximal P (_⊆_)
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
1605 -- MaximumSubset {L} {P} 0<L 0<P P⊆L PO SP = Zorn-lemma {P} {_⊆_} 0<P PO SP
63c1167b2343 fix comments
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1090
diff changeset
1606 --