mirror of
https://github.com/nmvdw/HITs-Examples
synced 2025-11-03 15:13:51 +01:00
Some cleaning in list representation
This commit is contained in:
@@ -11,15 +11,17 @@ Section length.
|
||||
apply (if a ∈_d X then n else (S n)).
|
||||
- intros X a n.
|
||||
simpl.
|
||||
simplify_isIn_d.
|
||||
destruct (dec (a ∈ X)) ; reflexivity.
|
||||
rewrite ?union_isIn_d, singleton_isIn_d_aa.
|
||||
reflexivity.
|
||||
- intros X a b n.
|
||||
simpl.
|
||||
simplify_isIn_d.
|
||||
rewrite ?union_isIn_d.
|
||||
destruct (m_dec_path a b) as [Hab | Hab].
|
||||
+ strip_truncations.
|
||||
rewrite Hab. simplify_isIn_d. reflexivity.
|
||||
+ rewrite ?singleton_isIn_d_false; auto.
|
||||
rewrite Hab.
|
||||
rewrite ?singleton_isIn_d_aa.
|
||||
reflexivity.
|
||||
+ rewrite ?singleton_isIn_d_false.
|
||||
++ simpl.
|
||||
destruct (a ∈_d X), (b ∈_d X) ; reflexivity.
|
||||
++ intro p. contradiction (Hab (tr p^)).
|
||||
|
||||
@@ -219,18 +219,12 @@ Section FSet_cons_rec.
|
||||
(Pcomm : forall X a b p, Pcons a ({|b|} ∪ X) (Pcons b X p)
|
||||
= Pcons b ({|a|} ∪ X) (Pcons a X p)).
|
||||
|
||||
Definition FSet_cons_rec (X : FSet A) : P.
|
||||
Proof.
|
||||
simple refine (FSetC_ind A (fun _ => P) _ Pe _ _ _ (FSet_to_FSetC X)) ; simpl.
|
||||
- intros a Y p.
|
||||
apply (Pcons a (FSetC_to_FSet Y) p).
|
||||
- intros.
|
||||
refine (transport_const _ _ @ _).
|
||||
apply Pdupl.
|
||||
- intros.
|
||||
refine (transport_const _ _ @ _).
|
||||
apply Pcomm.
|
||||
Defined.
|
||||
Definition FSet_cons_rec (X : FSet A) : P :=
|
||||
FSetC_prim_rec A P Pset Pe
|
||||
(fun a Y p => Pcons a (FSetC_to_FSet Y) p)
|
||||
(fun _ _ => Pdupl _ _)
|
||||
(fun _ _ _ => Pcomm _ _ _)
|
||||
(FSet_to_FSetC X).
|
||||
|
||||
Definition FSet_cons_beta_empty : FSet_cons_rec ∅ = Pe := idpath.
|
||||
|
||||
|
||||
Reference in New Issue
Block a user