HITs-Examples/FiniteSets/kuratowski/properties_decidable.v

146 lines
3.7 KiB
Coq
Raw Normal View History

2017-08-01 15:12:59 +02:00
(* Properties of [FSet A] where [A] has decidable equality *)
Require Import HoTT HitTactics.
From fsets Require Export properties extensionality operations_decidable.
Require Export lattice.
2017-08-01 15:12:59 +02:00
(* Lemmas relating operations to the membership predicate *)
Section operations_isIn.
2017-08-03 12:21:34 +02:00
Context {A : Type} `{DecidablePaths A} `{Univalence}.
2017-08-01 15:12:59 +02:00
2017-08-08 17:00:30 +02:00
Lemma ext : forall (S T : FSet A), (forall a, a _d S = a _d T) -> S = T.
2017-08-03 12:21:34 +02:00
Proof.
intros X Y H2.
apply fset_ext.
intro a.
specialize (H2 a).
2017-08-08 17:00:30 +02:00
unfold member_dec, fset_member_bool, dec in H2.
2017-08-03 12:21:34 +02:00
destruct (isIn_decidable a X) ; destruct (isIn_decidable a Y).
- apply path_iff_hprop ; intro ; assumption.
- contradiction (true_ne_false).
- contradiction (true_ne_false) ; apply H2^.
- apply path_iff_hprop ; intro ; contradiction.
Defined.
2017-08-01 15:12:59 +02:00
2017-08-03 12:21:34 +02:00
Lemma empty_isIn (a : A) :
2017-08-08 17:00:30 +02:00
a _d = false.
2017-08-03 12:21:34 +02:00
Proof.
reflexivity.
Defined.
2017-08-01 15:12:59 +02:00
2017-08-03 12:21:34 +02:00
Lemma L_isIn (a b : A) :
2017-08-07 16:49:46 +02:00
a {|b|} -> a = b.
2017-08-03 12:21:34 +02:00
Proof.
intros. strip_truncations. assumption.
Defined.
2017-08-01 15:12:59 +02:00
2017-08-03 12:21:34 +02:00
Lemma L_isIn_b_true (a b : A) (p : a = b) :
2017-08-08 17:00:30 +02:00
a _d {|b|} = true.
2017-08-03 12:21:34 +02:00
Proof.
2017-08-08 17:00:30 +02:00
unfold member_dec, fset_member_bool, dec.
2017-08-03 12:21:34 +02:00
destruct (isIn_decidable a {|b|}) as [n | n] .
- reflexivity.
- simpl in n.
contradiction (n (tr p)).
Defined.
2017-08-01 15:12:59 +02:00
2017-08-03 12:21:34 +02:00
Lemma L_isIn_b_aa (a : A) :
2017-08-08 17:00:30 +02:00
a _d {|a|} = true.
2017-08-03 12:21:34 +02:00
Proof.
apply L_isIn_b_true ; reflexivity.
Defined.
2017-08-01 15:12:59 +02:00
2017-08-03 12:21:34 +02:00
Lemma L_isIn_b_false (a b : A) (p : a <> b) :
2017-08-08 17:00:30 +02:00
a _d {|b|} = false.
2017-08-03 12:21:34 +02:00
Proof.
2017-08-08 17:00:30 +02:00
unfold member_dec, fset_member_bool, dec in *.
2017-08-03 12:21:34 +02:00
destruct (isIn_decidable a {|b|}).
- simpl in t.
strip_truncations.
contradiction.
- reflexivity.
Defined.
2017-08-03 12:21:34 +02:00
(* Union and membership *)
Lemma union_isIn_b (X Y : FSet A) (a : A) :
2017-08-08 17:00:30 +02:00
a _d (X Y) = orb (a _d X) (a _d Y).
2017-08-03 12:21:34 +02:00
Proof.
2017-08-08 17:00:30 +02:00
unfold member_dec, fset_member_bool, dec.
2017-08-03 12:21:34 +02:00
simpl.
destruct (isIn_decidable a X) ; destruct (isIn_decidable a Y) ; reflexivity.
Defined.
2017-08-01 15:12:59 +02:00
Lemma comprehension_isIn_b (Y : FSet A) (ϕ : A -> Bool) (a : A) :
2017-08-08 17:00:30 +02:00
a _d {|Y & ϕ|} = andb (a _d Y) (ϕ a).
2017-08-03 12:21:34 +02:00
Proof.
2017-08-08 17:00:30 +02:00
unfold member_dec, fset_member_bool, dec ; simpl.
2017-08-08 15:29:50 +02:00
destruct (isIn_decidable a {|Y & ϕ|}) as [t | t]
; destruct (isIn_decidable a Y) as [n | n] ; rewrite comprehension_isIn in t
2017-08-08 17:00:30 +02:00
; destruct (ϕ a) ; try reflexivity ; try contradiction
; try (contradiction (n t)) ; contradiction (t n).
2017-08-03 12:21:34 +02:00
Defined.
Lemma intersection_isIn_b (X Y: FSet A) (a : A) :
2017-08-08 17:00:30 +02:00
a _d (intersection X Y) = andb (a _d X) (a _d Y).
2017-08-03 12:21:34 +02:00
Proof.
apply comprehension_isIn_b.
2017-08-03 12:21:34 +02:00
Defined.
2017-08-01 15:12:59 +02:00
End operations_isIn.
(* Some suporting tactics *)
Ltac simplify_isIn_b :=
repeat (rewrite union_isIn_b
|| rewrite L_isIn_b_aa
|| rewrite intersection_isIn_b
|| rewrite comprehension_isIn_b).
2017-08-03 12:21:34 +02:00
Ltac toBool :=
repeat intro;
apply ext ; intros ;
simplify_isIn_b ; eauto with bool_lattice_hints typeclass_instances.
2017-08-01 15:12:59 +02:00
Section SetLattice.
Context {A : Type}.
Context {A_deceq : DecidablePaths A}.
Context `{Univalence}.
2017-08-08 15:29:50 +02:00
Instance fset_max : maximum (FSet A).
2017-08-01 15:12:59 +02:00
Proof.
2017-08-08 15:29:50 +02:00
intros x y.
apply (x y).
2017-08-01 15:12:59 +02:00
Defined.
2017-08-08 15:29:50 +02:00
Instance fset_min : minimum (FSet A).
2017-08-01 15:12:59 +02:00
Proof.
2017-08-08 15:29:50 +02:00
intros x y.
2017-08-08 17:00:30 +02:00
apply (x y).
2017-08-01 15:12:59 +02:00
Defined.
2017-08-08 15:29:50 +02:00
Instance fset_bot : bottom (FSet A).
2017-08-01 15:12:59 +02:00
Proof.
2017-08-08 15:29:50 +02:00
unfold bottom.
apply .
Defined.
2017-08-08 15:29:50 +02:00
Instance lattice_fset : Lattice (FSet A).
2017-08-01 15:12:59 +02:00
Proof.
2017-08-08 15:29:50 +02:00
split; toBool.
2017-08-01 15:12:59 +02:00
Defined.
2017-08-08 15:29:50 +02:00
End SetLattice.
2017-08-01 15:12:59 +02:00
(* With extensionality we can prove decidable equality *)
Section dec_eq.
Context (A : Type) `{DecidablePaths A} `{Univalence}.
Instance fsets_dec_eq : DecidablePaths (FSet A).
Proof.
intros X Y.
2017-08-07 16:49:46 +02:00
apply (decidable_equiv' ((Y X) * (X Y)) (eq_subset X Y)^-1).
apply decidable_prod.
2017-08-01 15:12:59 +02:00
Defined.
End dec_eq.