[Agda] Higher inductive-recursive definitions?

Altenkirch Thorsten psztxa at exmail.nottingham.ac.uk
Fri May 23 15:21:45 CEST 2014


I would like to combine higher inductive definitions (I.e. have path constructors) with induction-recursion. One application would be to define a closed universe which is univalent. However, I cannot see any reasonable way to define an eliminator.

Ok, I start with a simple universe(using Agda) - an inductive recursive definition

data U : Set
El : U -> Set

data U where
  nat : U
  pi : (a : U)(b : El a -> U) -> U

El nat = Nat
El (pi a b) = (x : El a) -> El (b x)

Now I can define an eliminator for the universe which allows me to define dependent functions by recursion over type codes:

ElimU : (X : U -> Set)
       -> (X nat)
       -> ((a : U) -> X a -> (b : El a -> U) -> ((x : El a) -> X (b x)) -> X (pi a b))
       -> (a : U) -> X a
ElimU X n p nat = n
ElimU X n p (pi a b) = p a (ElimU X n p a) b (£f x -> ElimU X n p (b x))

However, I also would like to add a path constructor, which identifies codes if the have the same semantics:

postulate
  eqU : forall {a b} -> El a == El b -> a == b

But I don't see a good way to modify the Eliminator. It seems that this corresponds to a condition on the eliminator as a whole.

An alternative is to quotient the universe afterwards. However, the problem is that in this case I cannot lift the pi constructor to the quotiented universe ¡V the usual problem when quotienting infinitary constructors. This can usually be overcome by defining the path constructors mutually..

Any ideas? Maybe the whole thing doesn't make sense semantically?

Thorsten
-------------- next part --------------
An HTML attachment was scrubbed...
URL: http://lists.chalmers.se/pipermail/agda/attachments/20140523/ddb191a8/attachment.html


More information about the Agda mailing list