Documentation

Init.Data.SInt.Basic

This module contains the definition of signed fixed width integer types as well as basic arithmetic and bitwise operations on top of it.

structure Int8 :

The type of signed 8-bit integers. This type has special support in the compiler to make it actually 8 bits rather than wrapping a Nat.

  • toUInt8 : UInt8

    Obtain the UInt8 that is 2's complement equivalent to the Int8.

Instances For
    structure Int16 :

    The type of signed 16-bit integers. This type has special support in the compiler to make it actually 16 bits rather than wrapping a Nat.

    Instances For
      structure Int32 :

      The type of signed 32-bit integers. This type has special support in the compiler to make it actually 32 bits rather than wrapping a Nat.

      Instances For
        structure Int64 :

        The type of signed 64-bit integers. This type has special support in the compiler to make it actually 64 bits rather than wrapping a Nat.

        Instances For
          structure ISize :

          A ISize is a signed integer with the size of a word for the platform's architecture.

          For example, if running on a 32-bit machine, ISize is equivalent to Int32. Or on a 64-bit machine, Int64.

          • toUSize : USize

            Obtain the USize that is 2's complement equivalent to the ISize.

          Instances For
            @[reducible, inline]
            abbrev Int8.size :

            The size of type Int8, that is, 2^8 = 256.

            Equations
            Instances For
              @[inline]

              Obtain the BitVec that contains the 2's complement representation of the Int8.

              Equations
              • x.toBitVec = x.toUInt8.toBitVec
              Instances For
                @[extern lean_int8_of_int]
                def Int8.ofInt (i : Int) :
                Instances For
                  @[extern lean_int8_of_nat]
                  def Int8.ofNat (n : Nat) :
                  Equations
                  Instances For
                    @[reducible, inline]
                    abbrev Int.toInt8 (i : Int) :
                    Equations
                    Instances For
                      @[reducible, inline]
                      abbrev Nat.toInt8 (n : Nat) :
                      Equations
                      Instances For
                        @[extern lean_int8_to_int]
                        def Int8.toInt (i : Int8) :
                        Instances For
                          @[inline]
                          def Int8.toNat (i : Int8) :

                          This function has the same behavior as Int.toNat for negative numbers. If you want to obtain the 2's complement representation use toBitVec.

                          Instances For
                            @[extern lean_int8_neg]
                            def Int8.neg (i : Int8) :
                            Instances For
                              Equations
                              instance instOfNatInt8 {n : Nat} :
                              @[extern lean_int8_add]
                              def Int8.add (a b : Int8) :
                              Instances For
                                @[extern lean_int8_sub]
                                def Int8.sub (a b : Int8) :
                                Instances For
                                  @[extern lean_int8_mul]
                                  def Int8.mul (a b : Int8) :
                                  Instances For
                                    @[extern lean_int8_div]
                                    def Int8.div (a b : Int8) :
                                    Instances For
                                      @[extern lean_int8_mod]
                                      def Int8.mod (a b : Int8) :
                                      Instances For
                                        @[extern lean_int8_land]
                                        def Int8.land (a b : Int8) :
                                        Equations
                                        • a.land b = { toUInt8 := { toBitVec := a.toBitVec &&& b.toBitVec } }
                                        Instances For
                                          @[extern lean_int8_lor]
                                          def Int8.lor (a b : Int8) :
                                          Instances For
                                            @[extern lean_int8_xor]
                                            def Int8.xor (a b : Int8) :
                                            Instances For
                                              @[extern lean_int8_shift_left]
                                              def Int8.shiftLeft (a b : Int8) :
                                              Equations
                                              • a.shiftLeft b = { toUInt8 := { toBitVec := a.toBitVec <<< b.toBitVec.smod 8 } }
                                              Instances For
                                                @[extern lean_int8_shift_right]
                                                Equations
                                                • a.shiftRight b = { toUInt8 := { toBitVec := a.toBitVec.sshiftRight' (b.toBitVec.smod 8) } }
                                                Instances For
                                                  @[extern lean_int8_complement]
                                                  Instances For
                                                    @[extern lean_int8_dec_eq]
                                                    def Int8.decEq (a b : Int8) :
                                                    Decidable (a = b)
                                                    Equations
                                                    • { toUInt8 := n }.decEq { toUInt8 := m } = if h : n = m then isTrue else isFalse
                                                    Instances For
                                                      def Int8.lt (a b : Int8) :
                                                      Instances For
                                                        def Int8.le (a b : Int8) :
                                                        Equations
                                                        • a.le b = (a.toBitVec.sle b.toBitVec = true)
                                                        Instances For
                                                          instance instAddInt8 :
                                                          Equations
                                                          instance instSubInt8 :
                                                          Equations
                                                          instance instMulInt8 :
                                                          Equations
                                                          instance instModInt8 :
                                                          Equations
                                                          instance instDivInt8 :
                                                          Equations
                                                          instance instLTInt8 :
                                                          Equations
                                                          Equations
                                                          Equations
                                                          @[extern lean_bool_to_int8]
                                                          def Bool.toInt8 (b : Bool) :
                                                          Instances For
                                                            @[extern lean_int8_dec_lt]
                                                            def Int8.decLt (a b : Int8) :
                                                            Decidable (a < b)
                                                            Equations
                                                            Instances For
                                                              @[extern lean_int8_dec_le]
                                                              def Int8.decLe (a b : Int8) :
                                                              Instances For
                                                                instance instDecidableLtInt8 (a b : Int8) :
                                                                Decidable (a < b)
                                                                Equations
                                                                instance instMaxInt8 :
                                                                Equations
                                                                instance instMinInt8 :
                                                                Equations
                                                                @[reducible, inline]
                                                                abbrev Int16.size :

                                                                The size of type Int16, that is, 2^16 = 65536.

                                                                Equations
                                                                Instances For
                                                                  @[inline]

                                                                  Obtain the BitVec that contains the 2's complement representation of the Int16.

                                                                  Instances For
                                                                    @[extern lean_int16_of_int]
                                                                    def Int16.ofInt (i : Int) :
                                                                    Instances For
                                                                      @[extern lean_int16_of_nat]
                                                                      def Int16.ofNat (n : Nat) :
                                                                      Instances For
                                                                        @[reducible, inline]
                                                                        abbrev Int.toInt16 (i : Int) :
                                                                        Equations
                                                                        Instances For
                                                                          @[reducible, inline]
                                                                          abbrev Nat.toInt16 (n : Nat) :
                                                                          Equations
                                                                          Instances For
                                                                            @[extern lean_int16_to_int]
                                                                            def Int16.toInt (i : Int16) :
                                                                            Equations
                                                                            • i.toInt = i.toBitVec.toInt
                                                                            Instances For
                                                                              @[inline]
                                                                              def Int16.toNat (i : Int16) :

                                                                              This function has the same behavior as Int.toNat for negative numbers. If you want to obtain the 2's complement representation use toBitVec.

                                                                              Equations
                                                                              • i.toNat = i.toInt.toNat
                                                                              Instances For
                                                                                @[extern lean_int16_to_int8]
                                                                                Instances For
                                                                                  @[extern lean_int8_to_int16]
                                                                                  Instances For
                                                                                    @[extern lean_int16_neg]
                                                                                    def Int16.neg (i : Int16) :
                                                                                    Equations
                                                                                    • i.neg = { toUInt16 := { toBitVec := -i.toBitVec } }
                                                                                    Instances For
                                                                                      instance instOfNatInt16 {n : Nat} :
                                                                                      Equations
                                                                                      @[extern lean_int16_add]
                                                                                      def Int16.add (a b : Int16) :
                                                                                      Instances For
                                                                                        @[extern lean_int16_sub]
                                                                                        def Int16.sub (a b : Int16) :
                                                                                        Equations
                                                                                        • a.sub b = { toUInt16 := { toBitVec := a.toBitVec - b.toBitVec } }
                                                                                        Instances For
                                                                                          @[extern lean_int16_mul]
                                                                                          def Int16.mul (a b : Int16) :
                                                                                          Equations
                                                                                          • a.mul b = { toUInt16 := { toBitVec := a.toBitVec * b.toBitVec } }
                                                                                          Instances For
                                                                                            @[extern lean_int16_div]
                                                                                            def Int16.div (a b : Int16) :
                                                                                            Instances For
                                                                                              @[extern lean_int16_mod]
                                                                                              def Int16.mod (a b : Int16) :
                                                                                              Equations
                                                                                              • a.mod b = { toUInt16 := { toBitVec := a.toBitVec.srem b.toBitVec } }
                                                                                              Instances For
                                                                                                @[extern lean_int16_land]
                                                                                                def Int16.land (a b : Int16) :
                                                                                                Equations
                                                                                                • a.land b = { toUInt16 := { toBitVec := a.toBitVec &&& b.toBitVec } }
                                                                                                Instances For
                                                                                                  @[extern lean_int16_lor]
                                                                                                  def Int16.lor (a b : Int16) :
                                                                                                  Instances For
                                                                                                    @[extern lean_int16_xor]
                                                                                                    def Int16.xor (a b : Int16) :
                                                                                                    Instances For
                                                                                                      @[extern lean_int16_shift_left]
                                                                                                      Instances For
                                                                                                        @[extern lean_int16_shift_right]
                                                                                                        Equations
                                                                                                        • a.shiftRight b = { toUInt16 := { toBitVec := a.toBitVec.sshiftRight' (b.toBitVec.smod 16) } }
                                                                                                        Instances For
                                                                                                          @[extern lean_int16_complement]
                                                                                                          Equations
                                                                                                          • a.complement = { toUInt16 := { toBitVec := ~~~a.toBitVec } }
                                                                                                          Instances For
                                                                                                            @[extern lean_int16_dec_eq]
                                                                                                            def Int16.decEq (a b : Int16) :
                                                                                                            Decidable (a = b)
                                                                                                            Instances For
                                                                                                              def Int16.lt (a b : Int16) :
                                                                                                              Equations
                                                                                                              • a.lt b = (a.toBitVec.slt b.toBitVec = true)
                                                                                                              Instances For
                                                                                                                def Int16.le (a b : Int16) :
                                                                                                                Instances For
                                                                                                                  Equations
                                                                                                                  Equations
                                                                                                                  Equations
                                                                                                                  instance instLEInt16 :
                                                                                                                  Equations
                                                                                                                  Equations
                                                                                                                  @[extern lean_bool_to_int16]
                                                                                                                  Instances For
                                                                                                                    @[extern lean_int16_dec_lt]
                                                                                                                    def Int16.decLt (a b : Int16) :
                                                                                                                    Decidable (a < b)
                                                                                                                    Instances For
                                                                                                                      @[extern lean_int16_dec_le]
                                                                                                                      def Int16.decLe (a b : Int16) :
                                                                                                                      Equations
                                                                                                                      Instances For
                                                                                                                        instance instDecidableLeInt16 (a b : Int16) :
                                                                                                                        Equations
                                                                                                                        @[reducible, inline]
                                                                                                                        abbrev Int32.size :

                                                                                                                        The size of type Int32, that is, 2^32 = 4294967296.

                                                                                                                        Equations
                                                                                                                        Instances For
                                                                                                                          @[inline]

                                                                                                                          Obtain the BitVec that contains the 2's complement representation of the Int32.

                                                                                                                          Instances For
                                                                                                                            @[extern lean_int32_of_int]
                                                                                                                            def Int32.ofInt (i : Int) :
                                                                                                                            Instances For
                                                                                                                              @[extern lean_int32_of_nat]
                                                                                                                              def Int32.ofNat (n : Nat) :
                                                                                                                              Instances For
                                                                                                                                @[reducible, inline]
                                                                                                                                abbrev Int.toInt32 (i : Int) :
                                                                                                                                Instances For
                                                                                                                                  @[reducible, inline]
                                                                                                                                  abbrev Nat.toInt32 (n : Nat) :
                                                                                                                                  Equations
                                                                                                                                  Instances For
                                                                                                                                    @[extern lean_int32_to_int]
                                                                                                                                    def Int32.toInt (i : Int32) :
                                                                                                                                    Instances For
                                                                                                                                      @[inline]
                                                                                                                                      def Int32.toNat (i : Int32) :

                                                                                                                                      This function has the same behavior as Int.toNat for negative numbers. If you want to obtain the 2's complement representation use toBitVec.

                                                                                                                                      Instances For
                                                                                                                                        @[extern lean_int32_to_int8]
                                                                                                                                        Equations
                                                                                                                                        Instances For
                                                                                                                                          @[extern lean_int32_to_int16]
                                                                                                                                          Equations
                                                                                                                                          Instances For
                                                                                                                                            @[extern lean_int8_to_int32]
                                                                                                                                            Equations
                                                                                                                                            Instances For
                                                                                                                                              @[extern lean_int16_to_int32]
                                                                                                                                              Equations
                                                                                                                                              Instances For
                                                                                                                                                @[extern lean_int32_neg]
                                                                                                                                                def Int32.neg (i : Int32) :
                                                                                                                                                Equations
                                                                                                                                                • i.neg = { toUInt32 := { toBitVec := -i.toBitVec } }
                                                                                                                                                Instances For
                                                                                                                                                  instance instOfNatInt32 {n : Nat} :
                                                                                                                                                  Equations
                                                                                                                                                  Equations
                                                                                                                                                  @[extern lean_int32_add]
                                                                                                                                                  def Int32.add (a b : Int32) :
                                                                                                                                                  Equations
                                                                                                                                                  • a.add b = { toUInt32 := { toBitVec := a.toBitVec + b.toBitVec } }
                                                                                                                                                  Instances For
                                                                                                                                                    @[extern lean_int32_sub]
                                                                                                                                                    def Int32.sub (a b : Int32) :
                                                                                                                                                    Equations
                                                                                                                                                    • a.sub b = { toUInt32 := { toBitVec := a.toBitVec - b.toBitVec } }
                                                                                                                                                    Instances For
                                                                                                                                                      @[extern lean_int32_mul]
                                                                                                                                                      def Int32.mul (a b : Int32) :
                                                                                                                                                      Instances For
                                                                                                                                                        @[extern lean_int32_div]
                                                                                                                                                        def Int32.div (a b : Int32) :
                                                                                                                                                        Equations
                                                                                                                                                        • a.div b = { toUInt32 := { toBitVec := a.toBitVec.sdiv b.toBitVec } }
                                                                                                                                                        Instances For
                                                                                                                                                          @[extern lean_int32_mod]
                                                                                                                                                          def Int32.mod (a b : Int32) :
                                                                                                                                                          Instances For
                                                                                                                                                            @[extern lean_int32_land]
                                                                                                                                                            def Int32.land (a b : Int32) :
                                                                                                                                                            Instances For
                                                                                                                                                              @[extern lean_int32_lor]
                                                                                                                                                              def Int32.lor (a b : Int32) :
                                                                                                                                                              Instances For
                                                                                                                                                                @[extern lean_int32_xor]
                                                                                                                                                                def Int32.xor (a b : Int32) :
                                                                                                                                                                Instances For
                                                                                                                                                                  @[extern lean_int32_shift_left]
                                                                                                                                                                  Equations
                                                                                                                                                                  • a.shiftLeft b = { toUInt32 := { toBitVec := a.toBitVec <<< b.toBitVec.smod 32 } }
                                                                                                                                                                  Instances For
                                                                                                                                                                    @[extern lean_int32_shift_right]
                                                                                                                                                                    Equations
                                                                                                                                                                    • a.shiftRight b = { toUInt32 := { toBitVec := a.toBitVec.sshiftRight' (b.toBitVec.smod 32) } }
                                                                                                                                                                    Instances For
                                                                                                                                                                      @[extern lean_int32_complement]
                                                                                                                                                                      Instances For
                                                                                                                                                                        @[extern lean_int32_dec_eq]
                                                                                                                                                                        def Int32.decEq (a b : Int32) :
                                                                                                                                                                        Decidable (a = b)
                                                                                                                                                                        Instances For
                                                                                                                                                                          def Int32.lt (a b : Int32) :
                                                                                                                                                                          Equations
                                                                                                                                                                          • a.lt b = (a.toBitVec.slt b.toBitVec = true)
                                                                                                                                                                          Instances For
                                                                                                                                                                            def Int32.le (a b : Int32) :
                                                                                                                                                                            Equations
                                                                                                                                                                            • a.le b = (a.toBitVec.sle b.toBitVec = true)
                                                                                                                                                                            Instances For
                                                                                                                                                                              Equations
                                                                                                                                                                              Equations
                                                                                                                                                                              Equations
                                                                                                                                                                              Equations
                                                                                                                                                                              @[extern lean_bool_to_int32]
                                                                                                                                                                              Equations
                                                                                                                                                                              • b.toInt32 = if b = true then 1 else 0
                                                                                                                                                                              Instances For
                                                                                                                                                                                @[extern lean_int32_dec_lt]
                                                                                                                                                                                def Int32.decLt (a b : Int32) :
                                                                                                                                                                                Decidable (a < b)
                                                                                                                                                                                Equations
                                                                                                                                                                                Instances For
                                                                                                                                                                                  @[extern lean_int32_dec_le]
                                                                                                                                                                                  def Int32.decLe (a b : Int32) :
                                                                                                                                                                                  Instances For
                                                                                                                                                                                    instance instDecidableLeInt32 (a b : Int32) :
                                                                                                                                                                                    Equations
                                                                                                                                                                                    Equations
                                                                                                                                                                                    @[reducible, inline]
                                                                                                                                                                                    abbrev Int64.size :

                                                                                                                                                                                    The size of type Int64, that is, 2^64 = 18446744073709551616.

                                                                                                                                                                                    Equations
                                                                                                                                                                                    Instances For
                                                                                                                                                                                      @[inline]

                                                                                                                                                                                      Obtain the BitVec that contains the 2's complement representation of the Int64.

                                                                                                                                                                                      Equations
                                                                                                                                                                                      • x.toBitVec = x.toUInt64.toBitVec
                                                                                                                                                                                      Instances For
                                                                                                                                                                                        @[extern lean_int64_of_int]
                                                                                                                                                                                        def Int64.ofInt (i : Int) :
                                                                                                                                                                                        Equations
                                                                                                                                                                                        Instances For
                                                                                                                                                                                          @[extern lean_int64_of_nat]
                                                                                                                                                                                          def Int64.ofNat (n : Nat) :
                                                                                                                                                                                          Equations
                                                                                                                                                                                          Instances For
                                                                                                                                                                                            @[reducible, inline]
                                                                                                                                                                                            abbrev Int.toInt64 (i : Int) :
                                                                                                                                                                                            Equations
                                                                                                                                                                                            Instances For
                                                                                                                                                                                              @[reducible, inline]
                                                                                                                                                                                              abbrev Nat.toInt64 (n : Nat) :
                                                                                                                                                                                              Instances For
                                                                                                                                                                                                @[extern lean_int64_to_int_sint]
                                                                                                                                                                                                def Int64.toInt (i : Int64) :
                                                                                                                                                                                                Instances For
                                                                                                                                                                                                  @[inline]
                                                                                                                                                                                                  def Int64.toNat (i : Int64) :

                                                                                                                                                                                                  This function has the same behavior as Int.toNat for negative numbers. If you want to obtain the 2's complement representation use toBitVec.

                                                                                                                                                                                                  Equations
                                                                                                                                                                                                  • i.toNat = i.toInt.toNat
                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                    @[extern lean_int64_to_int8]
                                                                                                                                                                                                    Equations
                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                      @[extern lean_int64_to_int16]
                                                                                                                                                                                                      Equations
                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                        @[extern lean_int64_to_int32]
                                                                                                                                                                                                        Equations
                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                          @[extern lean_int8_to_int64]
                                                                                                                                                                                                          Equations
                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                            @[extern lean_int16_to_int64]
                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                              @[extern lean_int32_to_int64]
                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                @[extern lean_int64_neg]
                                                                                                                                                                                                                def Int64.neg (i : Int64) :
                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                • i.neg = { toUInt64 := { toBitVec := -i.toBitVec } }
                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                  @[extern lean_int64_add]
                                                                                                                                                                                                                  def Int64.add (a b : Int64) :
                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                  • a.add b = { toUInt64 := { toBitVec := a.toBitVec + b.toBitVec } }
                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                    @[extern lean_int64_sub]
                                                                                                                                                                                                                    def Int64.sub (a b : Int64) :
                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                    • a.sub b = { toUInt64 := { toBitVec := a.toBitVec - b.toBitVec } }
                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                      @[extern lean_int64_mul]
                                                                                                                                                                                                                      def Int64.mul (a b : Int64) :
                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                      • a.mul b = { toUInt64 := { toBitVec := a.toBitVec * b.toBitVec } }
                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                        @[extern lean_int64_div]
                                                                                                                                                                                                                        def Int64.div (a b : Int64) :
                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                          @[extern lean_int64_mod]
                                                                                                                                                                                                                          def Int64.mod (a b : Int64) :
                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                          • a.mod b = { toUInt64 := { toBitVec := a.toBitVec.srem b.toBitVec } }
                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                            @[extern lean_int64_land]
                                                                                                                                                                                                                            def Int64.land (a b : Int64) :
                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                              @[extern lean_int64_lor]
                                                                                                                                                                                                                              def Int64.lor (a b : Int64) :
                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                @[extern lean_int64_xor]
                                                                                                                                                                                                                                def Int64.xor (a b : Int64) :
                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                  @[extern lean_int64_shift_left]
                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                  • a.shiftLeft b = { toUInt64 := { toBitVec := a.toBitVec <<< b.toBitVec.smod 64 } }
                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                    @[extern lean_int64_shift_right]
                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                      @[extern lean_int64_complement]
                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                        @[extern lean_int64_dec_eq]
                                                                                                                                                                                                                                        def Int64.decEq (a b : Int64) :
                                                                                                                                                                                                                                        Decidable (a = b)
                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                          def Int64.lt (a b : Int64) :
                                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                                            def Int64.le (a b : Int64) :
                                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                              instance instLEInt64 :
                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                              @[extern lean_bool_to_int64]
                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                              • b.toInt64 = if b = true then 1 else 0
                                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                                @[extern lean_int64_dec_lt]
                                                                                                                                                                                                                                                def Int64.decLt (a b : Int64) :
                                                                                                                                                                                                                                                Decidable (a < b)
                                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                                  @[extern lean_int64_dec_le]
                                                                                                                                                                                                                                                  def Int64.decLe (a b : Int64) :
                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                    instance instDecidableLtInt64 (a b : Int64) :
                                                                                                                                                                                                                                                    Decidable (a < b)
                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                    @[reducible, inline]
                                                                                                                                                                                                                                                    abbrev ISize.size :

                                                                                                                                                                                                                                                    The size of type ISize, that is, 2^System.Platform.numBits.

                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                                      @[inline]

                                                                                                                                                                                                                                                      Obtain the BitVec that contains the 2's complement representation of the ISize.

                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                      • x.toBitVec = x.toUSize.toBitVec
                                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                                        @[extern lean_isize_of_int]
                                                                                                                                                                                                                                                        def ISize.ofInt (i : Int) :
                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                          @[extern lean_isize_of_nat]
                                                                                                                                                                                                                                                          def ISize.ofNat (n : Nat) :
                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                                                            @[reducible, inline]
                                                                                                                                                                                                                                                            abbrev Int.toISize (i : Int) :
                                                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                                                              @[reducible, inline]
                                                                                                                                                                                                                                                              abbrev Nat.toISize (n : Nat) :
                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                                                @[extern lean_isize_to_int]
                                                                                                                                                                                                                                                                def ISize.toInt (i : ISize) :
                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                • i.toInt = i.toBitVec.toInt
                                                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                                                  @[inline]
                                                                                                                                                                                                                                                                  def ISize.toNat (i : ISize) :

                                                                                                                                                                                                                                                                  This function has the same behavior as Int.toNat for negative numbers. If you want to obtain the 2's complement representation use toBitVec.

                                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                                  • i.toNat = i.toInt.toNat
                                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                                    @[extern lean_isize_to_int32]
                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                                                      @[extern lean_isize_to_int64]

                                                                                                                                                                                                                                                                      Upcast ISize to Int64. This function is losless as ISize is either Int32 or Int64.

                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                                                        @[extern lean_int32_to_isize]

                                                                                                                                                                                                                                                                        Upcast Int32 to ISize. This function is losless as ISize is either Int32 or Int64.

                                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                                          @[extern lean_int64_to_isize]
                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                                                                            @[extern lean_isize_neg]
                                                                                                                                                                                                                                                                            def ISize.neg (i : ISize) :
                                                                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                                                                              instance instOfNatISize {n : Nat} :
                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                              @[extern lean_isize_add]
                                                                                                                                                                                                                                                                              def ISize.add (a b : ISize) :
                                                                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                                                                @[extern lean_isize_sub]
                                                                                                                                                                                                                                                                                def ISize.sub (a b : ISize) :
                                                                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                                                                  @[extern lean_isize_mul]
                                                                                                                                                                                                                                                                                  def ISize.mul (a b : ISize) :
                                                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                                                  • a.mul b = { toUSize := { toBitVec := a.toBitVec * b.toBitVec } }
                                                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                                                    @[extern lean_isize_div]
                                                                                                                                                                                                                                                                                    def ISize.div (a b : ISize) :
                                                                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                                                                      @[extern lean_isize_mod]
                                                                                                                                                                                                                                                                                      def ISize.mod (a b : ISize) :
                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                      • a.mod b = { toUSize := { toBitVec := a.toBitVec.srem b.toBitVec } }
                                                                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                                                                        @[extern lean_isize_land]
                                                                                                                                                                                                                                                                                        def ISize.land (a b : ISize) :
                                                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                                                          @[extern lean_isize_lor]
                                                                                                                                                                                                                                                                                          def ISize.lor (a b : ISize) :
                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                          • a.lor b = { toUSize := { toBitVec := a.toBitVec ||| b.toBitVec } }
                                                                                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                                                                                            @[extern lean_isize_xor]
                                                                                                                                                                                                                                                                                            def ISize.xor (a b : ISize) :
                                                                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                                                                            • a.xor b = { toUSize := { toBitVec := a.toBitVec ^^^ b.toBitVec } }
                                                                                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                                                                                              @[extern lean_isize_shift_left]
                                                                                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                                                                                @[extern lean_isize_shift_right]
                                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                                                                                  @[extern lean_isize_complement]
                                                                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                                                                  • a.complement = { toUSize := { toBitVec := ~~~a.toBitVec } }
                                                                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                                                                    @[extern lean_isize_dec_eq]
                                                                                                                                                                                                                                                                                                    def ISize.decEq (a b : ISize) :
                                                                                                                                                                                                                                                                                                    Decidable (a = b)
                                                                                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                                                                                      def ISize.lt (a b : ISize) :
                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                      • a.lt b = (a.toBitVec.slt b.toBitVec = true)
                                                                                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                                                                                        def ISize.le (a b : ISize) :
                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                        • a.le b = (a.toBitVec.sle b.toBitVec = true)
                                                                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                          instance instLTISize :
                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                          @[extern lean_bool_to_isize]
                                                                                                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                                                                                                            @[extern lean_isize_dec_lt]
                                                                                                                                                                                                                                                                                                            def ISize.decLt (a b : ISize) :
                                                                                                                                                                                                                                                                                                            Decidable (a < b)
                                                                                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                                                                                                              @[extern lean_isize_dec_le]
                                                                                                                                                                                                                                                                                                              def ISize.decLe (a b : ISize) :
                                                                                                                                                                                                                                                                                                              Instances For