mirror of https://github.com/nmvdw/HITs-Examples
Switch to the DecidablePaths typeclass
This commit is contained in:
parent
ebdaab4de0
commit
5abf0acf06
18
FinSets.v
18
FinSets.v
|
@ -331,9 +331,7 @@ End FinSet.
|
||||||
|
|
||||||
Section functions.
|
Section functions.
|
||||||
Parameter A : Type.
|
Parameter A : Type.
|
||||||
Context (A_eqdec: forall (x y : A), Decidable (x = y)).
|
Context {A_eqdec: DecidablePaths A}.
|
||||||
Definition deceq (x y : A) :=
|
|
||||||
if dec (x = y) then true else false.
|
|
||||||
|
|
||||||
Notation "{| x |}" := (L x).
|
Notation "{| x |}" := (L x).
|
||||||
Infix "∪" := U (at level 8, right associativity).
|
Infix "∪" := U (at level 8, right associativity).
|
||||||
|
@ -358,7 +356,7 @@ Proof.
|
||||||
intros a.
|
intros a.
|
||||||
hrecursion.
|
hrecursion.
|
||||||
- exact false.
|
- exact false.
|
||||||
- intro a'. apply (deceq a a').
|
- intro a'. destruct (dec (a = a')); [exact true | exact false].
|
||||||
- apply orb.
|
- apply orb.
|
||||||
- intros x y z. compute. destruct x; reflexivity.
|
- intros x y z. compute. destruct x; reflexivity.
|
||||||
- intros x y. compute. destruct x, y; reflexivity.
|
- intros x y. compute. destruct x, y; reflexivity.
|
||||||
|
@ -457,7 +455,6 @@ hinduction.
|
||||||
all: try (intros; apply path_forall; intro; apply set_path2).
|
all: try (intros; apply path_forall; intro; apply set_path2).
|
||||||
- intro Y. cbv. reflexivity.
|
- intro Y. cbv. reflexivity.
|
||||||
- intros a Y. cbn.
|
- intros a Y. cbn.
|
||||||
unfold deceq;
|
|
||||||
destruct (dec (a = a)); simpl.
|
destruct (dec (a = a)); simpl.
|
||||||
+ reflexivity.
|
+ reflexivity.
|
||||||
+ contradiction n. reflexivity.
|
+ contradiction n. reflexivity.
|
||||||
|
@ -494,7 +491,7 @@ hrecursion X; try (intros; apply set_path2).
|
||||||
- cbn. unfold intersection. intros a.
|
- cbn. unfold intersection. intros a.
|
||||||
hrecursion Y; try (intros; apply set_path2).
|
hrecursion Y; try (intros; apply set_path2).
|
||||||
+ cbn. reflexivity.
|
+ cbn. reflexivity.
|
||||||
+ cbn. intros. unfold deceq.
|
+ cbn. intros.
|
||||||
destruct (dec (a0 = a)).
|
destruct (dec (a0 = a)).
|
||||||
rewrite p. destruct (dec (a=a)).
|
rewrite p. destruct (dec (a=a)).
|
||||||
reflexivity.
|
reflexivity.
|
||||||
|
@ -505,16 +502,13 @@ hrecursion X; try (intros; apply set_path2).
|
||||||
+ cbn -[isIn]. intros Y1 Y2 IH1 IH2.
|
+ cbn -[isIn]. intros Y1 Y2 IH1 IH2.
|
||||||
rewrite IH1.
|
rewrite IH1.
|
||||||
rewrite IH2.
|
rewrite IH2.
|
||||||
symmetry.
|
apply (comprehension_union (L a)).
|
||||||
rewrite (comprehension_union (L a)).
|
|
||||||
reflexivity.
|
|
||||||
- intros X1 X2 IH1 IH2.
|
- intros X1 X2 IH1 IH2.
|
||||||
cbn.
|
cbn.
|
||||||
unfold intersection in *.
|
unfold intersection in *.
|
||||||
rewrite <- IH1.
|
rewrite <- IH1.
|
||||||
rewrite <- IH2.
|
rewrite <- IH2. symmetry.
|
||||||
rewrite comprehension_union.
|
apply comprehension_union.
|
||||||
reflexivity.
|
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
|
|
||||||
|
|
Loading…
Reference in New Issue