mirror of https://github.com/nmvdw/HITs-Examples
Length also uses merely decidable equality
This commit is contained in:
parent
39e2ce1c05
commit
7f9b2b7032
|
@ -1,8 +1,8 @@
|
||||||
Require Import HoTT HitTactics.
|
Require Import HoTT HitTactics prelude.
|
||||||
Require Import kuratowski.operations kuratowski.properties kuratowski.kuratowski_sets.
|
Require Import kuratowski.operations kuratowski.properties kuratowski.kuratowski_sets.
|
||||||
|
|
||||||
Section Length.
|
Section Length.
|
||||||
Context {A : Type} `{DecidablePaths A} `{Univalence}.
|
Context {A : Type} `{MerelyDecidablePaths A} `{Univalence}.
|
||||||
|
|
||||||
Definition length : FSet A -> nat.
|
Definition length : FSet A -> nat.
|
||||||
simple refine (FSet_cons_rec _ _ _ _ _ _).
|
simple refine (FSet_cons_rec _ _ _ _ _ _).
|
||||||
|
@ -16,11 +16,15 @@ Section Length.
|
||||||
- intros X a b n.
|
- intros X a b n.
|
||||||
simpl.
|
simpl.
|
||||||
simplify_isIn_d.
|
simplify_isIn_d.
|
||||||
destruct (dec (a = b)) as [Hab | Hab].
|
destruct (m_dec_path a b) as [Hab | Hab].
|
||||||
+ rewrite Hab. simplify_isIn_d. reflexivity.
|
(* destruct (dec (Trunc (-1) (a = b))) as [Hab | Hab]. *)
|
||||||
|
+ strip_truncations.
|
||||||
|
rewrite Hab. simplify_isIn_d. reflexivity.
|
||||||
+ rewrite ?singleton_isIn_d_false; auto.
|
+ rewrite ?singleton_isIn_d_false; auto.
|
||||||
++ simpl.
|
++ simpl.
|
||||||
destruct (a ∈_d X), (b ∈_d X) ; reflexivity.
|
destruct (a ∈_d X), (b ∈_d X) ; reflexivity.
|
||||||
++ intro p. contradiction (Hab p^).
|
++ intro p. contradiction (Hab (tr p^)).
|
||||||
|
++ intros p.
|
||||||
|
apply (Hab (tr p)).
|
||||||
Defined.
|
Defined.
|
||||||
End Length.
|
End Length.
|
Loading…
Reference in New Issue