# HG changeset patch # User Shinji KONO # Date 1625230815 -32400 # Node ID f094694aeec564dab280b626950e19db9217f8f8 # Parent 04f0b553db3418f793ff30e5d3538c5669a0082b ... diff -r 04f0b553db34 -r f094694aeec5 src/HyperReal.agda --- a/src/HyperReal.agda Fri Jul 02 17:21:27 2021 +0900 +++ b/src/HyperReal.agda Fri Jul 02 22:00:15 2021 +0900 @@ -8,6 +8,7 @@ open import Relation.Binary.PropositionalEquality hiding ( [_] ) open import Relation.Binary.Definitions open import Function.Bijection +open import Relation.Binary.Structures open import nat open import logic @@ -104,6 +105,9 @@ postulate _cmpn_ : ( i j : HyperNat ) → Dec ( i n≤ j ) +HNTotalOrder : IsTotalPreorder HyperNat ? _n≤_ +HNTotalOrder = ? + data HyperZ : Set where hz : HyperNat → HyperNat → HyperZ @@ -123,6 +127,8 @@ ≤→= {0} {0} z≤n z≤n = refl ≤→= {suc i} {suc j} (s≤s i_ : Bool → Bool → Bool +true <=> true = true +false <=> false = true +_ <=> _ = false + +open import Relation.Binary.PropositionalEquality + +¬t=f : (t : Bool ) → ¬ ( not t ≡ t) +¬t=f true () +¬t=f false () + +infixr 130 _\/_ +infixr 140 _/\_ + +≡-Bool-func : {A B : Bool } → ( A ≡ true → B ≡ true ) → ( B ≡ true → A ≡ true ) → A ≡ B +≡-Bool-func {true} {true} a→b b→a = refl +≡-Bool-func {false} {true} a→b b→a with b→a refl +... | () +≡-Bool-func {true} {false} a→b b→a with a→b refl +... | () +≡-Bool-func {false} {false} a→b b→a = refl + +bool-≡-? : (a b : Bool) → Dec ( a ≡ b ) +bool-≡-? true true = yes refl +bool-≡-? true false = no (λ ()) +bool-≡-? false true = no (λ ()) +bool-≡-? false false = yes refl + +¬-bool-t : {a : Bool} → ¬ ( a ≡ true ) → a ≡ false +¬-bool-t {true} ne = ⊥-elim ( ne refl ) +¬-bool-t {false} ne = refl + +¬-bool-f : {a : Bool} → ¬ ( a ≡ false ) → a ≡ true +¬-bool-f {true} ne = refl +¬-bool-f {false} ne = ⊥-elim ( ne refl ) + +¬-bool : {a : Bool} → a ≡ false → a ≡ true → ⊥ +¬-bool refl () + +lemma-∧-0 : {a b : Bool} → a /\ b ≡ true → a ≡ false → ⊥ +lemma-∧-0 {true} {true} refl () +lemma-∧-0 {true} {false} () +lemma-∧-0 {false} {true} () +lemma-∧-0 {false} {false} () + +lemma-∧-1 : {a b : Bool} → a /\ b ≡ true → b ≡ false → ⊥ +lemma-∧-1 {true} {true} refl () +lemma-∧-1 {true} {false} () +lemma-∧-1 {false} {true} () +lemma-∧-1 {false} {false} () + +bool-and-tt : {a b : Bool} → a ≡ true → b ≡ true → ( a /\ b ) ≡ true +bool-and-tt refl refl = refl + +bool-∧→tt-0 : {a b : Bool} → ( a /\ b ) ≡ true → a ≡ true +bool-∧→tt-0 {true} {true} refl = refl +bool-∧→tt-0 {false} {_} () + +bool-∧→tt-1 : {a b : Bool} → ( a /\ b ) ≡ true → b ≡ true +bool-∧→tt-1 {true} {true} refl = refl +bool-∧→tt-1 {true} {false} () +bool-∧→tt-1 {false} {false} () + +bool-or-1 : {a b : Bool} → a ≡ false → ( a \/ b ) ≡ b +bool-or-1 {false} {true} refl = refl +bool-or-1 {false} {false} refl = refl +bool-or-2 : {a b : Bool} → b ≡ false → (a \/ b ) ≡ a +bool-or-2 {true} {false} refl = refl +bool-or-2 {false} {false} refl = refl + +bool-or-3 : {a : Bool} → ( a \/ true ) ≡ true +bool-or-3 {true} = refl +bool-or-3 {false} = refl + +bool-or-31 : {a b : Bool} → b ≡ true → ( a \/ b ) ≡ true +bool-or-31 {true} {true} refl = refl +bool-or-31 {false} {true} refl = refl + +bool-or-4 : {a : Bool} → ( true \/ a ) ≡ true +bool-or-4 {true} = refl +bool-or-4 {false} = refl + +bool-or-41 : {a b : Bool} → a ≡ true → ( a \/ b ) ≡ true +bool-or-41 {true} {b} refl = refl + +bool-and-1 : {a b : Bool} → a ≡ false → (a /\ b ) ≡ false +bool-and-1 {false} {b} refl = refl +bool-and-2 : {a b : Bool} → b ≡ false → (a /\ b ) ≡ false +bool-and-2 {true} {false} refl = refl +bool-and-2 {false} {false} refl = refl + + diff -r 04f0b553db34 -r f094694aeec5 src/nat.agda --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/nat.agda Fri Jul 02 22:00:15 2021 +0900 @@ -0,0 +1,452 @@ +{-# OPTIONS --allow-unsolved-metas #-} +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 + +00 : {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}) + +y-x ¬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 + +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 ) + +