File: bug_3895.v

package info (click to toggle)
coq 8.20.1%2Bdfsg-1
  • links: PTS, VCS
  • area: main
  • in suites: forky, 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 (23 lines) | stat: -rw-r--r-- 949 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
Notation pr1 := (@projT1 _ _).
Notation compose := (fun g' f' x => g' (f' x)).
Notation "g 'o' f" := (compose g f) (at level 40, left associativity) :
function_scope.
Open Scope function_scope.
Definition ap {A B:Type} (f:A -> B) {x y:A} (p:x = y) : f x = f y := match p
with eq_refl => eq_refl end.
Definition pointwise_paths {A} {P:A->Type} (f g:forall x:A, P x) := forall x:A,
f x = g x.
Notation "f == g" := (pointwise_paths f g) (at level 70, no associativity) :
type_scope.
Theorem Univalence_implies_FunextNondep (A B : Type)
: forall f g : A -> B, f == g -> f = g.
Proof.
  intros f g p.
  pose (d := fun x : A => existT (fun xy => fst xy = snd xy) (f x, f x)
(eq_refl (f x))).
  pose (e := fun x : A => existT (fun xy => fst xy = snd xy) (f x, g x) (p x)).
  change f with ((snd o pr1) o d).
  change g with ((snd o pr1) o e).
  apply (ap (fun g => snd o pr1 o g)).
(* Used to raise a not Found due to a "typo" in solve_evar_evar *)
Abort.