Skip to content

Conversation

@Taneb
Copy link
Member

@Taneb Taneb commented Nov 23, 2025

Closes #2890 (eventually)

The proofs, while purely algebraic, are often very tedious, especially *-cong.

I've got basic arithmetic, ordering with addition, and divmod working!

record _≃_ (i j : ℤ) : Set where
constructor mk≃
field
eq : ℤ.minuend i ℕ.+ ℤ.subtrahend j ≡ ℤ.minuend j ℕ.+ ℤ.subtrahend i
Copy link
Collaborator

@jamesmckinna jamesmckinna Nov 27, 2025

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

First observation from Properties: use of sym to flip rewrites in the RHS of such equations... so suggest instead

Suggested change
eq : ℤ.minuend i ℕ.+ ℤ.subtrahend j ≡ ℤ.minuend j ℕ.+ ℤ.subtrahend i
eq : ℤ.minuend i ℕ.+ ℤ.subtrahend j ≡ ℤ.subtrahend i ℕ.+ ℤ.minuend j

?

Copy link
Member Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

The reason I did it this way was so that reflexivity is exactly refl. But you're right, that doesn't seem to be worth the tradeoff

Comment on lines +52 to +54
1ℤ :
1ℤ = 10

Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Also:

Suggested change
1ℤ :
1ℤ = 1 ⊖ 0
1ℤ :
1ℤ = 1 ⊖ 0
-1ℤ :
-1ℤ = 0 ⊖ 1

Copy link
Member Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Easy enough, but does it pass the Fairbairn threshold? -1ℤ = - 1ℤ, after all

Comment on lines +67 to +73
infix 8 ⁻_ ⁺_

⁺_ :
⁺ n = n ⊖ 0

⁻_ :
⁻ n = 0 ⊖ n
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

The symbol is in particular quite hard to discern as distinct form -. Would it be better to use Sign here (and reimplement the corresponding signAbs/SignAbs plumbing)?

Copy link
Member Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I mostly got annoyed at Agda not being able to tell if a section (x +_) was _+_ x or x (+_), so I wanted different symbols for them... the superscript-minus symbol is used in APL's number syntax for a negative, and that's what inspired the use of these here. I'm not attached to them, though.

a ℕ.% n ℕ.+ 0 ≡⟨ ℕ.+-identityʳ (a ℕ.% n) ⟩
a ℕ.% n ≡⟨ ℕ.m∸n+n≡m p ⟨
a ℕ.% n ℕ.∸ b ℕ.% n ℕ.+ b ℕ.% n ∎
... | no b%n≰a%n = ≃-trans property′ $ mk≃ $ begin
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Oh. Tempting to use a (Semi)Ring solver here?

Copy link
Member Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Do those work well when I'm using monus? Not that I ever really can manage to get them to work well for me...

open import Data.Product.Base
open import Function.Base
open import Relation.Binary.PropositionalEquality

Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

is it easier to derive each of these properties from the 'adjoint equivalence'/'graph relation' forms:

fromINT i ≡ j  i ≃ toINT j
i ≃ toINT j  fromINT i ≡ j

(plus, perhaps, some additional lemmas to 'explain' these two graph relations)?
(I think even the congs follow from this, but I'd need to check that...)

Copy link
Member Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

The second one is easy enough, it's hidden in my surjective proof. I'll have a go at the first one later.

As weak evidence that the congs don't follow from this: they're separate fields in the IsInverse structure in Function.Structures

Copy link
Collaborator

@jamesmckinna jamesmckinna left a comment

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

No proper review as yet, but some suggestions for consideration.
Also I need to look again at Landau.

Copy link
Collaborator

@JacquesCarette JacquesCarette left a comment

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Some are requests for changes, some are suggestions.

I have quite a few more things I'd like simplified, but I feel that maybe I should do those myself as they require some level of experimenting.

... | yes _ = ℕ.≤-<-trans (ℕ.m∸n≤m (a ℕ.% n) (b ℕ.% n)) (ℕ.m%n<n a n)
... | no b%n≰a%n = ℕ.∸-monoʳ-< {o = 0} (ℕ.m<n⇒0<n∸m (ℕ.≰⇒> b%n≰a%n)) (ℕ.≤-trans (ℕ.m∸n≤m (b ℕ.% n) (a ℕ.% n)) (ℕ.<⇒≤ (ℕ.m%n<n b n)))

property : a ⊖ b ≃ ⁺ remainder + quotient * ⁺ n
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I wonder if we could do a giant "control inversion": the computation of remainder and quotient and this here property all do a case analysis on the same thing. Maybe we could case once and then return all 3 pieces of information?

; property = property
}
where
open ≡-Reasoning
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Given that all of the reasoning below is done in a ℕ context and the +, *, etc of ℤ are never used, might it make sense to qualify ℤ and not ℕ instead? That would definitely shorten a lot of things.

((n ℕ.∸ (v ℕ.∸ w)) ℕ.+ x) ℕ.+ (v ℕ.+ (y ℕ.+ z)) ∎
where
w = a ℕ.% n
x = a ℕ./ n ℕ.* n ℕ.+ b ℕ./ n ℕ.* 0
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

can't the above proof be greatly simplified by first getting rid of those * 0 in x and y, and then reconstructing them later?

Also, isn't it the case that z + v = b ?

fromINT-cong {a INT.⊖ b} {c INT.⊖ d} (mk≃ a+d≡c+b) = begin
a ⊖ b ≡⟨ m-n≡m⊖n a b ⟨
+ a - + b ≡⟨ cong (_- + b) (+-identityʳ (+ a)) ⟨
(+ a + 0ℤ) - + b ≡⟨ cong (λ z (+ a + z) - + b) (+-inverseʳ (+ d)) ⟨
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

identity followed by inverse are exactly the (opposite direction) of the cancel lemmas.

(+ a + 0ℤ) - + b ≡⟨ cong (λ z (+ a + z) - + b) (+-inverseʳ (+ d)) ⟨
(+ a + (+ d - + d)) - + b ≡⟨ cong (_- + b) (+-assoc (+ a) (+ d) (- + d)) ⟨
(+ (a ℕ.+ d) - + d) - + b ≡⟨ cong (λ z (+ z - + d) - + b) a+d≡c+b ⟩
(+ (c ℕ.+ b) - + d) - + b ≡⟨ cong (_- + b) (+-assoc (+ c) (+ b) (- + d)) ⟩
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

aren't these 2 lines some kind of middle4 ?

(+ c + (+ b - + d)) - + b ≡⟨ cong (λ z (+ c + z) - + b) (+-comm (+ b) (- + d)) ⟩
(+ c + (- + d + + b)) - + b ≡⟨ cong (_- + b) (+-assoc (+ c) (- + d) (+ b)) ⟨
((+ c - + d) + + b) - + b ≡⟨ +-assoc (+ c - + d) (+ b) (- + b) ⟩
(+ c - + d) + (+ b - + b) ≡⟨ cong₂ _+_ (m-n≡m⊖n c d) (+-inverseʳ (+ b)) ⟩
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

cancel!

pred :
pred (a ⊖ b) = a ⊖ ℕ.suc b

0ℤ :
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I wonder if it would be useful to also have non-normalized 0 and 1, i.e. 0ℤn {n} = n ⊖ n with the n in the name being a super/sub-script? I think you are prematurely normalizing in places.

∣m+n-m+o∣≡∣n-o∣ zero n o = refl
∣m+n-m+o∣≡∣n-o∣ (suc m) n o = ∣m+n-m+o∣≡∣n-o∣ m n o

∣m+o-n+o∣≡∣m-n∣ : m n o ∣ m + o - n + o ∣ ≡ ∣ m - n ∣
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

maybe add parens in the lhs to make it easier to parse for humans?

Sign up for free to join this conversation on GitHub. Already have an account? Sign in to comment

Labels

None yet

Projects

None yet

Development

Successfully merging this pull request may close these issues.

[ add ] INT view of Data.Integer.Base.ℤ

3 participants