Basics for the Rational Numbers #
Rational numbers, implemented as a pair of integers num / den such that the
denominator is positive and the numerator and denominator are coprime.
- mk' :: (
- num : IntThe numerator of the rational number is an integer. 
- den : NatThe denominator of the rational number is a natural number. 
- The denominator is nonzero. 
- The numerator and denominator are coprime: it is in "reduced form". 
- )
Instances For
Equations
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- instHashableRat = { hash := instHashableRat.hash }
Equations
- instInhabitedRat = { default := { num := 0, den_nz := instInhabitedRat._proof_1, reduced := instInhabitedRat._proof_2 } }
Auxiliary definition for Rat.normalize. Constructs num / den as a rational number,
dividing both num and den by g (which is the gcd of the two) if it is not 1.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Construct a normalized Rat from a numerator and nonzero denominator.
This is a "smart constructor" that divides the numerator and denominator by
the gcd to ensure that the resulting rational number is normalized.
Equations
- Rat.normalize num den den_nz = Rat.maybeNormalize num den (num.natAbs.gcd den) ⋯ ⋯ ⋯ ⋯
Instances For
Construct a rational number from a numerator and denominator.
This is a "smart constructor" that divides the numerator and denominator by
the gcd to ensure that the resulting rational number is normalized, and returns
zero if den is zero.
Equations
- mkRat num den = if den_nz : den = 0 then { num := 0, den_nz := instInhabitedRat._proof_1, reduced := instInhabitedRat._proof_2 } else Rat.normalize num den den_nz
Instances For
Form the quotient n / d where n d : Int.
Equations
- Rat.divInt x✝ (Int.ofNat d) = inline (mkRat x✝ d)
- Rat.divInt x✝ (Int.negSucc d) = Rat.normalize (-x✝) d.succ ⋯
Instances For
Form the quotient n / d where n d : Int.
Equations
- Rat.«term_/._» = Lean.ParserDescr.trailingNode `Rat.«term_/._» 70 70 (Lean.ParserDescr.binary `andthen (Lean.ParserDescr.symbol " /. ") (Lean.ParserDescr.cat `term 71))
Instances For
Implements "scientific notation" 123.4e-5 for rational numbers. (This definition is
@[irreducible] because you don't want to unfold it. Use Rat.ofScientific_def,
Rat.ofScientific_true_def, or Rat.ofScientific_false_def instead.)
Instances For
Equations
- Rat.instOfScientific = { ofScientific := fun (m : Nat) (s : Bool) (e : Nat) => Rat.ofScientific (OfNat.ofNat m) s (OfNat.ofNat e) }
Equations
- a.instDecidableLt b = inferInstanceAs (Decidable (a.blt b = true))
Equations
- a.instDecidableLe b = inferInstanceAs (Decidable (b.blt a = false))
Multiplication of rational numbers. (This definition is @[irreducible] because you don't
want to unfold it. Use Rat.mul_def instead.)
Equations
- One or more equations did not get rendered due to their size.
Instances For
Addition of rational numbers. (This definition is @[irreducible] because you don't want to
unfold it. Use Rat.add_def instead.)
Equations
- One or more equations did not get rendered due to their size.
Instances For
Subtraction of rational numbers. (This definition is @[irreducible] because you don't want to
unfold it. Use Rat.sub_def instead.)
Equations
- One or more equations did not get rendered due to their size.