File: bug_2139.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 (25 lines) | stat: -rw-r--r-- 736 bytes parent folder | download | duplicates (5)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
(* Call of apply on <-> failed because of evars in elimination predicate *)
Generalizable Variables patch.

Class Patch (patch : Type) := {
    commute : patch -> patch -> Prop
}.

Parameter flip : forall	`{patchInstance	: Patch	patch}
                        {a b : patch},
                 commute a b <-> commute b a.

Lemma Foo : forall `{patchInstance : Patch patch}
                    {a b : patch},
            (commute a b)
         -> True.
Proof.
intros.
apply flip in H.

(* failed in well-formed arity check because elimination predicate of
   iff in (@flip _ _ _ _) had normalized evars while the ones in the
   type of (@flip _ _ _ _) itself had non-normalized evars *)

(* By the way, is the check necessary ? *)
Abort.