File: ConversionOrder.v

package info (click to toggle)
coq 8.20.1%2Bdfsg-1
  • links: PTS, VCS
  • area: main
  • in suites: sid, trixie
  • size: 44,116 kB
  • sloc: ml: 234,160; sh: 4,301; python: 3,270; ansic: 2,644; makefile: 882; lisp: 172; javascript: 63; xml: 24; sed: 2
file content (16 lines) | stat: -rw-r--r-- 513 bytes parent folder | download | duplicates (5)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
(* The kernel may convert application arguments right to left,
   resulting in ill-typed terms, but should be robust to them. *)

Inductive Hide := hide : forall A, A -> Hide.

Lemma foo : (hide Type Type) = (hide (nat -> Type) (fun x : nat => Type)).
Proof.
  Fail reflexivity.
  match goal with |- ?l = _ => exact_no_check (eq_refl l) end.
  Fail Defined.
Abort.

Definition HideMore (_:Hide) := 0.

Definition foo : HideMore (hide Type Type) = HideMore (hide (nat -> Type) (fun x : nat => Type))
  := eq_refl.