mirror of https://github.com/nmvdw/HITs-Examples
Some cleaning in list representation
This commit is contained in:
parent
0e9fcbc588
commit
d7a95697fb
|
@ -11,15 +11,17 @@ Section length.
|
||||||
apply (if a ∈_d X then n else (S n)).
|
apply (if a ∈_d X then n else (S n)).
|
||||||
- intros X a n.
|
- intros X a n.
|
||||||
simpl.
|
simpl.
|
||||||
simplify_isIn_d.
|
rewrite ?union_isIn_d, singleton_isIn_d_aa.
|
||||||
destruct (dec (a ∈ X)) ; reflexivity.
|
reflexivity.
|
||||||
- intros X a b n.
|
- intros X a b n.
|
||||||
simpl.
|
simpl.
|
||||||
simplify_isIn_d.
|
rewrite ?union_isIn_d.
|
||||||
destruct (m_dec_path a b) as [Hab | Hab].
|
destruct (m_dec_path a b) as [Hab | Hab].
|
||||||
+ strip_truncations.
|
+ strip_truncations.
|
||||||
rewrite Hab. simplify_isIn_d. reflexivity.
|
rewrite Hab.
|
||||||
+ rewrite ?singleton_isIn_d_false; auto.
|
rewrite ?singleton_isIn_d_aa.
|
||||||
|
reflexivity.
|
||||||
|
+ rewrite ?singleton_isIn_d_false.
|
||||||
++ simpl.
|
++ simpl.
|
||||||
destruct (a ∈_d X), (b ∈_d X) ; reflexivity.
|
destruct (a ∈_d X), (b ∈_d X) ; reflexivity.
|
||||||
++ intro p. contradiction (Hab (tr p^)).
|
++ 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)
|
(Pcomm : forall X a b p, Pcons a ({|b|} ∪ X) (Pcons b X p)
|
||||||
= Pcons b ({|a|} ∪ X) (Pcons a X p)).
|
= Pcons b ({|a|} ∪ X) (Pcons a X p)).
|
||||||
|
|
||||||
Definition FSet_cons_rec (X : FSet A) : P.
|
Definition FSet_cons_rec (X : FSet A) : P :=
|
||||||
Proof.
|
FSetC_prim_rec A P Pset Pe
|
||||||
simple refine (FSetC_ind A (fun _ => P) _ Pe _ _ _ (FSet_to_FSetC X)) ; simpl.
|
(fun a Y p => Pcons a (FSetC_to_FSet Y) p)
|
||||||
- intros a Y p.
|
(fun _ _ => Pdupl _ _)
|
||||||
apply (Pcons a (FSetC_to_FSet Y) p).
|
(fun _ _ _ => Pcomm _ _ _)
|
||||||
- intros.
|
(FSet_to_FSetC X).
|
||||||
refine (transport_const _ _ @ _).
|
|
||||||
apply Pdupl.
|
|
||||||
- intros.
|
|
||||||
refine (transport_const _ _ @ _).
|
|
||||||
apply Pcomm.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Definition FSet_cons_beta_empty : FSet_cons_rec ∅ = Pe := idpath.
|
Definition FSet_cons_beta_empty : FSet_cons_rec ∅ = Pe := idpath.
|
||||||
|
|
||||||
|
|
|
@ -79,6 +79,27 @@ Module Export FSetC.
|
||||||
indTy := _; recTy := _;
|
indTy := _; recTy := _;
|
||||||
H_inductor := FSetC_ind A; H_recursor := FSetC_rec A
|
H_inductor := FSetC_ind A; H_recursor := FSetC_rec A
|
||||||
}.
|
}.
|
||||||
|
|
||||||
|
Section FSetC_prim_recursion.
|
||||||
|
Variable (A : Type)
|
||||||
|
(P : Type)
|
||||||
|
(H : IsHSet P)
|
||||||
|
(nil : P)
|
||||||
|
(cns : A -> FSetC A -> P -> P)
|
||||||
|
(duplP : forall (a : A) (X : FSetC A) (x : P),
|
||||||
|
cns a (a ;; X) (cns a X x) = (cns a X x))
|
||||||
|
(commP : forall (a b: A) (X : FSetC A) (x: P),
|
||||||
|
cns a (b ;; X) (cns b X x) = cns b (a ;; X) (cns a X x)).
|
||||||
|
|
||||||
|
(* Recursion principle *)
|
||||||
|
Definition FSetC_prim_rec : FSetC A -> P.
|
||||||
|
Proof.
|
||||||
|
simple refine (FSetC_ind A (fun _ => P) (fun _ => H) nil cns _ _ );
|
||||||
|
try (intros; simple refine ((transport_const _ _) @ _ )); cbn.
|
||||||
|
- apply duplP.
|
||||||
|
- apply commP.
|
||||||
|
Defined.
|
||||||
|
End FSetC_prim_recursion.
|
||||||
End FSetC.
|
End FSetC.
|
||||||
|
|
||||||
Infix ";;" := Cns (at level 8, right associativity).
|
Infix ";;" := Cns (at level 8, right associativity).
|
||||||
|
|
Loading…
Reference in New Issue