Documentation

Mathlib.Data.Nat.BinaryRec

Binary recursion on Nat #

This file defines binary recursion on Nat.

Main results #

def Nat.bit (b : Bool) :
NatNat

bit b appends the digit b to the binary representation of its natural number input.

Equations
Instances For
    theorem Nat.shiftRight_one (n : Nat) :
    n >>> 1 = n / 2
    @[simp]
    theorem Nat.bit_testBit_zero_shiftRight_one (n : Nat) :
    Nat.bit (n.testBit 0) (n >>> 1) = n
    @[simp]
    theorem Nat.bit_eq_zero_iff {n : Nat} {b : Bool} :
    Nat.bit b n = 0 n = 0 b = false
    @[inline]
    def Nat.bitCasesOn {motive : NatSort u} (n : Nat) (h : (b : Bool) → (n : Nat) → motive (Nat.bit b n)) :
    motive n

    For a predicate motive : Nat → Sort u, if instances can be constructed for natural numbers of the form bit b n, they can be constructed for any given natural number.

    Equations
    Instances For
      @[irreducible, specialize #[]]
      def Nat.binaryRec {motive : NatSort u} (z : motive 0) (f : (b : Bool) → (n : Nat) → motive nmotive (Nat.bit b n)) (n : Nat) :
      motive n

      A recursion principle for bit representations of natural numbers. For a predicate motive : Nat → Sort u, if instances can be constructed for natural numbers of the form bit b n, they can be constructed for all natural numbers.

      Equations
      Instances For
        @[specialize #[]]
        def Nat.binaryRec' {motive : NatSort u} (z : motive 0) (f : (b : Bool) → (n : Nat) → (n = 0b = true)motive nmotive (Nat.bit b n)) (n : Nat) :
        motive n

        The same as binaryRec, but the induction step can assume that if n=0, the bit being appended is true

        Equations
        Instances For
          @[specialize #[]]
          def Nat.binaryRecFromOne {motive : NatSort u} (z₀ : motive 0) (z₁ : motive 1) (f : (b : Bool) → (n : Nat) → n 0motive nmotive (Nat.bit b n)) (n : Nat) :
          motive n

          The same as binaryRec, but special casing both 0 and 1 as base cases

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For
            theorem Nat.bit_val (b : Bool) (n : Nat) :
            Nat.bit b n = 2 * n + b.toNat
            @[simp]
            theorem Nat.bit_div_two (b : Bool) (n : Nat) :
            Nat.bit b n / 2 = n
            @[simp]
            theorem Nat.bit_mod_two (b : Bool) (n : Nat) :
            Nat.bit b n % 2 = b.toNat
            @[simp]
            theorem Nat.bit_shiftRight_one (b : Bool) (n : Nat) :
            Nat.bit b n >>> 1 = n
            theorem Nat.testBit_bit_zero (b : Bool) (n : Nat) :
            (Nat.bit b n).testBit 0 = b
            @[simp]
            theorem Nat.bitCasesOn_bit {motive : NatSort u} (h : (b : Bool) → (n : Nat) → motive (Nat.bit b n)) (b : Bool) (n : Nat) :
            Nat.bitCasesOn (Nat.bit b n) h = h b n
            @[simp]
            theorem Nat.binaryRec_zero {motive : NatSort u} (z : motive 0) (f : (b : Bool) → (n : Nat) → motive nmotive (Nat.bit b n)) :
            @[simp]
            theorem Nat.binaryRec_one {motive : NatSort u} (z : motive 0) (f : (b : Bool) → (n : Nat) → motive nmotive (Nat.bit b n)) :
            Nat.binaryRec z f 1 = f true 0 z
            theorem Nat.binaryRec_eq {motive : NatSort u} {z : motive 0} {f : (b : Bool) → (n : Nat) → motive nmotive (Nat.bit b n)} (b : Bool) (n : Nat) (h : f false 0 z = z (n = 0b = true)) :
            Nat.binaryRec z f (Nat.bit b n) = f b n (Nat.binaryRec z f n)
            @[deprecated Nat.binaryRec_eq]
            theorem Nat.binaryRec_eq' {motive : NatSort u} {z : motive 0} {f : (b : Bool) → (n : Nat) → motive nmotive (Nat.bit b n)} (b : Bool) (n : Nat) (h : f false 0 z = z (n = 0b = true)) :
            Nat.binaryRec z f (Nat.bit b n) = f b n (Nat.binaryRec z f n)

            Alias of Nat.binaryRec_eq.