{-# OPTIONS --cubical --safe #-}
module CantorNormalForm.Wellfoundedness where
open import Cubical.Foundations.Prelude
open import Cubical.Data.Sum
open import Cubical.Data.Sigma
open import Cubical.Data.Empty
renaming (rec to β₯-rec)
open import Cubical.Relation.Nullary
open import Cubical.Relation.Binary
open BinaryRelation
open import Cubical.Induction.WellFounded
renaming (WellFounded to isWellFounded;
Acc to isAcc)
open import CantorNormalForm.Base
πIsAcc : isAcc _<_ π
πIsAcc = acc (Ξ» _ ())
mutual
Lm[fst-acc] : β {a a'} β isAcc _<_ a' β a β‘ a'
β β {b x} β isAcc _<_ b β x < a' β (r : x β₯ left b)
β isAcc _<_ (Ο^ x + b [ r ])
Lm[fst-acc] {a} {a'} (acc ΞΎ) a=a' {b} {x} acα΅ x<a' r = acc goal
where
goal : β z β z < Ο^ x + b [ r ] β isAcc _<_ z
goal π <β = πIsAcc
goal (Ο^ c + d [ s ]) (<β c<y) = Lm[fst-acc] (ΞΎ x x<a') refl IH c<y s
where
IH : isAcc _<_ d
IH = goal d (<-trans (right< c d s) (<β c<y))
goal (Ο^ c + d [ s ]) (<β c=y d<b) = Lm[snd-acc] (ΞΎ x x<a') c=y acα΅ d<b s
Lm[snd-acc] : β {a a'} β isAcc _<_ a' β a β‘ a'
β β {c y} β isAcc _<_ c β y < c β (r : a β₯ left y)
β isAcc _<_ (Ο^ a + y [ r ])
Lm[snd-acc] {a} {a'} acα΅ a=a' {c} {y} (acc ΞΎαΆ) y<c r = acc goal
where
goal : β z β z < Ο^ a + y [ r ] β isAcc _<_ z
goal π <β = πIsAcc
goal (Ο^ b + d [ t ]) (<β b<a) = Lm[fst-acc] acα΅ a=a' IH (subst (b <_) a=a' b<a) t
where
IH : isAcc _<_ d
IH = goal d (<-trans (right< b d t) (<β b<a))
goal (Ο^ b + d [ t ]) (<β b=a d<y) = Lm[snd-acc] acα΅ (b=a β a=a') (ΞΎαΆ y y<c) d<y t
<-is-wellfounded : isWellFounded _<_
<-is-wellfounded π = πIsAcc
<-is-wellfounded (Ο^ a + b [ r ]) = acc goal
where
goal : β z β z < Ο^ a + b [ r ] β isAcc _<_ z
goal π <β = πIsAcc
goal (Ο^ c + d [ s ]) (<β c<a) =
Lm[fst-acc] (<-is-wellfounded a) refl IH c<a s
where
IH : isAcc _<_ d
IH = goal d (<-trans (right< c d s) (<β c<a))
goal (Ο^ c + d [ s ]) (<β c=a d<b) =
Lm[snd-acc] (<-is-wellfounded a) c=a (<-is-wellfounded b) d<b s
open WFI <-is-wellfounded
renaming (induction to wf-ind)
public