Idris2Doc : Data.Linear.Inverse

Data.Linear.Inverse(source)

This module is based on the content of the functional pearl
How to Take the Inverse of a Type
by Daniel Marshall and Dominic Orchard

Definitions

Inverse : Type->Type
  Inverse is a shorthand for 'multiplicative skew inverse'

Totality: total
Visibility: export
mkInverse : (a-@ ()) -@Inversea
Totality: total
Visibility: export
divide : a-@ (Inversea-@ ())
  And this is the morphism witnessing the multiplicative skew inverse

Totality: total
Visibility: export
doubleNegation : a-@Inverse (Inversea)
Totality: total
Visibility: export
maybeNeg : LMaybea-@ (Inversea-@LMaybe (Inversea))
  We only use the inverse of a if the LMaybe is actually Just.m

Totality: total
Visibility: export
lid : a-@a
Totality: total
Visibility: public export
lcompose : (b-@c) -@ ((a-@b) -@ (a-@c))
Totality: total
Visibility: public export
comap : (b-@a) -@ (Inversea-@Inverseb)
Totality: total
Visibility: public export
comapId : (u : Inversea) ->comaplidu=lidu
Totality: total
Visibility: export
comapComp : (u : Inversea) ->comap (lcomposegf) u=lcompose (comapf) (comapg) u
Totality: total
Visibility: export
munit : () -@Inverse ()
Totality: total
Visibility: export
mmult : Inversea-@ (Inverseb-@Inverse (LPairab))
Totality: total
Visibility: export
Pow : Type->INTEGER->Type
Totality: total
Visibility: public export
zip : Powan-@ (Powbn-@Pow (LPairab) n)
Totality: total
Visibility: public export
Fixity Declaration: infixr operator, level 6
powPositiveL : (m : Nat) -> (n : Nat) ->Powa (cast (m+n)) -@LPair (Powa (castm)) (Powa (castn))
Totality: total
Visibility: export
powPositiveR : (m : Nat) -> (n : Nat) ->Powa (castm) -@ (Powa (castn) -@Powa (cast (m+n)))
Totality: total
Visibility: export
powNegateL : (m : Nat) ->Powa (negate (castm)) ->Pow (Inversea) (castm)
Totality: total
Visibility: export
powNegateR : (m : Nat) ->Pow (Inversea) (castm) ->Powa (negate (castm))
Totality: total
Visibility: export
powNegativeL : (m : Nat) -> (n : Nat) ->Powa (negate (cast (m+n))) -@LPair (Powa (negate (castm))) (Powa (negate (castn)))
Totality: total
Visibility: export
powNegativeR : (m : Nat) -> (n : Nat) ->Powa (negate (castm)) -@ (Powa (negate (castn)) -@Powa (negate (cast (m+n))))
Totality: total
Visibility: export
powAnnihilate : (m : Nat) ->Powa (castm) -@ (Powa (negate (castm)) -@ ())
Totality: total
Visibility: export
powMinus : (m : Nat) -> (n : Nat) ->Powa (castm) -@ (Powa (negate (castn)) -@Powa (castm-castn))
Totality: total
Visibility: export