File: bug_5434.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 (19 lines) | stat: -rw-r--r-- 661 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
(* About binders which remain unnamed after typing *)

Global Set Asymmetric Patterns.

Definition proj2_sig_map {A} {P Q : A -> Prop} (f : forall a, P a -> Q a) (x :
@sig A P) : @sig A Q
  := let 'exist a p := x in exist Q a (f a p).
Axioms (feBW' : Type) (g : Prop -> Prop) (f' : feBW' -> Prop).
Definition foo := @proj2_sig_map feBW' (fun  H  => True = f' _) (fun H =>
 g True = g (f' H))
                                 (fun (a : feBW') (p : (fun H : feBW' => True =
 f' H) a) => @f_equal Prop Prop g True (f' a) p).
Print foo.
Goal True.
  lazymatch type of foo with
  | sig (fun a : ?A => ?P) -> _
    => pose (fun a : A => a = a /\ P = P)
  end.
Abort.