File: bug_14014.v

package info (click to toggle)
coq-doc 8.16.1-1
  • links: PTS, VCS
  • area: non-free
  • in suites: bookworm
  • size: 42,788 kB
  • sloc: ml: 219,673; sh: 4,035; python: 3,372; ansic: 2,529; makefile: 728; lisp: 279; javascript: 87; xml: 24; sed: 2
file content (12 lines) | stat: -rw-r--r-- 663 bytes parent folder | download | duplicates (2)
1
2
3
4
5
6
7
8
9
10
11
12
Inductive SFalse : SProp := .
Definition adjust_of_sprop {A} {x y : A} (pf : x <> y) : x <> y
  := fun e : x = y => SFalse_ind (fun _ => False) (match pf e return SFalse with end).

Definition adjust_of_sprop_idempotent {A x y pf} : @adjust_of_sprop A x y (@adjust_of_sprop A x y pf) = @adjust_of_sprop A x y pf.
Proof. cbv [adjust_of_sprop]; reflexivity. Defined.
Definition adjust_of_sprop_idempotent' {A x y pf} : @adjust_of_sprop A x y (@adjust_of_sprop A x y pf) = @adjust_of_sprop A x y pf
  := Eval cbv in @adjust_of_sprop_idempotent A x y pf.

Inductive sF : SProp := .

Definition ff (x y:sF) : match x return nat with end = match y with end := eq_refl.