{-# OPTIONS --cubical-compatible --safe #-} module nat where open import Data.Nat open import Data.Nat.Properties open import Data.Empty open import Relation.Nullary open import Relation.Binary.PropositionalEquality open import Relation.Binary.Core open import Relation.Binary.Definitions open import logic open import Level hiding ( zero ; suc ) nat-<> : { x y : ℕ } → x < y → y < x → ⊥ nat-<> (s≤s x x : { x y : ℕ } → x ≤ y → y < x → ⊥ nat-≤> (s≤s x x→¬< : {x y : ℕ } → (x < y ) → ¬ ( y < x ) >→¬< (s≤s x→¬< x y → minus x y + y ≡ x minus+n {x} {zero} _ = trans (sym (+-comm zero _ )) refl minus+n {zero} {suc y} (s≤s ()) minus+n {suc x} {suc y} (s≤s lt) = begin minus (suc x) (suc y) + suc y ≡⟨ +-comm _ (suc y) ⟩ suc y + minus x y ≡⟨ cong ( λ k → suc k ) ( begin y + minus x y ≡⟨ +-comm y _ ⟩ minus x y + y ≡⟨ minus+n {x} {y} lt ⟩ x ∎ ) ⟩ suc x ∎ where open ≡-Reasoning <-minus-0 : {x y z : ℕ } → z + x < z + y → x < y <-minus-0 {x} {suc _} {zero} lt = lt <-minus-0 {x} {y} {suc z} (s≤s lt) = <-minus-0 {x} {y} {z} lt <-minus : {x y z : ℕ } → x + z < y + z → x < y <-minus {x} {y} {z} lt = <-minus-0 ( subst₂ ( λ j k → j < k ) (+-comm x _) (+-comm y _ ) lt ) x≤x+y : {z y : ℕ } → z ≤ z + y x≤x+y {zero} {y} = z≤n x≤x+y {suc z} {y} = s≤s (x≤x+y {z} {y}) x≤y+x : {z y : ℕ } → z ≤ y + z x≤y+x {z} {y} = subst (λ k → z ≤ k ) (+-comm _ y ) x≤x+y x≤x+sy : {x y : ℕ} → x < x + suc y x≤x+sy {x} {y} = begin suc x ≤⟨ x≤x+y ⟩ suc x + y ≡⟨ cong (λ k → k + y) (+-comm 1 x ) ⟩ (x + 1) + y ≡⟨ (+-assoc x 1 _) ⟩ x + suc y ∎ where open ≤-Reasoning <-plus : {x y z : ℕ } → x < y → x + z < y + z <-plus {zero} {suc y} {z} (s≤s z≤n) = s≤s (subst (λ k → z ≤ k ) (+-comm z _ ) x≤x+y ) <-plus {suc x} {suc y} {z} (s≤s lt) = s≤s (<-plus {x} {y} {z} lt) <-plus-0 : {x y z : ℕ } → x < y → z + x < z + y <-plus-0 {x} {y} {z} lt = subst₂ (λ j k → j < k ) (+-comm _ z) (+-comm _ z) ( <-plus {x} {y} {z} lt ) ≤-plus : {x y z : ℕ } → x ≤ y → x + z ≤ y + z ≤-plus {0} {y} {zero} z≤n = z≤n ≤-plus {0} {y} {suc z} z≤n = subst (λ k → z < k ) (+-comm _ y ) x≤x+y ≤-plus {suc x} {suc y} {z} (s≤s lt) = s≤s ( ≤-plus {x} {y} {z} lt ) ≤-plus-0 : {x y z : ℕ } → x ≤ y → z + x ≤ z + y ≤-plus-0 {x} {y} {zero} lt = lt ≤-plus-0 {x} {y} {suc z} lt = s≤s ( ≤-plus-0 {x} {y} {z} lt ) x+y ¬a ¬b c = case2 (0 : {x y : ℕ } → x < y → 0 < minus y x minus>0 {zero} {suc _} (s≤s z≤n) = s≤s z≤n minus>0 {suc x} {suc y} (s≤s lt) = minus>0 {x} {y} lt minus>0→x0→x ¬a ¬b c = ⊥-elim ( nat-≡< (sym (minus<=0 {y} (≤-trans refl-≤s c ))) lt ) minus+y-y : {x y : ℕ } → (x + y) - y ≡ x minus+y-y {zero} {y} = minus<=0 {zero + y} {y} ≤-refl minus+y-y {suc x} {y} = begin (suc x + y) - y ≡⟨ sym (minus+1 {_} {y} x≤y+x) ⟩ suc ((x + y) - y) ≡⟨ cong suc (minus+y-y {x} {y}) ⟩ suc x ∎ where open ≡-Reasoning minus+yx-yz : {x y z : ℕ } → (y + x) - (y + z) ≡ x - z minus+yx-yz {x} {zero} {z} = refl minus+yx-yz {x} {suc y} {z} = minus+yx-yz {x} {y} {z} minus+xy-zy : {x y z : ℕ } → (x + y) - (z + y) ≡ x - z minus+xy-zy {x} {y} {z} = subst₂ (λ j k → j - k ≡ x - z ) (+-comm y x) (+-comm y z) (minus+yx-yz {x} {y} {z}) +cancel ¬a ¬b c = subst ( λ k → k < y ) (sym (minus<=0 {y} {x} (≤-trans (≤-trans refl-≤s refl-≤s) c))) 0 ¬a ¬b c = +m= {_} {_} {suc y * z} ( begin minus x (suc y) * z + suc y * z ≡⟨ sym (proj₂ *-distrib-+ z (minus x (suc y) ) _) ⟩ ( minus x (suc y) + suc y ) * z ≡⟨ cong (λ k → k * z) (minus+n {x} {suc y} (s≤s c)) ⟩ x * z ≡⟨ sym (minus+n {x * z} {suc y * z} (s≤s (lt c))) ⟩ minus (x * z) (suc y * z) + suc y * z ∎ ) where open ≡-Reasoning lt : {x y z : ℕ } → suc y ≤ x → z + y * z ≤ x * z lt {x} {y} {z} le = *≤ le distr-minus-*' : {z x y : ℕ } → z * (minus x y) ≡ minus (z * x) (z * y) distr-minus-*' {z} {x} {y} = begin z * (minus x y) ≡⟨ *-comm _ (x - y) ⟩ (minus x y) * z ≡⟨ distr-minus-* {x} {y} {z} ⟩ minus (x * z) (y * z) ≡⟨ cong₂ (λ j k → j - k ) (*-comm x z ) (*-comm y z) ⟩ minus (z * x) (z * y) ∎ where open ≡-Reasoning minus- : {x y z : ℕ } → suc x > z + y → minus (minus x y) z ≡ minus x (y + z) minus- {x} {y} {z} gt = +m= {_} {_} {z} ( begin minus (minus x y) z + z ≡⟨ minus+n {_} {z} lemma ⟩ minus x y ≡⟨ +m= {_} {_} {y} ( begin minus x y + y ≡⟨ minus+n {_} {y} lemma1 ⟩ x ≡⟨ sym ( minus+n {_} {z + y} gt ) ⟩ minus x (z + y) + (z + y) ≡⟨ sym ( +-assoc (minus x (z + y)) _ _ ) ⟩ minus x (z + y) + z + y ∎ ) ⟩ minus x (z + y) + z ≡⟨ cong (λ k → minus x k + z ) (+-comm _ y ) ⟩ minus x (y + z) + z ∎ ) where open ≡-Reasoning lemma1 : suc x > y lemma1 = x+y z lemma = <-minus {_} {_} {y} ( subst ( λ x → z + y < suc x ) (sym (minus+n {x} {y} lemma1 )) gt ) minus-* : {M k n : ℕ } → n < k → minus k (suc n) * M ≡ minus (minus k n * M ) M minus-* {zero} {k} {n} lt = begin minus k (suc n) * zero ≡⟨ *-comm (minus k (suc n)) zero ⟩ zero * minus k (suc n) ≡⟨⟩ 0 * minus k n ≡⟨ *-comm 0 (minus k n) ⟩ minus (minus k n * 0 ) 0 ∎ where open ≡-Reasoning minus-* {suc m} {k} {n} lt with <-cmp k 1 minus-* {suc m} {.0} {zero} lt | tri< (s≤s z≤n) ¬b ¬c = refl minus-* {suc m} {.0} {suc n} lt | tri< (s≤s z≤n) ¬b ¬c = refl minus-* {suc zero} {.1} {zero} lt | tri≈ ¬a refl ¬c = refl minus-* {suc (suc m)} {.1} {zero} lt | tri≈ ¬a refl ¬c = minus-* {suc m} {1} {zero} lt minus-* {suc m} {.1} {suc n} (s≤s ()) | tri≈ ¬a refl ¬c minus-* {suc m} {k} {n} lt | tri> ¬a ¬b c = begin minus k (suc n) * M ≡⟨ distr-minus-* {k} {suc n} {M} ⟩ minus (k * M ) ((suc n) * M) ≡⟨⟩ minus (k * M ) (M + n * M ) ≡⟨ cong (λ x → minus (k * M) x) (+-comm M _ ) ⟩ minus (k * M ) ((n * M) + M ) ≡⟨ sym ( minus- {k * M} {n * M} (lemma lt) ) ⟩ minus (minus (k * M ) (n * M)) M ≡⟨ cong (λ x → minus x M ) ( sym ( distr-minus-* {k} {n} )) ⟩ minus (minus k n * M ) M ∎ where M = suc m lemma : {n k m : ℕ } → n < k → suc (k * suc m) > suc m + n * suc m lemma {zero} {suc k} {m} (s≤s lt) = s≤s (s≤s (subst (λ x → x ≤ m + k * suc m) (+-comm 0 _ ) x≤x+y )) lemma {suc n} {suc k} {m} lt = begin suc (suc m + suc n * suc m) ≡⟨⟩ suc ( suc (suc n) * suc m) ≤⟨ ≤-plus-0 {_} {_} {1} (*≤ lt ) ⟩ suc (suc k * suc m) ∎ where open ≤-Reasoning open ≡-Reasoning x=y+z→x-z=y : {x y z : ℕ } → x ≡ y + z → x - z ≡ y x=y+z→x-z=y {x} {zero} {.x} refl = minus<=0 {x} {x} refl-≤ -- x ≡ suc (y + z) → (x ≡ y + z → x - z ≡ y) → (x - z) ≡ suc y x=y+z→x-z=y {suc x} {suc y} {zero} eq = begin -- suc x ≡ suc (y + zero) → (suc x - zero) ≡ suc y suc x - zero ≡⟨ refl ⟩ suc x ≡⟨ eq ⟩ suc y + zero ≡⟨ +-comm _ zero ⟩ suc y ∎ where open ≡-Reasoning x=y+z→x-z=y {suc x} {suc y} {suc z} eq = x=y+z→x-z=y {x} {suc y} {z} ( begin x ≡⟨ cong pred eq ⟩ pred (suc y + suc z) ≡⟨ +-comm _ (suc z) ⟩ suc z + y ≡⟨ cong suc ( +-comm _ y ) ⟩ suc y + z ∎ ) where open ≡-Reasoning m*1=m : {m : ℕ } → m * 1 ≡ m m*1=m {zero} = refl m*1=m {suc m} = cong suc m*1=m +-cancel-1 : (x y z : ℕ ) → x + y ≡ x + z → y ≡ z +-cancel-1 zero y z eq = eq +-cancel-1 (suc x) y z eq = +-cancel-1 x y z (cong pred eq ) +-cancel-0 : (x y z : ℕ ) → y + x ≡ z + x → y ≡ z +-cancel-0 x y z eq = +-cancel-1 x y z (trans (+-comm x y) (trans eq (sym (+-comm x z)) )) *-cancel-left : {x y z : ℕ } → x > 0 → x * y ≡ x * z → y ≡ z *-cancel-left {suc x} {zero} {zero} lt eq = refl *-cancel-left {suc x} {zero} {suc z} lt eq = ⊥-elim ( nat-≡< eq (s≤s (begin x * zero ≡⟨ *-comm x _ ⟩ zero ≤⟨ z≤n ⟩ z + x * suc z ∎ ))) where open ≤-Reasoning *-cancel-left {suc x} {suc y} {zero} lt eq = ⊥-elim ( nat-≡< (sym eq) (s≤s (begin x * zero ≡⟨ *-comm x _ ⟩ zero ≤⟨ z≤n ⟩ _ ∎ ))) where open ≤-Reasoning *-cancel-left {suc x} {suc y} {suc z} lt eq with cong pred eq ... | eq1 = cong suc (*-cancel-left {suc x} {y} {z} lt (+-cancel-0 x _ _ (begin y + x * y + x ≡⟨ +-assoc y _ _ ⟩ y + (x * y + x) ≡⟨ cong (λ k → y + (k + x)) (*-comm x _) ⟩ y + (y * x + x) ≡⟨ cong (_+_ y) (+-comm _ x) ⟩ y + (x + y * x ) ≡⟨ refl ⟩ y + suc y * x ≡⟨ cong (_+_ y) (*-comm (suc y) _) ⟩ y + x * suc y ≡⟨ eq1 ⟩ z + x * suc z ≡⟨ refl ⟩ _ ≡⟨ sym ( cong (_+_ z) (*-comm (suc z) _) ) ⟩ _ ≡⟨ sym ( cong (_+_ z) (+-comm _ x)) ⟩ z + (z * x + x) ≡⟨ sym ( cong (λ k → z + (k + x)) (*-comm x _) ) ⟩ z + (x * z + x) ≡⟨ sym ( +-assoc z _ _) ⟩ z + x * z + x ∎ ))) where open ≡-Reasoning record Finduction {n m : Level} (P : Set n ) (Q : P → Set m ) (f : P → ℕ) : Set (n Level.⊔ m) where field fzero : {p : P} → f p ≡ zero → Q p pnext : (p : P ) → P decline : {p : P} → 0 < f p → f (pnext p) < f p ind : {p : P} → Q (pnext p) → Q p y ¬a ¬b () ... | tri≈ ¬a b ¬c = Finduction.fzero I (sym b) ... | tri< lt _ _ = f-induction0 p (f p) ( ¬a ¬b c = ⊥-elim ( nat-≤> le c ) record Ninduction {n m : Level} (P : Set n ) (Q : P → Set m ) (f : P → ℕ) : Set (n Level.⊔ m) where field pnext : (p : P ) → P fzero : {p : P} → f (pnext p) ≡ zero → Q p decline : {p : P} → 0 < f p → f (pnext p) < f p ind : {p : P} → Q (pnext p) → Q p s≤s→≤ : { i j : ℕ} → suc i ≤ suc j → i ≤ j s≤s→≤ (s≤s lt) = lt n-induction : {n m : Level} {P : Set n } → {Q : P → Set m } → (f : P → ℕ) → Ninduction P Q f → (p : P ) → Q p n-induction {n} {m} {P} {Q} f I p = f-induction0 p (f (Ninduction.pnext I p)) ≤-refl where f-induction0 : (p : P) → (x : ℕ) → (f (Ninduction.pnext I p)) ≤ x → Q p f-induction0 p zero lt = Ninduction.fzero I {p} (fi0 _ lt) f-induction0 p (suc x) le with <-cmp (f (Ninduction.pnext I p)) (suc x) ... | tri< (s≤s a) ¬b ¬c = f-induction0 p x a ... | tri≈ ¬a b ¬c = Ninduction.ind I (f-induction0 (Ninduction.pnext I p) x (s≤s→≤ nle) ) where f>0 : 0 < f (Ninduction.pnext I p) f>0 = subst (λ k → 0 < k ) (sym b) ( s≤s z≤n ) nle : suc (f (Ninduction.pnext I (Ninduction.pnext I p))) ≤ suc x nle = subst (λ k → suc (f (Ninduction.pnext I (Ninduction.pnext I p))) ≤ k) b (Ninduction.decline I {Ninduction.pnext I p} f>0 ) ... | tri> ¬a ¬b c = ⊥-elim ( nat-≤> le c ) record Factor (n m : ℕ ) : Set where field factor : ℕ remain : ℕ is-factor : factor * n + remain ≡ m record Dividable (n m : ℕ ) : Set where field factor : ℕ is-factor : factor * n + 0 ≡ m open Factor DtoF : {n m : ℕ} → Dividable n m → Factor n m DtoF {n} {m} record { factor = f ; is-factor = fa } = record { factor = f ; remain = 0 ; is-factor = fa } FtoD : {n m : ℕ} → (fc : Factor n m) → remain fc ≡ 0 → Dividable n m FtoD {n} {m} record { factor = f ; remain = r ; is-factor = fa } refl = record { factor = f ; is-factor = fa } --divdable^2 : ( n k : ℕ ) → Dividable k ( n * n ) → Dividable k n --divdable^2 n k dn2 = {!!} decf : { n k : ℕ } → ( x : Factor k (suc n) ) → Factor k n decf {n} {k} record { factor = f ; remain = r ; is-factor = fa } = decf1 {n} {k} f r fa where decf1 : { n k : ℕ } → (f r : ℕ) → (f * k + r ≡ suc n) → Factor k n decf1 {n} {k} f (suc r) fa = -- this case must be the first record { factor = f ; remain = r ; is-factor = ( begin -- fa : f * k + suc r ≡ suc n f * k + r ≡⟨ cong pred ( begin suc ( f * k + r ) ≡⟨ +-comm _ r ⟩ r + suc (f * k) ≡⟨ sym (+-assoc r 1 _) ⟩ (r + 1) + f * k ≡⟨ cong (λ t → t + f * k ) (+-comm r 1) ⟩ (suc r ) + f * k ≡⟨ +-comm (suc r) _ ⟩ f * k + suc r ≡⟨ fa ⟩ suc n ∎ ) ⟩ n ∎ ) } where open ≡-Reasoning decf1 {n} {zero} (suc f) zero fa = ⊥-elim ( nat-≡< fa ( begin suc (suc f * zero + zero) ≡⟨ cong suc (+-comm _ zero) ⟩ suc (f * 0) ≡⟨ cong suc (*-comm f zero) ⟩ suc zero ≤⟨ s≤s z≤n ⟩ suc n ∎ )) where open ≤-Reasoning decf1 {n} {suc k} (suc f) zero fa = record { factor = f ; remain = k ; is-factor = ( begin -- fa : suc (k + f * suc k + zero) ≡ suc n f * suc k + k ≡⟨ +-comm _ k ⟩ k + f * suc k ≡⟨ +-comm zero _ ⟩ (k + f * suc k) + zero ≡⟨ cong pred fa ⟩ n ∎ ) } where open ≡-Reasoning div0 : {k : ℕ} → Dividable k 0 div0 {k} = record { factor = 0; is-factor = refl } div= : {k : ℕ} → Dividable k k div= {k} = record { factor = 1; is-factor = ( begin k + 0 * k + 0 ≡⟨ trans ( +-comm _ 0) ( +-comm _ 0) ⟩ k ∎ ) } where open ≡-Reasoning div1 : { k : ℕ } → k > 1 → ¬ Dividable k 1 div1 {k} k>1 record { factor = (suc f) ; is-factor = fa } = ⊥-elim ( nat-≡< (sym fa) ( begin 2 ≤⟨ k>1 ⟩ k ≡⟨ +-comm 0 _ ⟩ k + 0 ≡⟨ refl ⟩ 1 * k ≤⟨ *-mono-≤ {1} {suc f} (s≤s z≤n ) ≤-refl ⟩ suc f * k ≡⟨ +-comm 0 _ ⟩ suc f * k + 0 ∎ )) where open ≤-Reasoning div+div : { i j k : ℕ } → Dividable k i → Dividable k j → Dividable k (i + j) ∧ Dividable k (j + i) div+div {i} {j} {k} di dj = ⟪ div+div1 , subst (λ g → Dividable k g) (+-comm i j) div+div1 ⟫ where fki = Dividable.factor di fkj = Dividable.factor dj div+div1 : Dividable k (i + j) div+div1 = record { factor = fki + fkj ; is-factor = ( begin (fki + fkj) * k + 0 ≡⟨ +-comm _ 0 ⟩ (fki + fkj) * k ≡⟨ *-distribʳ-+ k fki _ ⟩ fki * k + fkj * k ≡⟨ cong₂ ( λ i j → i + j ) (+-comm 0 (fki * k)) (+-comm 0 (fkj * k)) ⟩ (fki * k + 0) + (fkj * k + 0) ≡⟨ cong₂ ( λ i j → i + j ) (Dividable.is-factor di) (Dividable.is-factor dj) ⟩ i + j ∎ ) } where open ≡-Reasoning div-div : { i j k : ℕ } → k > 1 → Dividable k i → Dividable k j → Dividable k (i - j) ∧ Dividable k (j - i) div-div {i} {j} {k} k>1 di dj = ⟪ div-div1 di dj , div-div1 dj di ⟫ where div-div1 : {i j : ℕ } → Dividable k i → Dividable k j → Dividable k (i - j) div-div1 {i} {j} di dj = record { factor = fki - fkj ; is-factor = ( begin (fki - fkj) * k + 0 ≡⟨ +-comm _ 0 ⟩ (fki - fkj) * k ≡⟨ distr-minus-* {fki} {fkj} ⟩ (fki * k) - (fkj * k) ≡⟨ cong₂ ( λ i j → i - j ) (+-comm 0 (fki * k)) (+-comm 0 (fkj * k)) ⟩ (fki * k + 0) - (fkj * k + 0) ≡⟨ cong₂ ( λ i j → i - j ) (Dividable.is-factor di) (Dividable.is-factor dj) ⟩ i - j ∎ ) } where open ≡-Reasoning fki = Dividable.factor di fkj = Dividable.factor dj open _∧_ div+1 : { i k : ℕ } → k > 1 → Dividable k i → ¬ Dividable k (suc i) div+1 {i} {k} k>1 d d1 = div1 k>1 div+11 where div+11 : Dividable k 1 div+11 = subst (λ g → Dividable k g) (minus+y-y {1} {i} ) ( proj2 (div-div k>1 d d1 ) ) div 1 → m > 0 → m < k → ¬ Dividable k m div1 m>0 m (div0 ) div 0 → m > 0 → (d : Dividable k m ) → Dividable.factor d > 0 00 m>0 d with Dividable.factor d in eq1 ... | zero = ⊥-elim ( nat-≡< ff1 m>0 ) where ff1 : 0 ≡ m ff1 = begin 0 ≡⟨⟩ 0 * k + 0 ≡⟨ cong (λ j → j * k + 0) (sym eq1) ⟩ Dividable.factor d * k + 0 ≡⟨ Dividable.is-factor d ⟩ m ∎ where open ≡-Reasoning ... | suc t = s≤s z≤n div→k≤m : { m k : ℕ } → k > 1 → m > 0 → Dividable k m → m ≥ k div→k≤m {m} {k} k>1 m>0 d with <-cmp m k ... | tri< a ¬b ¬c = ⊥-elim ( div1 m>0 a d ) ... | tri≈ ¬a refl ¬c = ≤-refl ... | tri> ¬a ¬b c = 1 → Dec (Dividable k m ) decD {k} {m} k>1 = n-induction {_} {_} {ℕ} {λ m → Dec (Dividable k m ) } F I m where F : ℕ → ℕ F m = m F0 : ( m : ℕ ) → F (m - k) ≡ 0 → Dec (Dividable k m ) F0 0 eq = yes record { factor = 0 ; is-factor = refl } F0 (suc m) eq with <-cmp k (suc m) ... | tri< a ¬b ¬c = yes record { factor = 1 ; is-factor = subst (λ g → 1 * k + 0 ≡ g ) (sym (i-j=0→i=j ( ¬a ¬b c = no ( λ d → ⊥-elim (div1 (s≤s z≤n ) c d) ) decl : {m : ℕ } → 0 < m → m - k < m decl {m} 01 ) 0 ¬a ¬b k

¬a ¬b₁ c = no (λ d → not-div p (Dividable.factor d) a c (Dividable.is-factor d) ) where not-div : (p f : ℕ) → p < k → 0 < p → f * k + 0 ≡ p → ⊥ not-div (suc p) (suc f) p1 d div=)) ) I : Ninduction ℕ _ F I = record { pnext = λ p → p - k ; fzero = λ {m} eq → F0 m eq ; decline = λ {m} lt → decl lt ; ind = λ {p} prev → ind p prev }