Skip to content

Commit

Permalink
Connectedness.v: remove a universe variable from conn_point_incl
Browse files Browse the repository at this point in the history
  • Loading branch information
jdchristensen committed Feb 27, 2023
1 parent 99ebf02 commit 68587ee
Showing 1 changed file with 2 additions and 1 deletion.
3 changes: 2 additions & 1 deletion theories/Truncations/Connectedness.v
Original file line number Diff line number Diff line change
Expand Up @@ -77,8 +77,9 @@ Definition conn_point_incl `{Univalence} {n : trunc_index} {A : Type@{u}} (a0:A)
`{IsConnected n.+1 A}
: IsConnMap n (unit_name a0).
Proof.
(* Coq can find [conn_map_to_unit_isconnected] via typeclass search, but we manually pose it to get rid of an unneeded universe variable. *)
(* Coq can find [conn_map_to_unit_isconnected] and [isconnected_contr] via typeclass search, but we manually pose them to get rid of an unneeded universe variable. *)
pose conn_map_to_unit_isconnected@{u u}.
pose isconnected_contr@{u u}.
rapply (OO_cancelL_conn_map@{u u u u} (Tr n.+1) (Tr n) (unit_name a0) (const_tt A)).
apply O_lex_leq_Tr.
Defined.
Expand Down

0 comments on commit 68587ee

Please sign in to comment.