File: setoid_unif.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 (28 lines) | stat: -rw-r--r-- 918 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
26
27
28
(* An example of unification in rewrite which uses eager substitution
   of metas (provided by Pierre-Marie).

   Put in the test suite as an indication of what the use metas
   eagerly flag provides, even though the concrete cases that use it
   are seldom. Today supported thanks to a new flag for using evars
   eagerly, after this variant of setoid rewrite started to use clause
   environments based on evars (fbbe491cfa157da627) *)

Require Import Setoid.

Parameter elt : Type.
Parameter T : Type -> Type.
Parameter empty : forall A, T A.
Parameter MapsTo : forall A : Type, elt -> A -> T A -> Prop.

(* Definition In A x t := exists e, MapsTo A x e t. *)
Axiom In : forall A, A -> T A -> Prop.
Axiom foo : forall A x, In A x (empty A) <-> False.

Record R := { t : T unit; s : unit }.
Definition Empty := {| t := empty unit; s := tt |}.

Goal forall x, ~ In _ x (t Empty).
Proof.
intros x.
rewrite foo.
Abort.