diff --git a/tests/telescopes.v b/tests/telescopes.v index e282b131e68b71ae081e9eb55942b416597a45f0..d231e29188b5ea3d5da6f0ca130643a23559640e 100644 --- a/tests/telescopes.v +++ b/tests/telescopes.v @@ -52,3 +52,30 @@ Definition no_bump@{u} (t : tele@{u}) : Type@{u} := tele_arg@{u} t. Lemma texist_exist_universes (X : Type) (P : TeleS (λ _ : X, TeleO) → Prop) : texist P ↔ ex P. Proof. by rewrite texist_exist. Qed. + +(** [tele_arg ..] notation tests. + These tests mainly test type annotations and casts in the [tele_arg] + notations. + We test that Coq can typecheck literal telescope arguments in two ways: + - tactic unification/old unification using [exact] + - evarconv/new unification using [refine] + *) +Example tele_arg_notation_0 : [tele]. +assert_succeeds exact [tele_arg]. +assert_succeeds refine [tele_arg]. +Abort. + +Example tele_arg_notation_1 : [tele (_:nat)]. +assert_succeeds exact [tele_arg 0]. +assert_succeeds refine [tele_arg 0]. +Abort. + +Example tele_arg_notation_2 : [tele (_ : bool) (_ : nat)]. +assert_succeeds exact [tele_arg true; 0]. +assert_succeeds refine [tele_arg true; 0]. +Abort. + +Example tele_arg_notation_2_dep : [tele (b : bool) (_ : if b then nat else False)]. +assert_succeeds exact [tele_arg true; 0]. +assert_succeeds refine [tele_arg true; 0]. +Abort. diff --git a/theories/telescopes.v b/theories/telescopes.v index 3dad07edcfa4b22ad8e041d9bb0cadeefa653e11..2378216ea18776b08f38b416d35a16500db3c278 100644 --- a/theories/telescopes.v +++ b/theories/telescopes.v @@ -52,8 +52,20 @@ Fixpoint tele_arg@{u} (t : tele@{u}) : Type@{u} := | TeleS f => tele_arg_cons (λ x, tele_arg (f x)) end. Global Arguments tele_arg _ : simpl never. -Notation TargO := tt (only parsing). -Notation TargS a b := (@TeleArgCons _ (λ x, tele_arg _) a b) (only parsing). + +(* Coq has no idea that [unit] and [tele_arg_cons] have anything to do with + telescopes. This only becomes a problem when concrete telescope arguments + (of concrete telescopes) need to be typechecked. To work around this, we + annotate the notations below with extra information to guide unification. + *) + +(* The cast in the notation below is necessary to make Coq understand that + [TargO] can be unified with [tele_arg TeleO]. *) +Notation TargO := (tt : tele_arg TeleO) (only parsing). +(* The casts and annotations are necessary for Coq to typecheck nested [TargS] + as well as the final [TargO] in a chain of [TargS]. *) +Notation TargS a b := + ((@TeleArgCons _ (λ x, tele_arg (_ x)) a b) : (tele_arg (TeleS _))) (only parsing). Coercion tele_arg : tele >-> Sortclass. Lemma tele_arg_ind (P : ∀ TT, tele_arg TT → Prop) :