Make set_eq and instance direclty
This commit is contained in:
parent
ab51715fd3
commit
73e87b6b1d
91
czf.v
91
czf.v
|
@ -15,21 +15,21 @@ Definition branch (a : V) : ind a -> V :=
|
||||||
end.
|
end.
|
||||||
|
|
||||||
(** * Equality *)
|
(** * Equality *)
|
||||||
Fixpoint set_eq (a : V) (b : V) : Prop :=
|
Instance set_eq : Equiv V :=
|
||||||
match (a, b) with
|
fix go a b : Prop :=
|
||||||
(sup A f, sup B g) =>
|
let _ : Equiv V := @go in
|
||||||
(forall (x:A), exists (y: B), (set_eq (f x) (g y)))
|
match a, b with
|
||||||
/\ (forall (y:B), exists (x: A), (set_eq (f x) (g y)))
|
| sup A f, sup B g =>
|
||||||
|
(∀ x:A, ∃ y: B, f x ≡ g y) ∧ (∀ y:B, ∃ x: A, f x ≡ g y)
|
||||||
end.
|
end.
|
||||||
Instance equiv_V : Equiv V := set_eq.
|
|
||||||
|
|
||||||
Lemma eq1 : forall x, x ≡ x.
|
Lemma eq1 : forall (x : V), x ≡ x.
|
||||||
Proof.
|
Proof.
|
||||||
intros. induction x.
|
intros. induction x.
|
||||||
split; intros; [exists x | exists y]; apply H.
|
split; intros; [exists x | exists y]; apply H.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
Lemma eq2 : forall x y, x ≡ y -> y ≡ x.
|
Lemma eq2 : forall (x y : V), x ≡ y -> y ≡ x.
|
||||||
Proof.
|
Proof.
|
||||||
intro x. induction x as [A f IHx].
|
intro x. induction x as [A f IHx].
|
||||||
intro y. destruct y as [B g].
|
intro y. destruct y as [B g].
|
||||||
|
@ -43,7 +43,7 @@ Proof.
|
||||||
exists y0. apply (IHx y (g y0)). assumption.
|
exists y0. apply (IHx y (g y0)). assumption.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
Lemma eq3 : forall x y z, x≡y -> y≡z -> x≡z.
|
Lemma eq3 : forall (x y z : V), x≡y -> y≡z -> x≡z.
|
||||||
Proof.
|
Proof.
|
||||||
intro. induction x as [A f X].
|
intro. induction x as [A f X].
|
||||||
intro. destruct y as [B g].
|
intro. destruct y as [B g].
|
||||||
|
@ -60,10 +60,10 @@ Proof.
|
||||||
exists z. apply (X z (g y0) (h y)); assumption.
|
exists z. apply (X z (g y0) (h y)); assumption.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
Instance set_eq_refl : Reflexive equiv_V := eq1.
|
Instance set_eq_refl : Reflexive set_eq := eq1.
|
||||||
Instance set_eq_sym : Symmetric equiv_V := eq2.
|
Instance set_eq_sym : Symmetric set_eq := eq2.
|
||||||
Instance set_eq_trans : Transitive equiv_V := eq3.
|
Instance set_eq_trans : Transitive set_eq := eq3.
|
||||||
Instance set_eq_equivalence : Equivalence equiv_V.
|
Instance set_eq_equivalence : Equivalence set_eq.
|
||||||
Proof. split; apply _. Qed.
|
Proof. split; apply _. Qed.
|
||||||
|
|
||||||
Lemma set_eq_unfold a b :
|
Lemma set_eq_unfold a b :
|
||||||
|
@ -75,56 +75,24 @@ Proof.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
(** * Membership *)
|
(** * Membership *)
|
||||||
|
|
||||||
Lemma test1 (a b c : V) :
|
|
||||||
a ≡ b →
|
|
||||||
b ≡ c →
|
|
||||||
a ≡ c.
|
|
||||||
Proof.
|
|
||||||
intros ? H2.
|
|
||||||
rewrite <-H2.
|
|
||||||
apply H.
|
|
||||||
Qed.
|
|
||||||
|
|
||||||
Lemma test2 (a b c : V) :
|
|
||||||
equiv_V a b →
|
|
||||||
equiv_V b c →
|
|
||||||
equiv_V a c.
|
|
||||||
Proof.
|
|
||||||
intros ? H2.
|
|
||||||
Fail rewrite <-H2.
|
|
||||||
Abort.
|
|
||||||
|
|
||||||
Instance elem_of_V : ElemOf V V := λ (a b : V),
|
Instance elem_of_V : ElemOf V V := λ (a b : V),
|
||||||
exists (y : ind b), a ≡ (branch b y).
|
exists (y : ind b), a ≡ (branch b y).
|
||||||
|
|
||||||
Instance elem_of_V_proper : Proper ((≡) ==> (≡) ==> (=)) elem_of_V.
|
Instance elem_of_V_proper : Proper ((≡) ==> (≡) ==> (impl)) elem_of_V.
|
||||||
Proof. Admitted.
|
Proof.
|
||||||
|
intros a a' Ha b b' Hb.
|
||||||
|
intros [y Hy].
|
||||||
|
destruct b as [B g].
|
||||||
|
destruct b' as [B' g'].
|
||||||
|
simpl in *.
|
||||||
|
destruct Hb as [i1 i2].
|
||||||
|
destruct (i1 y) as [y' Hy'].
|
||||||
|
exists y'. simpl. rewrite <- Hy'.
|
||||||
|
rewrite <- Ha. apply Hy.
|
||||||
|
Qed.
|
||||||
|
|
||||||
(* Instance elem_of_V_proper : Proper ((=) ==> (≡) ==> (impl)) elem_of_V. *)
|
(* Here we have to use `elem_of` instead of `elem_of_V` *)
|
||||||
(* Proof. *)
|
Instance Acc_elem_of_proper : Proper ((≡) ==> (impl)) (Acc (elem_of : V -> V -> Prop)).
|
||||||
(* intros a a' Ha b b' Hb. *)
|
|
||||||
(* apply propositional_extensionality. split. *)
|
|
||||||
(* rewrite <-Ha. *)
|
|
||||||
(* { intros [y Hy]. *)
|
|
||||||
(* destruct b as [B g]. *)
|
|
||||||
(* destruct b' as [B' g']. *)
|
|
||||||
(* simpl in *. *)
|
|
||||||
(* destruct Hb as [i1 i2]. *)
|
|
||||||
(* destruct (i1 y) as [y' Hy']. *)
|
|
||||||
(* exists y'. simpl. unfold equiv. admit. } *)
|
|
||||||
(* { intros [y' Hy']. *)
|
|
||||||
(* destruct b as [B g]. *)
|
|
||||||
(* destruct b' as [B' g']. *)
|
|
||||||
(* simpl in *. *)
|
|
||||||
(* destruct Hb as [i1 i2]. *)
|
|
||||||
(* destruct (i2 y') as [y Hy]. *)
|
|
||||||
(* exists y. simpl. symmetry. *)
|
|
||||||
(* unfold equiv. *)
|
|
||||||
(* admit. } *)
|
|
||||||
(* Admitted. *)
|
|
||||||
|
|
||||||
Instance Acc_elem_of_proper : Proper ((≡) ==> (impl)) (Acc elem_of_V).
|
|
||||||
Proof.
|
Proof.
|
||||||
intros a a' Ha.
|
intros a a' Ha.
|
||||||
intros acc.
|
intros acc.
|
||||||
|
@ -135,9 +103,8 @@ Proof.
|
||||||
destruct Ha as [i1 i2].
|
destruct Ha as [i1 i2].
|
||||||
constructor. intros x' [y' Hx'].
|
constructor. intros x' [y' Hx'].
|
||||||
eapply IH. 2:{ reflexivity. }
|
eapply IH. 2:{ reflexivity. }
|
||||||
(* TODO: does not work *)
|
rewrite Hx'.
|
||||||
(* rewrite Hx'. *)
|
(* erewrite (elem_of_V_proper x' (branch a' y')); eauto. *)
|
||||||
erewrite (elem_of_V_proper x' (branch a' y')); eauto.
|
|
||||||
destruct (i2 y') as [y Hy].
|
destruct (i2 y') as [y Hy].
|
||||||
exists y. symmetry. apply Hy.
|
exists y. symmetry. apply Hy.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
Loading…
Reference in New Issue