Rename rotation constructors

This commit is contained in:
Kiana Sheibani 2022-10-21 10:03:43 -04:00
parent 342cff97dd
commit 1ad4c1f13c
Signed by: toki
GPG key ID: 6CB106C25E86A9F7
13 changed files with 108 additions and 94 deletions

View file

@ -273,8 +273,6 @@ export %inline
(!!) : Array s a -> Coords s -> a (!!) : Array s a -> Coords s -> a
arr !! is = index is arr arr !! is = index is arr
-- TODO: Create set/update at index functions
||| Update the entry at the given coordinates using the function. ||| Update the entry at the given coordinates using the function.
export export
indexUpdate : Coords s -> (a -> a) -> Array s a -> Array s a indexUpdate : Coords s -> (a -> a) -> Array s a -> Array s a
@ -452,7 +450,7 @@ resize s' def arr = fromFunction' s' (getOrder arr) (fromMaybe def . (arr !?) .
||| |||
||| @ s' The shape to resize the array to ||| @ s' The shape to resize the array to
export export
-- TODO: Come up with a solution that doesn't use `believe_me` or trip over some -- HACK: Come up with a solution that doesn't use `believe_me` or trip over some
-- weird bug in the type-checker -- weird bug in the type-checker
resizeLTE : (s' : Vect rk Nat) -> (0 ok : NP Prelude.id (zipWith LTE s' s)) => resizeLTE : (s' : Vect rk Nat) -> (0 ok : NP Prelude.id (zipWith LTE s' s)) =>
Array {rk} s a -> Array s' a Array {rk} s a -> Array s' a

View file

@ -62,7 +62,7 @@ hvectorLinear v = rewrite plusCommutative 1 n in vector (v ++ [0])
export export
fromHomogeneous : HVector n a -> Vector n a fromHomogeneous : HVector n a -> Vector n a
fromHomogeneous = vector . init . toVect fromHomogeneous = vector . init . toVect
-- TODO: Find an implementation for `fromHomogeneous` that doesn't suck -- HACK: Find an implementation for `fromHomogeneous` that doesn't suck
export export
@ -116,17 +116,17 @@ translationH {n} v with (viewShape v)
_ | Shape [n] = hmatrix identity v _ | Shape [n] = hmatrix identity v
export export
rotation2DH : Double -> HMatrix' 2 Double rotate2DH : Double -> HMatrix' 2 Double
rotation2DH = matrixToH . rotation2D rotate2DH = matrixToH . rotate2D
export export
rotation3DXH : Double -> HMatrix' 3 Double rotate3DXH : Double -> HMatrix' 3 Double
rotation3DXH = matrixToH . rotation3DX rotate3DXH = matrixToH . rotate3DX
export export
rotation3DYH : Double -> HMatrix' 3 Double rotate3DYH : Double -> HMatrix' 3 Double
rotation3DYH = matrixToH . rotation3DY rotate3DYH = matrixToH . rotate3DY
export export
rotation3DZH : Double -> HMatrix' 3 Double rotate3DZH : Double -> HMatrix' 3 Double
rotation3DZH = matrixToH . rotation3DZ rotate3DZH = matrixToH . rotate3DZ

View file

@ -71,26 +71,26 @@ permuteM p = permuteInAxis 0 p (repeatDiag 1 0)
||| Construct the matrix that scales a vector by the given value. ||| Construct the matrix that scales a vector by the given value.
export export
scaling : {n : _} -> Num a => a -> Matrix' n a scale : {n : _} -> Num a => a -> Matrix' n a
scaling x = repeatDiag x 0 scale x = repeatDiag x 0
||| Calculate the rotation matrix of an angle. ||| Calculate the rotation matrix of an angle.
export export
rotation2D : Double -> Matrix' 2 Double rotate2D : Double -> Matrix' 2 Double
rotation2D a = matrix [[cos a, - sin a], [sin a, cos a]] rotate2D a = matrix [[cos a, - sin a], [sin a, cos a]]
export export
rotation3DX : Double -> Matrix' 3 Double rotate3DX : Double -> Matrix' 3 Double
rotation3DX a = matrix [[1,0,0], [0, cos a, - sin a], [0, sin a, cos a]] rotate3DX a = matrix [[1,0,0], [0, cos a, - sin a], [0, sin a, cos a]]
export export
rotation3DY : Double -> Matrix' 3 Double rotate3DY : Double -> Matrix' 3 Double
rotation3DY a = matrix [[cos a, 0, sin a], [0,1,0], [- sin a, 0, cos a]] rotate3DY a = matrix [[cos a, 0, sin a], [0,1,0], [- sin a, 0, cos a]]
export export
rotation3DZ : Double -> Matrix' 3 Double rotate3DZ : Double -> Matrix' 3 Double
rotation3DZ a = matrix [[cos a, - sin a, 0], [sin a, cos a, 0], [0,0,1]] rotate3DZ a = matrix [[cos a, - sin a, 0], [sin a, cos a, 0], [0,0,1]]
-------------------------------------------------------------------------------- --------------------------------------------------------------------------------

View file

@ -16,8 +16,12 @@ public export
Affine : Nat -> Type -> Type Affine : Nat -> Type -> Type
Affine = Transform TAffine Affine = Transform TAffine
export
isAffine : FieldCmp a => HMatrix' n a -> Bool
isAffine mat = isHMatrix mat && invertible (getMatrix mat)
export export
fromHMatrix : FieldCmp a => HMatrix' n a -> Maybe (Affine n a) fromHMatrix : FieldCmp a => HMatrix' n a -> Maybe (Affine n a)
fromHMatrix mat = if invertible (getMatrix mat) fromHMatrix mat = if isAffine mat
then Just (unsafeMkTrans mat) then Just (unsafeMkTrans mat)
else Nothing else Nothing

View file

@ -15,3 +15,6 @@ import Data.NumIdr.Transform.Transform
public export public export
Isometry : Nat -> Type -> Type Isometry : Nat -> Type -> Type
Isometry = Transform TIsometry Isometry = Transform TIsometry
-- TODO: Add Isometry constructors

View file

@ -18,13 +18,13 @@ Orthonormal = Transform TOrthonormal
export export
isOrthonormal : Eq a => Num a => Matrix' n a -> Bool isOrthonormal' : Eq a => Num a => Matrix' n a -> Bool
isOrthonormal {n} mat with (viewShape mat) isOrthonormal' {n} mat with (viewShape mat)
_ | Shape [n,n] = identity == fromFunction [n,n] (\[i,j] => getColumn i mat `dot` getColumn j mat) _ | Shape [n,n] = identity == fromFunction [n,n] (\[i,j] => getColumn i mat `dot` getColumn j mat)
export export
fromMatrix : Eq a => Num a => Matrix' n a -> Maybe (Orthonormal n a) fromMatrix : Eq a => Num a => Matrix' n a -> Maybe (Orthonormal n a)
fromMatrix mat = if isOrthonormal mat then Just (unsafeMkTrans (matrixToH mat)) fromMatrix mat = if isOrthonormal' mat then Just (unsafeMkTrans (matrixToH mat))
else Nothing else Nothing
@ -48,4 +48,4 @@ reflectZ = reflect 2
export export
reflectNormal : (Neg a, Fractional a) => Vector n a -> Orthonormal n a reflectNormal : (Neg a, Fractional a) => Vector n a -> Orthonormal n a
reflectNormal {n} v with (viewShape v) reflectNormal {n} v with (viewShape v)
_ | Shape [n] = unsafeMkTrans $ matrixToH $ (identity - (2 / normSq v) *. outer v v) _ | Shape [n] = unsafeMkTrans $ matrixToH $ identity - (2 / normSq v) *. outer v v

View file

@ -90,24 +90,26 @@ export
-- Arithmetic operations -- Arithmetic operations
-------------------------------------------------------------------------------- --------------------------------------------------------------------------------
infixr 8 +.
infixl 8 .+
infixl 8 -.
-- Affine space operations -- Affine space operations
-- These seem to cause issues with interface resolution -- These would have been named simply (+) and (-), but that caused
-- too many issues with interface resolution.
-- namespace Left export
-- export (+.) : Num a => Vector n a -> Point n a -> Point n a
-- (+) : Num a => Vector n a -> Point n a -> Point n a a +. MkPoint b = MkPoint (zipWith (+) a b)
-- a + MkPoint b = MkPoint (zipWith (+) a b)
-- namespace Right export
-- export (.+) : Num a => Point n a -> Vector n a -> Point n a
-- (+) : Num a => Point n a -> Vector n a -> Point n a MkPoint a .+ b = MkPoint (zipWith (+) a b)
-- MkPoint a + b = MkPoint (zipWith (+) a b)
-- export export
-- (-) : Neg a => Point n a -> Point n a -> Vector n a (-.) : Neg a => Point n a -> Point n a -> Vector n a
-- MkPoint a - MkPoint b = zipWith (-) a b MkPoint a -. MkPoint b = zipWith (-) a b
-------------------------------------------------------------------------------- --------------------------------------------------------------------------------
@ -150,8 +152,8 @@ Traversable (Point n) where
export export
Show a => Show (Point n a) where Show a => Show (Point n a) where
showPrec d (MkPoint v) = showCon d "point " $ showPrec d (MkPoint v) = showCon d "point" $
show $ PrimArray.toList $ getPrim v showArg $ PrimArray.toList $ getPrim v
export export
Cast a b => Cast (Point n a) (Point n b) where Cast a b => Cast (Point n a) (Point n b) where

View file

@ -15,3 +15,5 @@ import Data.NumIdr.Transform.Transform
public export public export
Rigid : Nat -> Type -> Type Rigid : Nat -> Type -> Type
Rigid = Transform TRigid Rigid = Transform TRigid
-- TODO: Add Rigid constructors

View file

@ -8,6 +8,7 @@ import Data.NumIdr.Matrix
import Data.NumIdr.Homogeneous import Data.NumIdr.Homogeneous
import Data.NumIdr.Transform.Point import Data.NumIdr.Transform.Point
import Data.NumIdr.Transform.Transform import Data.NumIdr.Transform.Transform
import Data.NumIdr.Transform.Orthonormal
%default total %default total
@ -17,6 +18,27 @@ Rotation : Nat -> Type -> Type
Rotation = Transform TRotation Rotation = Transform TRotation
-- HACK: Replace with more efficient method
export export
isRotation' : Matrix' n a -> Bool isRotation' : FieldCmp a => Matrix' n a -> Bool
isRotation' mat = isRotation' mat = isOrthonormal mat && det mat == 1
fromMatrix : FieldCmp a => Matrix' n a -> Maybe (Rotation n a)
fromMatrix mat = if isRotation' mat then Just (unsafeMkTrans $ matrixToH mat)
else Nothing
export
rotate2D : Num a => Double -> Rotation 2 Double
rotate2D = unsafeMkTrans . rotate2DH
export
rotate3DX : Num a => Double -> Rotation 3 Double
rotate3DX = unsafeMkTrans . rotate3DXH
export
rotate3DY : Num a => Double -> Rotation 3 Double
rotate3DY = unsafeMkTrans . rotate3DYH
export
rotate3DZ : Num a => Double -> Rotation 3 Double
rotate3DZ = unsafeMkTrans . rotate3DZH

View file

@ -16,46 +16,22 @@ import Data.NumIdr.Transform.Point
-------------------------------------------------------------------------------- --------------------------------------------------------------------------------
public export export
data TransType = TAffine | TIsometry | TRigid | TTranslation TransType : Type
| TLinear | TOrthonormal | TRotation | TTrivial TransType = (Fin 4, Bool)
%name TransType ty namespace TransType
export
TAffine, TIsometry, TRigid, TTranslation,
public export TLinear, TOrthonormal, TRotation, TTrivial : TransType
Show TransType where TAffine = (3, True)
show TAffine = "TAffine" TIsometry = (2, True)
show TIsometry = "TIsometry" TRigid = (1, True)
show TRigid = "TRigid" TTranslation = (0, True)
show TTranslation = "TTranslation" TLinear = (3, False)
show TLinear = "TLinear" TOrthonormal = (2, False)
show TOrthonormal = "TOrthonormal" TRotation = (1, False)
show TRotation = "TRotation" TTrivial = (0, False)
show TTrivial = "TTrivial"
-- Lower numbers can be coerced to higher numbers
toSignature : TransType -> (Fin 4, Bool)
toSignature TAffine = (3, True)
toSignature TIsometry = (2, True)
toSignature TRigid = (1, True)
toSignature TTranslation = (0, True)
toSignature TLinear = (3, False)
toSignature TOrthonormal = (2, False)
toSignature TRotation = (1, False)
toSignature TTrivial = (0, False)
public export
fromSignature : (Fin 4, Bool) -> TransType
fromSignature (3, True) = TAffine
fromSignature (2, True) = TIsometry
fromSignature (1, True) = TRigid
fromSignature (0, True) = TTranslation
fromSignature (3, False) = TLinear
fromSignature (2, False) = TOrthonormal
fromSignature (1, False) = TRotation
fromSignature (0, False) = TTrivial
-------------------------------------------------------------------------------- --------------------------------------------------------------------------------
@ -65,17 +41,19 @@ fromSignature (0, False) = TTrivial
public export public export
(:<) : TransType -> TransType -> Bool (:<) : TransType -> TransType -> Bool
x :< y with (toSignature x, toSignature y) (xn, xb) :< (yn, yb) = (xn <= yn) && (xb >= yb)
_ | ((xn, xb), (yn, yb)) = (xn <= yn) && (xb >= yb)
public export public export
transMult : TransType -> TransType -> TransType transMult : TransType -> TransType -> TransType
transMult x y with (toSignature x, toSignature y) transMult (xn, xb) (yn, yb) = (max xn yn, xb && yb)
_ | ((xn, xb), (yn, yb)) = fromSignature (max xn yn, xb && yb)
public export public export
linearizeType : TransType -> TransType linearizeType : TransType -> TransType
linearizeType = fromSignature . mapSnd (const False) . toSignature linearizeType = mapSnd (const False)
public export
delinearizeType : TransType -> TransType
delinearizeType = mapSnd (const True)
export export
@ -98,6 +76,11 @@ linearize : Num a => Transform ty n a -> Transform (linearizeType ty) n a
linearize {n} (MkTrans _ mat) with (viewShape mat) linearize {n} (MkTrans _ mat) with (viewShape mat)
_ | Shape [S n,S n] = MkTrans _ (hmatrix (getMatrix mat) (zeros _)) _ | Shape [S n,S n] = MkTrans _ (hmatrix (getMatrix mat) (zeros _))
export
setTranslation : Num a => Vector n a -> Transform ty n a
-> Transform (delinearizeType ty) n a
setTranslation v (MkTrans _ mat) = MkTrans _ (hmatrix (getMatrix mat) v)
-------------------------------------------------------------------------------- --------------------------------------------------------------------------------
-- Interface implementations -- Interface implementations

View file

@ -18,13 +18,13 @@ Translation = Transform TTranslation
export export
isTranslation : (Eq a, Num a) => HMatrix' n a -> Bool isTranslation : Eq a => Num a => HMatrix' n a -> Bool
isTranslation {n} mat with (viewShape mat) isTranslation {n} mat with (viewShape mat)
_ | Shape [S n,S n] = isHMatrix mat && getMatrix mat == identity _ | Shape [S n,S n] = isHMatrix mat && getMatrix mat == identity
export export
fromVector : Num a => Vector n a -> Translation n a translate : Num a => Vector n a -> Translation n a
fromVector v = unsafeMkTrans (translationH v) translate v = unsafeMkTrans (translationH v)
export export
fromHMatrix : (Eq a, Num a) => HMatrix' n a -> Maybe (Translation n a) fromHMatrix : (Eq a, Num a) => HMatrix' n a -> Maybe (Translation n a)

View file

@ -18,6 +18,6 @@ Trivial = Transform TTrivial
export export
isTrivial : (Eq a, Num a) => HMatrix' n a -> Bool isTrivial : Eq a => Num a => HMatrix' n a -> Bool
isTrivial {n} mat with (viewShape mat) isTrivial {n} mat with (viewShape mat)
_ | Shape [S n,S n] = mat == identity _ | Shape [S n,S n] = mat == identity

View file

@ -43,7 +43,7 @@ swapElems i j v =
export export
permuteVect : Permutation n -> Vect n a -> Vect n a permuteVect : Permutation n -> Vect n a -> Vect n a
permuteVect p = foldMap @{%search} @{mon} (\(i,j) => swapElems i j) p.swaps permuteVect p = foldMap @{%search} @{mon} (uncurry swapElems) p.swaps
export export
@ -54,7 +54,7 @@ swapValues i j x = if x == cast i then cast j
export export
permuteValues : Permutation n -> Nat -> Nat permuteValues : Permutation n -> Nat -> Nat
permuteValues p = foldMap @{%search} @{mon} (\(i,j) => swapValues i j) p.swaps permuteValues p = foldMap @{%search} @{mon} (uncurry swapValues) p.swaps