diff --git a/CaseStudies/Convergence/Algorithm_noB.v b/CaseStudies/Convergence/Algorithm_noB.v index f347a298a879ab5e551f668edaf4a5fec426665e..cdc2905d87a5f96ee6869ed063c9e8b6b01aa016 100644 --- a/CaseStudies/Convergence/Algorithm_noB.v +++ b/CaseStudies/Convergence/Algorithm_noB.v @@ -31,9 +31,9 @@ Typeclasses eauto := (bfs). Section ConvergenceAlgo. (** There are [ub] good robots and no byzantine one. *) -Context {gt0 : greater_than 0}. +Context {k : nat} {ltc_0_k : 0 <c k}. -Instance MyRobots : Names := Robots ub 0. +Instance MyRobots : Names := Robots k 0. Instance NoByz : NoByzantine. Proof using . now split. Qed. @@ -79,7 +79,7 @@ Implicit Type pt : location. (** As there are robots, the observation can never be empty. *) Lemma obs_non_empty : forall config pt, obs_from_config config pt =/= @empty location _ _ _. -Proof using . +Proof using ltc_0_k. intros config pt. rewrite obs_from_config_ignore_snd. intro Habs. pose (g := fin0 : G). @@ -159,7 +159,7 @@ Theorem round_simplify : forall da config, SSYNC_da da -> == fun id => if da.(activate) id then isobarycenter (@elements location _ _ _ (!! config)) else config id. -Proof using . +Proof using ltc_0_k. intros da config HSSYNC. rewrite SSYNC_round_simplify; trivial; []. intro id. pattern id. apply no_byz. clear id. intro g. unfold round. destruct_match; try reflexivity; []. @@ -196,7 +196,7 @@ Qed. Lemma contained_next : forall da c r config, SSYNC_da da -> contained c r config -> contained c r (round convergeR2 da config). -Proof using . +Proof using ltc_0_k. intros da c r config HSSYNC Hconfig g. rewrite round_simplify; trivial; []. destruct_match. @@ -206,7 +206,7 @@ Qed. Lemma converge_forever : forall d c r config, SSYNC d -> contained c r config -> imprisoned c r (execute convergeR2 d config). -Proof using . +Proof using ltc_0_k. cofix Hcorec. intros d c r config [] Hrec. constructor. - apply Hrec. - apply Hcorec; auto using contained_next. @@ -218,7 +218,7 @@ Qed. (************************) Theorem convergence_FSYNC : solution_FSYNC convergeR2. -Proof using . +Proof using ltc_0_k. intros config d [Hfair ?]. exists (isobarycenter (elements (obs_from_config (Observation := set_observation) config 0))). intros ε Hε. diff --git a/CaseStudies/Exploration/ImpossibilityKDividesN.v b/CaseStudies/Exploration/ImpossibilityKDividesN.v index 905f3a5e30fdb4f0b4083a9ce2db85229eab420e..9ba8b338f923f2d5611e339694b069a37d8119f0 100644 --- a/CaseStudies/Exploration/ImpossibilityKDividesN.v +++ b/CaseStudies/Exploration/ImpossibilityKDividesN.v @@ -17,11 +17,9 @@ Typeclasses eauto := (bfs). Section Exploration. (** Given an abitrary ring *) -Context {_n : greater_than 2}. +Context {n : nat} {ltc_2_n : 2 <c n}. (** There are k good robots and no byzantine ones. *) -Context {_k : greater_than 0}. -Notation n := (@ub 2 _n). -Notation k := (@ub 0 _k). +Context {k : nat} {ltc_0_k : 0 <c k}. Instance Robots : Names := Robots k 0. (** Assumptions on the number of robots: it strictly divides the ring size. *) @@ -29,17 +27,17 @@ Hypothesis kdn : (n mod k = 0). Hypothesis k_inf_n : (k < n). Lemma local_subproof1 : n = k * (n / k). -Proof using kdn. apply Nat.div_exact. apply neq_ub_lb. apply kdn. Qed. +Proof using kdn ltc_0_k. apply Nat.div_exact. apply neq_u_l. apply kdn. Qed. Lemma local_subproof2 : n / k ≠0. -Proof using kdn. - intros Habs. eapply (@neq_ub_0 2 _n). rewrite local_subproof1, Habs. +Proof using kdn ltc_2_n ltc_0_k. + intros Habs. eapply @neq_u_0. apply ltc_2_n. rewrite local_subproof1, Habs. apply Nat.mul_0_r. Qed. Lemma local_subproof3 : ∀ m : nat, m < k -> m * (n / k) < n. -Proof using kdn. +Proof using kdn ltc_2_n ltc_0_k. intros * H. rewrite local_subproof1 at 2. apply mult_lt_compat_r. apply H. apply Nat.neq_0_lt_0, local_subproof2. Qed. @@ -77,19 +75,19 @@ Definition ref_config : configuration := Lemma ref_config_injective : Util.Preliminary.injective eq equiv (fun id => get_location (ref_config id)). -Proof using kdn. +Proof using kdn ltc_0_k. intros id1 id2. apply (no_byz id2), (no_byz id1). clear id1 id2. intros g1 g2 Heq. f_equal. unfold ref_config, create_ref_config in Heq. - eapply mod2fin_bounded_diff_inj in Heq. 2,3: split. 2,4: apply Nat.le_0_l. - apply Nat.mul_cancel_r, fin2nat_inj in Heq. apply Heq. apply local_subproof2. + eapply mod2fin_betweenI in Heq. 2,3: split. 2,4: apply Nat.le_0_l. + apply Nat.mul_cancel_r, fin2natI in Heq. apply Heq. apply local_subproof2. all: apply local_subproof3, fin_lt. Qed. (** Translating [ref_config] by multiples of [n / k] does not change its observation. *) Lemma obs_asbf_ref_config : forall g, - !! (map_config (asbf (create_ref_config g)â»Â¹) ref_config) == !! ref_config. -Proof using kdn. + !! (map_config (λ x, subf x (create_ref_config g)) ref_config) == !! ref_config. +Proof using kdn ltc_0_k. unfold obs_from_config, MultisetObservation.multiset_observation, MultisetObservation.make_multiset. intro g. f_equiv. rewrite 2 config_list_spec, 4 map_map. @@ -105,25 +103,25 @@ Proof using kdn. split; intros [id [Hpt _]]; revert Hpt; apply (no_byz id); clear id; intros g' Hpt. + pose (id' := subf g' g). change (fin k) with G in id'. exists (Good id'). split. 2:{ rewrite InA_Leibniz. apply In_names. } - rewrite <- Hpt. cbn-[subf create_ref_config]. split. 2: reflexivity. - unfold create_ref_config, Datatypes.id, id'. apply fin2nat_inj. - rewrite 2 subf2nat, 3 mod2fin2nat, Nat.add_mod_idemp_l, + rewrite <- Hpt. cbn-[create_ref_config BijectionInverse]. split. 2: reflexivity. + change G with (fin k) in *. unfold create_ref_config, Datatypes.id, id'. + apply fin2natI. rewrite 2 subf2nat, 3 mod2fin2nat, Nat.add_mod_idemp_l, 2 (Nat.mod_small (_ * _)). rewrite local_subproof1 at 3 5. rewrite <- Nat.mul_sub_distr_r, <- Nat.mul_add_distr_r, - Nat.mul_mod_distr_r. reflexivity. 1,5: apply neq_ub_0. + Nat.mul_mod_distr_r. reflexivity. 1,5: apply neq_u_0. apply local_subproof2. all: apply local_subproof3. apply fin_lt. apply mod2fin_lt. + pose (id' := addf g' g). change (fin k) with G in id'. exists (Good id'). split. 2:{ rewrite InA_Leibniz. apply In_names. } - rewrite <- Hpt. cbn. split. 2: reflexivity. unfold create_ref_config, id'. - apply fin2nat_inj. rewrite subf2nat, 3 mod2fin2nat, addf2nat, + rewrite <- Hpt. cbn. split. 2: reflexivity. unfold create_ref_config, id', Datatypes.id. + apply fin2natI. rewrite subf2nat, 3 mod2fin2nat, addf2nat, Nat.add_mod_idemp_l, 2 (Nat.mod_small (_ * _)). rewrite local_subproof1 at 2 4. rewrite <- Nat.mul_sub_distr_r, <- Nat.mul_add_distr_r, Nat.mul_mod_distr_r. apply Nat.mul_cancel_r. all: cycle 1. rewrite Nat.add_mod_idemp_l, <- Nat.add_assoc, le_plus_minus_r, <- Nat.add_mod_idemp_r, Nat.mod_same, Nat.add_0_r. apply Nat.mod_small. 4: apply Nat.lt_le_incl. - 2,3,5,6,10: apply neq_ub_0. 3,6: apply local_subproof2. + 2,3,5,6,10: apply neq_u_0. 3,6: apply local_subproof2. 3,4: apply local_subproof3. all: apply fin_lt. Qed. @@ -171,15 +169,16 @@ Section NoMove. Hypothesis Hmove : move == SelfLoop. Lemma round_id : round r da ref_config == ref_config. -Proof using Hmove kdn. +Proof using Hmove kdn ltc_0_k. rewrite FSYNC_round_simplify; try (now split); []. apply no_byz_eq. intro g. - cbn-[BijectionInverse equiv]. erewrite obs_from_config_ignore_snd, - obs_asbf_ref_config, Hmove, move_along_0. apply Bijection.retraction_section. + cbn-[create_ref_config Bijection.BijectionInverse equiv]. + erewrite transvE, asbfVE, obs_from_config_ignore_snd, obs_asbf_ref_config, Hmove, move_along_0. + rewrite Bijection.inv_inv, asbfE. apply subfVKV. Qed. Lemma NeverVisited_ref_config : ∀ e, e == execute r bad_demon ref_config -> exists pt, ~ Will_be_visited pt e. -Proof using Hmove kdn k_inf_n. +Proof using Hmove kdn k_inf_n ltc_0_k. intros e Heq_e. exists (mod2fin 1). intro Hl. induction Hl as [e [g Hvisited] | e Hlater IHvisited]. * rewrite Heq_e in Hvisited. cbn in Hvisited. @@ -188,20 +187,19 @@ Proof using Hmove kdn k_inf_n. rewrite Nat.mul_mod_distr_r, (Nat.mod_small 1), Nat.mul_eq_1. intros [_ Habs]. contradiction (lt_not_le _ _ k_inf_n). rewrite local_subproof1, Habs, Nat.mul_1_r. reflexivity. 3: apply local_subproof2. - 2: apply neq_ub_lb. erewrite <- smaller_bound_ub. apply lt_lb_ub. - Unshelve. auto. + 2: apply neq_u_l. eapply @lt_l_u. apply lt_s_u. auto. * apply IHvisited. rewrite Heq_e, execute_tail. rewrite round_id. f_equiv. Qed. Lemma never_visited : ~(∀ pt, Will_be_visited pt (execute r bad_demon ref_config)). -Proof using Hmove kdn k_inf_n. +Proof using Hmove kdn k_inf_n ltc_0_k. intros Hw. destruct (NeverVisited_ref_config (reflexivity _)) as [pt Hpt]. apply Hpt, Hw. Qed. Theorem no_exploration_idle : ~ Explore_and_Stop r. -Proof using Hmove k_inf_n kdn. +Proof using Hmove k_inf_n kdn ltc_0_k. intros Habs. destruct (Habs bad_demon ref_config) as [Hexpl _]. apply FSYNC_implies_Fair, FYSNC_setting. apply ref_config_injective. now apply never_visited. @@ -225,7 +223,9 @@ Lemma f_config_merge : ∀ config m1 m2, f_config (f_config config m1) m2 == f_config config (m1 + m2). Proof using k_inf_n kdn. intros. unfold f_config. rewrite map_config_merge; autoclass; []. - intros id. cbn-[equiv]. apply addm_comp. + intros id. cbn-[equiv]. rewrite 3 asbmE, <- (addm_mod (config id)), <- mod2fin2nat, + <- addmA, addm2nat, mod2fin2nat, Nat.add_mod_idemp_l. + apply addm_mod. apply neq_u_0. Qed. Lemma f_config_swap : ∀ config m1 m2, @@ -246,10 +246,10 @@ Qed. Lemma f_config_injective_N : ∀ config m1 m2, f_config config m1 == f_config config m2 -> m1 mod n == m2 mod n. -Proof using kdn k_inf_n. +Proof using kdn k_inf_n ltc_0_k. unfold f_config. intros * Heq. specialize (Heq (Good fin0)). - unshelve eapply (@addm_bounded_diff_inj 2 _n). 5: rewrite 2 addm_mod. - 5: apply Heq. 2,3: split. 2,4: apply Nat.le_0_l. all: apply mod2fin_lt. + eapply (@addm_betweenI 2 n ltc_2_n). 3: rewrite 2 addm_mod. + 3: apply Heq. all: split. 1,3: apply Nat.le_0_l. all: apply mod2fin_lt. Qed. Lemma f_config_modulo : ∀ config m, @@ -271,11 +271,11 @@ Definition equiv_config config1 config2 : Prop Lemma equiv_config_m_sym : ∀ m config1 config2, equiv_config_m m config1 config2 - -> equiv_config_m (@oppm 2 _n m) config2 config1. + -> equiv_config_m (@oppm 2 n ltc_2_n m) config2 config1. Proof using k_inf_n kdn. unfold equiv_config_m. intros * Hequiv. unshelve erewrite Hequiv, f_config_merge, <- f_config_modulo, (proj2 (Nat.mod_divide _ _ _)). - apply neq_ub_0. symmetry. apply f_config_0. apply oppm_divide. + apply neq_u_0. symmetry. apply f_config_0. apply divide_addn_oppm. Qed. Lemma equiv_config_m_trans : ∀ m1 m2 config1 config2 config3, @@ -289,7 +289,7 @@ Qed. Instance equiv_config_equiv : Equivalence equiv_config. Proof using k_inf_n kdn. split. + intro config. exists 0. unfold equiv_config_m. now rewrite f_config_0. - + intros config1 config2 [m Hm]. exists (@oppm 2 _n m). + + intros config1 config2 [m Hm]. exists (@oppm 2 n ltc_2_n m). apply (equiv_config_m_sym Hm). + intros ? ? ? [m1 Hm1] [m2 Hm2]. exists (m1 + m2). revert Hm1 Hm2. apply equiv_config_m_trans. @@ -328,10 +328,10 @@ Theorem equiv_config_m_round : ∀ m config1 config2, Proof using k_inf_n kdn. unfold equiv_config_m. intros * Hequiv. unfold f_config. rewrite 2 (FSYNC_round_simplify r _ FSYNC_one). intros id. apply (no_byz id). - clear id. intro g. cbn-[equiv BijectionInverse]. + clear id. intro g. cbn-[equiv Bijection.BijectionInverse]. setoid_rewrite <- config1_obs_equiv. 2,3: reflexivity. rewrite Hequiv. - cbn-[f_config equiv]. rewrite <- 2 subf_move_along'. - unfold f_config, map_config. cbn. rewrite 2 subfVKV, addm_move_along. + setoid_rewrite transvVE. rewrite 2 transvE, 2 asbfE, 2 asbfVE, asbmE. cbn-[f_config equiv]. + rewrite <- 2 subf_move_along'. unfold f_config, map_config. cbn. rewrite 2 subfVKV, addm_move_along. apply move_along_compat. reflexivity. apply (pgm_compat r), obs_from_config_compat. 2: reflexivity. intros id. apply subf_addm_addm. Qed. @@ -354,7 +354,7 @@ Proof using . Qed. Lemma AlwaysEquiv_sym : ∀ m (e1 e2 : execution), - AlwaysEquiv m e1 e2 -> AlwaysEquiv (@oppm 2 _n m) e2 e1. + AlwaysEquiv m e1 e2 -> AlwaysEquiv (@oppm 2 n ltc_2_n m) e2 e1. Proof using k_inf_n kdn. cofix Hcoind. intros m1 e1 e2 [Hnow Hlater]. constructor. + now apply equiv_config_m_sym. @@ -452,21 +452,25 @@ Hypothesis Hmove : move =/= SelfLoop. (** After a round, the configuration obtained from ref_config is equivalent to ref_config. *) Lemma round_simplify : round r da ref_config == f_config ref_config target. -Proof using k_inf_n kdn. +Proof using k_inf_n kdn ltc_0_k. rewrite (FSYNC_round_simplify _ _ FSYNC_one). apply no_byz_eq. intro g. - cbn-[BijectionInverse mod2fin equiv]. rewrite obs_from_config_ignore_snd, - obs_asbf_ref_config. cbn-[mod2fin equiv]. unfold target. rewrite subff, - 2 move_along0_, addfC, addm_addf, mod2finK. reflexivity. + cbn-[Bijection.BijectionInverse mod2fin equiv]. + setoid_rewrite transvVE. rewrite transvE, asbfE, asbfVE. + rewrite obs_from_config_ignore_snd, obs_asbf_ref_config. + unfold target. rewrite subff, 2 move_along0_, addfC. unfold f_config. + rewrite asbmE. cbn. rewrite addm_addf, <- dir2nodE. reflexivity. Qed. Corollary round_ref_config : equiv_config_m target ref_config (round r da ref_config). -Proof using k_inf_n kdn. apply round_simplify. Qed. +Proof using k_inf_n kdn ltc_0_k. apply round_simplify. Qed. Corollary AlwaysEquiv_ref_config : AlwaysEquiv target (execute r bad_demon ref_config) (tl (execute r bad_demon ref_config)). -Proof using k_inf_n kdn. apply execute_equiv_compat, round_simplify. Qed. +Proof using k_inf_n kdn ltc_0_k. + apply execute_equiv_compat, round_simplify. +Qed. (** An execution that is always moving cannot stop. *) Lemma AlwaysMoving_not_WillStop : ∀ e, AlwaysMoving e -> ~Will_stop e. @@ -477,14 +481,14 @@ Qed. (** The starting configuration is always moving. *) Lemma ref_config_AlwaysMoving : AlwaysMoving (execute r bad_demon ref_config). -Proof using Hmove k_inf_n kdn. +Proof using Hmove k_inf_n kdn ltc_0_k. generalize (AlwaysEquiv_refl (execute r bad_demon ref_config)). generalize 0 at 1. generalize (execute r bad_demon ref_config) at 1 3. cofix Hcoind. intros e m Hequiv. constructor. + clear Hcoind. rewrite Hequiv. intros [Hstop _]. contradiction Hmove. unfold Stall in Hstop. rewrite execute_tail, round_simplify in Hstop. simpl hd in Hstop. apply (move_along_I fin0). rewrite move_along_0. - apply fin2nat_inj. setoid_rewrite <- Nat.mod_small. 2,3: apply fin_lt. + apply fin2natI. setoid_rewrite <- Nat.mod_small. 2,3: apply fin_lt. eapply f_config_injective_N. rewrite fin02nat, f_config_0. symmetry. apply Hstop. + eapply (Hcoind _ (addm (oppf target) m)). clear Hcoind. rewrite addm2nat, @@ -496,7 +500,7 @@ Qed. (** Final theorem when robots move. *) Theorem no_exploration_moving : ~ Explore_and_Stop r. -Proof using Hmove k_inf_n kdn. +Proof using Hmove k_inf_n kdn ltc_0_k. intros Habs. unfold Explore_and_Stop in *. destruct (Habs bad_demon ref_config) as [_ Hstop]. apply FSYNC_implies_Fair. apply FYSNC_setting. unfold Valid_starting_config. apply ref_config_injective. @@ -509,7 +513,7 @@ End DoesMove. In the asynchronous model, if the number of robots [kG] divides the size [n] of the ring, then the exploration with stop of a n-node ring is not possible. *) Theorem no_exploration : ~ Explore_and_Stop r. -Proof using k_inf_n kdn. +Proof using k_inf_n kdn ltc_0_k. destruct (move =?= SelfLoop) as [Hmove | Hmove]. + now apply no_exploration_idle. + now apply no_exploration_moving. diff --git a/CaseStudies/Exploration/Tower.v b/CaseStudies/Exploration/Tower.v index 316d99995e38a08b82c203f56d1df3433840de28..d40bcb1a4727afcb36a096c52edb36ec5fb1662c 100644 --- a/CaseStudies/Exploration/Tower.v +++ b/CaseStudies/Exploration/Tower.v @@ -27,11 +27,9 @@ Typeclasses eauto := (bfs). Section Tower. (** Given an abitrary ring *) -Context {_n : greater_than 2}. +Context {n : nat} {ltc_2_n : 2 <c n}. (** There are k good robots and no byzantine ones. *) -Context {_k : greater_than 0}. -Notation n := (@ub 2 _n). -Notation k := (@ub 0 _k). +Context {k : nat} {ltc_0_k : 0 <c k}. Instance Robots : Names := Robots k 0. (** Assumptions on the number of robots: it is non zero and strictly divides the ring size. *) diff --git a/CaseStudies/Gathering/InR2/FSyncFlexNoMultAlgorithm.v b/CaseStudies/Gathering/InR2/FSyncFlexNoMultAlgorithm.v index 285174c16c6038ea4c734291618ff70c32c2640c..d3a78325e9765916b40c8ca6c1b9d7e9d3558fae 100644 --- a/CaseStudies/Gathering/InR2/FSyncFlexNoMultAlgorithm.v +++ b/CaseStudies/Gathering/InR2/FSyncFlexNoMultAlgorithm.v @@ -42,10 +42,10 @@ Typeclasses eauto := (bfs) 10. (** ** Framework of the correctness proof: a finite set with at least two elements **) Section GatheringInR2. -Context {gt1 : greater_than 1}. +Context {k : nat} {ltc_1_k : 1 <c k}. (** There are n good robots and no byzantine ones. *) -Instance MyRobots : Names := Robots ub 0. +Instance MyRobots : Names := Robots k 0. Instance NoByz : NoByzantine. Proof using . now split. Qed. @@ -141,14 +141,14 @@ Proof using . intros. eapply map_injective_elements; autoclass; apply Similarity (** Spectra can never be empty as the number of robots is non null. *) Lemma obs_non_nil : forall config, !! config =/= empty. -Proof using . +Proof using ltc_1_k. intros config Habs. specialize (Habs (get_location (config (Good g1)))). rewrite empty_spec in Habs. rewrite <- Habs. apply pos_in_config. Qed. Lemma elements_non_nil : forall config, elements (!! config) <> nil. -Proof using . intro. rewrite elements_nil. apply obs_non_nil. Qed. +Proof using ltc_1_k. intro. rewrite elements_nil. apply obs_non_nil. Qed. (* We need to unfold [obs_is_ok] for rewriting *) Definition obs_from_config_spec : forall (config : configuration) pt, @@ -215,7 +215,7 @@ Instance measure_compat : Proper (equiv ==> eq) measure. Proof using . intros ? ? Heq. unfold measure. now rewrite Heq. Qed. Lemma measure_nonneg : forall config, 0 <= measure config. -Proof using . +Proof using ltc_1_k. intros config. unfold measure. destruct (elements (!! config)) as [| pt l] eqn:Heq. + contradiction (elements_non_nil _ Heq). @@ -225,7 +225,7 @@ Qed. (** The minimum value 0 is reached only on gathered configurations. *) Lemma gathered_elements : forall config pt, gathered_at pt config <-> PermutationA (@equiv _ location_Setoid) (elements (!! config)) (pt :: nil). -Proof using . +Proof using ltc_1_k. intros config pt. split; intro H. * apply NoDupA_equivlistA_PermutationA; autoclass. @@ -247,7 +247,7 @@ split; intro H. Qed. Lemma gathered_measure : forall config, measure config = 0%R <-> exists pt, gathered_at pt config. -Proof using . +Proof using ltc_1_k. intros config. split; intro H. * unfold measure, max_dist_obs in *. assert (Hnil := elements_non_nil config). @@ -351,7 +351,7 @@ Theorem round_lt_config : forall da config, FSYNC_da da -> delta <= measure config -> measure (round ffgatherR2 da config) <= measure config - delta. -Proof using . +Proof using ltc_1_k. intros da config Hdelta HFSync Hnotdone. set (elems := elements (!! config)). set (C := isobarycenter elems). @@ -568,7 +568,7 @@ Theorem round_last_step : forall da config, FSYNC_da da -> measure config <= delta -> measure (round ffgatherR2 da config) == 0. -Proof using . +Proof using ltc_1_k. intros da config Hdelta HFSync Hlt. unfold measure. set (elems := (elements (!! config))). @@ -628,7 +628,7 @@ Proof using . unfold lt_config. apply wf_inverse_image, lt_wf. Qed. Lemma lt_config_decrease : 0 < delta -> forall config1 config2, measure config1 <= measure config2 - delta -> lt_config delta config1 config2. -Proof using . +Proof using ltc_1_k. intros Hdelta config1 config2 Hle. unfold lt_config. rewrite <- Z2Nat.inj_lt. + apply Zup_lt. field_simplify; try lra. unfold Rdiv. apply Rmult_le_compat. @@ -674,7 +674,7 @@ Qed. *) Lemma gathered_at_elements : forall config pt, gathered_at pt config -> PermutationA equiv (elements (!! config)) (pt :: nil). -Proof using . +Proof using ltc_1_k. intros config pt Hgather. apply NoDupA_equivlistA_PermutationA; autoclass. + apply elements_NoDupA. @@ -689,7 +689,7 @@ Qed. Lemma gathered_at_forever : forall da config pt, FSYNC_da da -> gathered_at pt config -> gathered_at pt (round ffgatherR2 da config). -Proof using . +Proof using ltc_1_k. intros da config pt Hda Hgather g. rewrite round_simplify; trivial; []. cbn zeta. rewrite (gathered_at_elements Hgather), isobarycenter_singleton. @@ -698,7 +698,7 @@ Qed. Lemma gathered_at_OK : forall d conf pt, FSYNC (similarity_demon2demon d) -> gathered_at pt conf -> Gather pt (execute ffgatherR2 d conf). -Proof using . +Proof using ltc_1_k. cofix Hind. intros d conf pt [] Hgather. constructor. + clear Hind. simpl. assumption. + rewrite execute_tail. apply Hind; now try apply gathered_at_forever. @@ -733,7 +733,7 @@ Qed. (** The final theorem. *) Theorem FSGathering_in_R2 : forall d, delta > 0 -> FSYNC (similarity_demon2demon d) -> FullSolGathering ffgatherR2 d. -Proof using . +Proof using ltc_1_k. intros d Hdelta HFS config. revert d HFS. pattern config. apply (well_founded_ind (wf_lt_config Hdelta)). clear config. intros config Hind [da d] [Hda HFS]. diff --git a/CaseStudies/Gathering/InR2/Viglietta.v b/CaseStudies/Gathering/InR2/Viglietta.v index 45b0f3bfb005bd0b49255094980872991fab465b..666c6d8e52570c04a633e92c43bd4db15aeac2cc 100644 --- a/CaseStudies/Gathering/InR2/Viglietta.v +++ b/CaseStudies/Gathering/InR2/Viglietta.v @@ -43,7 +43,7 @@ Lemma id_case : forall id, id = Good r0 \/ id = Good r1. Proof using . intros [[[| [| ?]] ?] | []]; simpl; solve [ exfalso; lia - | now left + right; f_equal; apply fin2nat_inj ]. + | now left + right; f_equal; apply fin2natI ]. Qed. (** The space is R², so that we can define similarities and they respect middle points. *) @@ -369,9 +369,9 @@ destruct (config (Good r0)) as [pt1 l1] eqn:Hr0, assert (Hpt : pt1 =/= pt2). { intro Habs. rewrite Habs in Hr0. apply Hnotgather. exists (get_location (config (Good r0))). intros [[| [| ?]] ?]; simpl. - + unfold r0. do 4 f_equal. now apply fin2nat_inj. + + unfold r0. do 4 f_equal. now apply fin2natI. + transitivity (Datatypes.id (fst (config (Good r1)))). - - unfold r1. simpl. do 4 f_equal. now apply fin2nat_inj. + - unfold r1. simpl. do 4 f_equal. now apply fin2natI. - now rewrite Hr0, Hr1. + exfalso. lia. } destruct l1 eqn:Hl1; [| destruct l2 eqn:Hl2]. diff --git a/CaseStudies/LifeLine/Algorithm.v b/CaseStudies/LifeLine/Algorithm.v index eeb4bce139872c6d043cf440ebeaf1f0501cad72..0f2f06728e5712f25794f199fdc94abbeea0352c 100644 --- a/CaseStudies/LifeLine/Algorithm.v +++ b/CaseStudies/LifeLine/Algorithm.v @@ -42,7 +42,7 @@ Definition scout : ident := Byz scout_B. (** The only Byzantine robot is the scout. *) Lemma byz_is_scout : ∀ b : B, b = scout_B. Proof using . - intros [b Hb]. change B with (fin 1) in *. apply fin2nat_inj. cbn. + intros [b Hb]. change B with (fin 1) in *. apply fin2natI. cbn. destruct b as [|b]. reflexivity. inv Hb. inv H0. Qed. diff --git a/Models/ContinuousGraph.v b/Models/ContinuousGraph.v index e7b3f8b71044a59b5cb13a3a6b5fe683ef33eff5..b18a6acaebdc700a3deea32ca04214b1dbfb6401 100644 --- a/Models/ContinuousGraph.v +++ b/Models/ContinuousGraph.v @@ -17,29 +17,23 @@ *) (**************************************************************************) -Require Import Utf8_core. -Require Import Arith_base. -Require Import Reals. -Require Import Lia. -Require Import Psatz. -Require Import SetoidList. -Require Import Pactole.Setting. -Require Import Pactole.Spaces.Graph. -Require Import Pactole.Spaces.Isomorphism. +Require Import Utf8_core Arith_base Reals Lia Psatz SetoidList. +From Pactole Require Import Setting Util.SetoidDefs Util.Bijection. +From Pactole Require Import Spaces.Graph Spaces.Isomorphism + Spaces.ThresholdIsomorphism. Require Import Pactole.Models.Flexible. Notation "x == y" := (equiv x y). Typeclasses eauto := (bfs). -Remove Hints eq_setoid : typeclass_instances. Section CGF. Context {V E : Type}. -Context `{Names}. -Context {G : Graph V E}. +Context {N : Names}. +Context {TG : ThresholdGraph V E}. -Instance LocationV : Location := { location := V }. +Instance LocationV : Location := NodesLoc. (** We do not want to use the default equivalence on [E] because we only need equality of source, target and threshold on the edge. *) @@ -51,7 +45,10 @@ Global Instance E_src_tgt_thd_EqDec : EqDec E_src_tgt_thd_Setoid := (precompose_EqDec threshold). Global Instance E_subrelation : subrelation (@equiv E E_Setoid) (@equiv E E_src_tgt_thd_Setoid). -Proof using . intros ? ? Heq. split; simpl; now rewrite Heq. Qed. +Proof using . intros ? ? Heq; split; simpl; now rewrite Heq. Qed. + +Global Instance E_subrelation' : subrelation (@equiv E E_src_tgt_thd_Setoid) (@equiv E E_Setoid). +Proof using . intros ?? [H _]. apply simple_graph, H. Qed. Global Instance src_compat : Proper (equiv ==> equiv) src. Proof using . intros ? ? Heq. apply Heq. Qed. @@ -64,14 +61,14 @@ Proof using . intros ? ? Heq. apply Heq. Qed. (* Since iso_E gives a bijection that comes with its setoid, we need to be lower level to change it from [E_Setoid] to [E_src_tgt_thd_Setoid]. *) -Global Instance iso_E_compat : forall iso, - Proper (equiv ==> equiv) (iso_E iso). +Global Instance iso_E_compat : + ∀ iso : threshold_isomorphism TG, Proper (equiv ==> equiv) (iso_E iso). Proof using . -intros iso ? ? [[Hsrc Htgt] Hthd]. -repeat split; unfold equiv in *; cbn -[equiv] in *. -- now rewrite <- 2 (proj1 (iso_morphism _ _)), Hsrc. -- now rewrite <- 2 (proj2 (iso_morphism _ _)), Htgt. -- rewrite <- 2 iso_threshold. now f_equiv. + intros iso ? ? [[Hsrc Htgt] Hthd]. + repeat split; unfold equiv in *; cbn -[equiv] in *. + - now rewrite <- 2 (proj1 (iso_morphism _ _)), Hsrc. + - now rewrite <- 2 (proj2 (iso_morphism _ _)), Htgt. + - rewrite <- iso_threshold, Hthd. apply iso_threshold. Qed. @@ -87,20 +84,19 @@ simple refine {| equiv := fun l l' => | OnEdge e p, OnEdge e' p' => e == e' /\ p == p' | _, _ => False end |}; autoclass; []. -Proof. split. -+ now intros []. -+ intros [] [] Heq; simpl in *; decompose [False and] Heq; repeat split; now symmetry. -+ intros [] [] [] Heq1 Heq2; simpl in *; - decompose [False and] Heq1; decompose [False and] Heq2; repeat split; etransitivity; eauto. +Proof using . split. + + now intros []. + + intros [] [] Heq; simpl in *; decompose [False and] Heq; repeat split; now symmetry. + + intros [] [] [] Heq1 Heq2; simpl in *; + decompose [False and] Heq1; decompose [False and] Heq2; repeat split; etransitivity; eauto. Defined. Global Instance locG_EqDec: EqDec locG_Setoid. Proof using . - intros [l1 | e1 p1] [l2 | e2 p2]; simpl. -+ apply equiv_dec. -+ intuition. -+ intuition. -+ destruct (e1 =?= e2); [destruct (p1 =?= p2) |]; intuition. + intros [l1 | e1 p1] [l2 | e2 p2]. 2,3: right; intros H; inversion H. + cbn. apply equiv_dec. destruct (e1 =?= e2) as [H1 | H1]. + destruct (p1 =?= p2) as [H2 | H2]; [left | right]. split. apply H1. apply H2. + intros [_ H3]. apply H2, H3. right. intros [H2 _]. apply H1, H2. Qed. Instance LocationG : Location := { location := loc }. @@ -115,53 +111,32 @@ Global Instance OnEdge_compat : Proper (equiv ==> equiv ==> equiv) OnEdge. Proof using . repeat intro. auto. Qed. (** We can use an isomorphism to build a bijection on a continuous graph. *) -Definition bijectionG (iso : isomorphism G) : Bijection.bijection loc. -simple refine {| Bijection.section := fun pt => match pt with +Definition bijectionG (iso : threshold_isomorphism TG) : bijection loc. +simple refine {| section := fun pt => match pt with | OnVertex v => OnVertex (iso.(iso_V) v) | OnEdge e p => OnEdge (iso.(iso_E) e) p end; - Bijection.retraction := fun pt => match pt with - | OnVertex v => OnVertex (Bijection.retraction iso.(iso_V) v) - | OnEdge e p => OnEdge (Bijection.retraction iso.(iso_E) e) p + retraction := fun pt => match pt with + | OnVertex v => OnVertex (retraction iso.(iso_V) v) + | OnEdge e p => OnEdge (retraction iso.(iso_E) e) p end |}. -Proof. -* intros [| e1 p1] [| e2 p2] Heq; simpl in Heq; trivial. - + now repeat f_equiv. - + destruct Heq as [[[Hsrc Htgt] Hthd] Hp]. repeat split; simpl. - - rewrite <- (proj1 (iso_morphism iso e1)), Hsrc. apply iso_morphism. - - rewrite <- (proj2 (iso_morphism iso e1)), Htgt. apply iso_morphism. - - now rewrite <- 2 iso_threshold, Hthd. - - assumption. -* intros [| e1 p1] [| e2 p2] ; simpl in *; try tauto; [|]; split; intro Heq. - + rewrite <- Heq. apply Bijection.retraction_section. - + rewrite <- Heq. apply Bijection.section_retraction. - + destruct Heq as [[[Hsrc Htgt] Hthd] Hp]. repeat split. - - change (Bijection.retraction (iso_E iso)) with (Bijection.section (iso_E (inverse iso))). - rewrite <- (proj1 (iso_morphism _ e2)). simpl. - now rewrite <- (Bijection.Inversion iso), (proj1 (iso_morphism _ e1)). - - change (Bijection.retraction (iso_E iso)) with (Bijection.section (iso_E (inverse iso))). - rewrite <- (proj2 (iso_morphism _ e2)). simpl. - now rewrite <- (Bijection.Inversion iso), (proj2 (iso_morphism _ e1)). - - change (Bijection.retraction (iso_E iso)) with (Bijection.section (iso_E (inverse iso))). - rewrite <- (iso_threshold _ e2), <- Hthd, iso_threshold. - simpl. now rewrite Bijection.retraction_section. - - auto. - + destruct Heq as [[[Hsrc Htgt] Hthd] Hp]. repeat split. - - now rewrite <- (proj1 (iso_morphism _ e1)), <- Hsrc, (proj1 (iso_morphism _ _)), - Bijection.section_retraction. - - now rewrite <- (proj2 (iso_morphism _ e1)), <- Htgt, (proj2 (iso_morphism _ _)), - Bijection.section_retraction. - - now rewrite <- (iso_threshold _ e1), <- Hthd, (iso_threshold _ _), - Bijection.section_retraction. - - auto. +Proof using . + * intros [l1 | e1 p1] [l2 | e2 p2] H. 2,3: inversion H. + + cbn in H. rewrite H. reflexivity. + + destruct H as [H1 H2]. rewrite H1, H2. reflexivity. + * intros [l1 | e1 p1] [l2 | e2 p2]. all: split. all: intros H. 3-6: inversion H. + + cbn in H. rewrite <- H, retraction_section. reflexivity. + + cbn in H. rewrite <- H, section_retraction. reflexivity. + + destruct H as [H1 H2]. erewrite (proj1 (Inversion _ _ _)). + rewrite H2. reflexivity. apply E_subrelation', H1. + + destruct H as [H1 H2]. erewrite (proj2 (Inversion _ _ _)). + rewrite H2. reflexivity. apply E_subrelation', H1. Defined. Global Instance bijectionG_compat : Proper (equiv ==> equiv) bijectionG. Proof using . -intros iso1 iso2 Hiso []; simpl. -+ apply Hiso. -+ repeat split; apply Graph.src_compat || apply Graph.tgt_compat - || apply Graph.threshold_compat; apply Hiso. + intros iso1 iso2 Hiso [l | e p]. apply Hiso. split. + apply E_subrelation, Hiso. reflexivity. Qed. (** ** Translation of locations *) @@ -169,23 +144,15 @@ Qed. Definition location_G2V (loc : locG) : locV := match loc with | OnVertex l => l - | OnEdge e p => if Rle_dec (threshold e) p then Graph.tgt e else Graph.src e + | OnEdge e p => if Rle_dec (threshold e) p then tgt e else src e end. Global Instance location_G2V_compat : Proper (equiv ==> equiv) location_G2V. Proof using . -unfold location_G2V. intros [l1 | e1 p1] [l2 | e2 p2] Hxy; try tauto; []. -destruct Hxy as [Hexy Hpxy], - (Rle_dec (threshold e1) p1) eqn:Hx, - (Rle_dec (threshold e2) p2) eqn:Hy. -+ apply Hexy. -+ assert (Ht := proj2 Hexy). - assert (Hr : (threshold e1 <= p1)%R) by assumption. - now rewrite Ht, Hpxy in Hr. -+ assert (Hr : (threshold e2 <= p2)%R) by assumption. - assert (Ht := proj2 Hexy). - now rewrite <- Ht, <- Hpxy in Hr. -+ apply Hexy. + unfold location_G2V. intros [l1 | e1 p1] [l2 | e2 p2] H. 2-3: inversion H. + apply H. destruct H as [He Hp], (Rle_dec (threshold e1) p1) as [H1 | H1], + (Rle_dec (threshold e2) p2) as [H2 | H2]. 1,4: apply He. contradict H2. + rewrite <- Hp, <- He. apply H1. contradict H1. rewrite Hp, He. apply H2. Qed. Definition location_V2G : locV -> locG := OnVertex. @@ -286,21 +253,22 @@ Proof using . intros ? ? []. unfold state_V2G. now split. Qed. Definition state_G2V (state : stateG) : stateV := match state with | SOnVertex v e p => exist valid_stateV (v, e) p - | SOnEdge e p => if Rle_dec (@threshold locV E G e) p + | SOnEdge e p => if Rle_dec (@threshold locV E TG e) p then exist valid_stateV (Graph.tgt e, e) ltac:(now right) else exist valid_stateV (Graph.src e, e) ltac:(now left) end. Global Instance state_G2V_compat : Proper (equiv ==> equiv) state_G2V. Proof using . -intros [v e p | e p] [v' e' p' | e' p'] Hstate; auto; []. -destruct Hstate as [[[Hsrc Htgt] Hthd] Hp]. simpl. -destruct (Rle_dec (threshold e) p), (Rle_dec (threshold e') p'); -repeat split; simpl in *; rewrite ?Hsrc, ?Htgt, ?Hthd in *; try reflexivity; [|]; -destruct p, p'; simpl in *; subst; contradiction. + intros [v1 e1 p1 | e1 p1] [v2 e2 p2 | e2 p2] H. + 2,3: inversion H. apply H. destruct H as [He Hp]. cbn[state_G2V]. + destruct (Rle_dec (threshold e1) p1) as [H1 | H1], + (Rle_dec (threshold e2) p2) as [H2 | H2]. all: split; cbn[proj1_sig fst snd]. + 1,2,4,6,7,8: apply He. contradict H2. rewrite <- Hp, <- He. apply H1. + contradict H1. rewrite Hp, He. apply H2. Qed. -Lemma state_V2G2V : forall state, state_G2V (state_V2G state) == state. +Lemma state_V2G2V : ∀ state, state_G2V (state_V2G state) == state. Proof using . intro. simpl. repeat (split; try reflexivity). Qed. (** ** On configurations *) @@ -308,31 +276,32 @@ Proof using . intro. simpl. repeat (split; try reflexivity). Qed. (** The precondition for liftable changes of frame is that they must come from isomorphisms (which must not change the thresholds). *) Local Instance InfoV : @State LocationV stateV. -simple refine {| - get_location := fun state => fst (proj1_sig state); - state_Setoid := stateV_Setoid; - precondition := fun f => sigT (fun iso => f == iso.(iso_V) /\ iso_T iso == @Bijection.id R _); - lift := fun f state => exist _ (projT1 f (fst (proj1_sig state)), - iso_E (projT1 (projT2 f)) (snd (proj1_sig state))) _ |}; autoclass. -Proof. -+ abstract (destruct f as [f [iso [Hiso ?]]], state as [state [Hcase | Hcase]]; +Proof using . + simple refine {| + get_location := λ state, fst (proj1_sig state); + state_Setoid := stateV_Setoid; + precondition := λ f, sigT (λ iso : threshold_isomorphism TG, + f == iso.(iso_V) /\ iso_T iso == Bijection.id); + lift := λ f state, exist _ (projT1 f (fst (proj1_sig state)), + iso_E (projT1 (projT2 f)) (snd (proj1_sig state))) _ |}; autoclass. + + abstract (destruct f as [f [iso [Hiso ?]]], state as [state [Hcase | Hcase]]; cbn; left + right; rewrite Hiso, Hcase; cbn; apply iso_morphism). -+ intros ? ? Heq. apply Heq. -+ (* lift_compat *) - intros [f [iso1 [Hiso1 Ht1]]] [g [iso2 [Hiso2 Ht2]]] Heq [] [] [Heq1 [Heq2 Heq3]]. - cbn in *. repeat split. - - now apply Heq. - - rewrite <- (proj1 (iso_morphism iso1 _)), <- Hiso1, - <- (proj1 (iso_morphism iso2 _)), <- Hiso2. - now apply Heq. - - rewrite <- (proj2 (iso_morphism iso1 _)), <- Hiso1, - <- (proj2 (iso_morphism iso2 _)), <- Hiso2. - now apply Heq. - - now rewrite <- 2 iso_threshold, Ht1, Ht2. + + intros ? ? Heq. apply Heq. + + (* lift_compat *) + intros [f [iso1 [Hiso1 Ht1]]] [g [iso2 [Hiso2 Ht2]]] Heq [] [] [Heq1 [Heq2 Heq3]]. + cbn in *. repeat split. + - now apply Heq. + - rewrite <- (proj1 (iso_morphism iso1 _)), <- Hiso1, + <- (proj1 (iso_morphism iso2 _)), <- Hiso2. + now apply Heq. + - rewrite <- (proj2 (iso_morphism iso1 _)), <- Hiso1, + <- (proj2 (iso_morphism iso2 _)), <- Hiso2. + now apply Heq. + - now rewrite <- 2 iso_threshold, Ht1, Ht2. Defined. Definition good_iso_of f iso := f == Bijection.section (bijectionG iso) - /\ iso_T iso == @Bijection.id R _. + /\ iso_T iso == Bijection.id. Definition preconditionG := fun f => sigT (good_iso_of f). Definition liftG (f : sigT preconditionG) state := @@ -471,25 +440,24 @@ Proof using . intro. simpl. repeat split; reflexivity. Qed. (** ** Demonic schedulers **) (** Acceptable frame changes must not change the thresholds. *) -Definition stable_threshold iso := iso_T iso == @Bijection.id R _. +Definition stable_threshold iso := iso_T iso == Bijection.id. -Definition stable_threshold_inverse : forall iso, +Definition stable_threshold_inverse : ∀ iso, stable_threshold iso -> stable_threshold (inverse iso). Proof using . -intros iso Hstable x. unfold stable_threshold in *. simpl in *. -now rewrite <- (Hstable x), Bijection.retraction_section. + unfold stable_threshold. intros iso Hstable x. apply (injective (iso_T iso)). + rewrite Hstable at 2. apply section_retraction. Qed. (** Frame choice: graph isomorphisms not changing thresholds *) Global Instance FrameChoiceIsomorphismV : @frame_choice LocationV (sig stable_threshold) := {| - frame_choice_bijection := fun f => @iso_V locV E G (proj1_sig f); - frame_choice_Setoid := sig_Setoid (@isomorphism_Setoid locV E G); - frame_choice_bijection_compat := - fun f g => @iso_V_compat locV E G (proj1_sig f) (proj1_sig g) |}. + frame_choice_bijection := λ f : sig stable_threshold, @iso_V locV E TG (proj1_sig f); + frame_choice_Setoid := sig_Setoid (@threshold_isomorphism_Setoid locV E TG); + frame_choice_bijection_compat := ltac:(intros ?? H; apply H) |}. Global Instance FrameChoiceIsomorphismG : @frame_choice LocationG (sig stable_threshold) := {| frame_choice_bijection := fun f => bijectionG (proj1_sig f); - frame_choice_Setoid := sig_Setoid (@isomorphism_Setoid locV E G); + frame_choice_Setoid := sig_Setoid (@threshold_isomorphism_Setoid locV E TG); frame_choice_bijection_compat := fun f g => bijectionG_compat (proj1_sig f) (proj1_sig g) |}. (** The demon update choice only contains the movement ratio, either a boolean or a ratio. *) diff --git a/Models/GraphEquivalence.v b/Models/GraphEquivalence.v index 36e5d6e21841716e83379cd502423707bd4bfe39..bbe2bc1b9a1a6c267160ef2ff1308a7adbd284a4 100644 --- a/Models/GraphEquivalence.v +++ b/Models/GraphEquivalence.v @@ -20,6 +20,7 @@ Require Import SetoidList. Require Import Pactole.Setting. Require Import Pactole.Spaces.Graph. Require Import Pactole.Spaces.Isomorphism. +Require Import Pactole.Spaces.ThresholdIsomorphism. Require Import Pactole.Models.ContinuousGraph. Require Import Pactole.Models.DiscreteGraph. @@ -32,7 +33,7 @@ Section GraphEquivalence. Context (V E : Type). Context {NN : Names}. -Context {G : Graph V E}. +Context {TG : ThresholdGraph V E}. (** We assume that the graph contains loops from each node to itself. *) Variable self_loop : V -> E. @@ -147,24 +148,18 @@ simple refine {| end in let e := snd (proj1_sig (config id)) in Rle_bool (threshold e) (current_ratio + da.(choose_inactive) (config_V2G config) id) |}; -try exact G; autoclass. -Proof. -+ intros config g. exists (proj1_sig (change_frame da (config_V2G config) g)). - split; try reflexivity; []. apply proj2_sig. -+ intros config g. exists (inverse (proj1_sig (change_frame da (config_V2G config) g))). - split; try reflexivity; []. apply stable_threshold_inverse, proj2_sig. -+ intros config1 config2 Hconfig gg g ?. subst gg. now rewrite Hconfig. -+ intros config1 config2 Hconfig gg g ? pt1 pt2 Hpt. - f_equiv; try apply Hpt; []. - f_equiv. now apply (choose_update_compat da); f_equiv. -+ intros config1 config2 Hconfig id1 id2 Hid. simpl in Hid. subst id1. - assert (Hpt := Hconfig id2). - destruct Hpt as [Hpt [[Hsrc Htgt] Hthd]], - (proj1_sig (config1 id2)) as [pt1 e1], - (proj1_sig (config2 id2)) as [pt2 e2]. - simpl in Hpt, Hsrc, Htgt, Hthd. - destruct_match; simpl; rewrite Hthd; do 3 f_equiv; - apply (choose_inactive_compat da); trivial; reflexivity. +try exact TG; autoclass. +Proof using . + + intros config g. exists (proj1_sig (change_frame da (config_V2G config) g)). + split; try reflexivity; []. apply proj2_sig. + + intros config g. exists (inverse (proj1_sig (change_frame da (config_V2G config) g))). + split; try reflexivity; []. apply stable_threshold_inverse, proj2_sig. + + intros config1 config2 Hconfig gg g ?. subst gg. now rewrite Hconfig. + + intros config1 config2 Hconfig gg g ? pt1 pt2 Hpt. + rewrite Hpt, Hconfig, H. reflexivity. + + intros config1 config2 Hconfig id1 id2 Hid. simpl in Hid. subst id1. + destruct (Cconfig id2) as [v1 e1 p1 | e1 p1]. all: cbn-[ratio_0]. + all: rewrite Hconfig. all: reflexivity. Defined. Instance da_C2D_compat : Proper (equiv ==> equiv ==> equiv) da_C2D. @@ -371,15 +366,14 @@ simpl activate. destruct_match. * cbn -[equiv precondition_satisfied]. rewrite Hconfig. simpl snd. transitivity (iso_E Ciso e); [| transitivity (iso_E Diso e)]. + f_equiv. - + apply E_subrelation, (Isomorphism.iso_E_compat Hiso e). + + apply E_subrelation, Hiso. + repeat split. - unfold equiv. cbn -[equiv]. now rewrite <- (proj1 (iso_morphism Diso e)), <- HDiso, (proj1 (iso_morphism _ e)). - unfold equiv. cbn -[equiv]. now rewrite <- (proj2 (iso_morphism Diso e)), <- HDiso, (proj2 (iso_morphism _ e)). - - unfold equiv. cbn -[equiv]. - rewrite <- 2 iso_threshold. unfold Diso. rewrite (proj2_sig Dframe_choice). - destruct (precondition_satisfied da config g) as [? [? Ht]]. simpl. now rewrite Ht. } + - cbn. rewrite <- 2 iso_threshold. unfold Diso. rewrite (proj2_sig Dframe_choice). + destruct (precondition_satisfied da config g) as [? [? Ht]]. simpl. now rewrite Ht. } assert (Hlocal_state : Clocal_state == state_V2G Dlocal_state). { unfold Clocal_state. rewrite Hlocal_config. reflexivity. } assert (Hobs : Cobs == Dobs). @@ -551,9 +545,9 @@ simpl activate. destruct_match_eq Hactive. assert (HDiso : projT1 (precondition_satisfied (da_C2D da config) (config_G2V config) g) == Diso). { destruct (precondition_satisfied (da_C2D da config) (config_G2V config) g) as [Diso' [HDf HDt]] eqn:HDiso. simpl projT1. - pose (iso_OKV := fun f (iso : @isomorphism (@location LocationV) E G) => + pose (iso_OKV := fun f (iso : @threshold_isomorphism (@location LocationV) E TG) => frame_choice_bijection f == iso.(iso_V) - /\ iso_T iso == @Bijection.id R R_Setoid). + /\ iso_T iso == Bijection.id). change (projT1 (existT (iso_OKV _) Diso' (conj HDf HDt)) == proj1_sig Dframe_choice). fold (iso_OKV (change_frame (da_C2D da config) (config_G2V config) g)) in HDiso. change (precondition_satisfied (da_C2D da config) (config_G2V config) g = @@ -573,22 +567,18 @@ simpl activate. destruct_match_eq Hactive. simpl state_G2V. repeat split. + simpl fst. rewrite Hframe. unfold config_G2V. transitivity (Ciso v); reflexivity || now symmetry; apply (HCiso (OnVertex v)). - + cbn -[precondition_satisfied]. - rewrite <- 2 (proj1 (iso_morphism _ _)), - (iso_V_compat HDiso (src e)), - (iso_V_compat Hiso (src e)). - symmetry. apply (HCiso (OnVertex (src e))). - + cbn -[precondition_satisfied]. - rewrite <- 2 (proj2 (iso_morphism _ _)), - (iso_V_compat HDiso (tgt e)), - (iso_V_compat Hiso (tgt e)). - symmetry. apply (HCiso (OnVertex (tgt e))). + + cbn-[precondition_satisfied]. rewrite <- 2 (proj1 (iso_morphism _ _)). + etransitivity. 2: symmetry; apply (HCiso (OnVertex (src e))). + etransitivity. 2: apply Hiso. apply HDiso. + + cbn-[precondition_satisfied]. rewrite <- 2 (proj2 (iso_morphism _ _)). + etransitivity. 2: symmetry; apply (HCiso (OnVertex (tgt e))). + etransitivity. 2: apply Hiso. apply HDiso. + cbn -[precondition_satisfied]. rewrite <- 2 iso_threshold. now rewrite (proj2 (projT2 (precondition_satisfied da config g))), (proj2 (projT2 (precondition_satisfied (da_C2D da config) (config_G2V config) g))). * (* OnEdge *) simpl liftG. simpl state_G2V. - assert (Heq : threshold ((iso_E (projT1 (precondition_satisfied da config g))) e) = threshold e). + assert (Heq : threshold ((iso_E (projT1 (precondition_satisfied da config g))) e) == threshold e). { now rewrite <- iso_threshold, (proj2 (projT2 (precondition_satisfied da config g))). } destruct (Rle_dec (threshold e) p); destruct_match; try (rewrite <- Heq in *; contradiction); [|]. @@ -639,7 +629,7 @@ simpl activate. destruct_match_eq Hactive. then true else false). { unfold Dchoice, choose_update, da_C2D. rewrite Hlocal_config, Hchoose_update. rewrite Hlocal_robot_decision at 2. - assert (Hthd := proj2 Hlocal_robot_decision). hnf in Hthd. rewrite Hthd. + rewrite Hlocal_robot_decision. destruct (Rle_dec (threshold Clocal_robot_decision) Cchoice) as [Hle | Hlt]. - rewrite <- Rle_bool_true_iff in Hle. apply Hle. - rewrite <- Rle_bool_false_iff in Hlt. apply Hlt. } @@ -695,12 +685,12 @@ simpl activate. destruct_match_eq Hactive. assert (H0 : proj_ratio Cchoice = 0%R). { assert (Hbounds := ratio_bounds Cchoice). simpl in *; lra. } assert (proj_ratio Cchoice < threshold Clocal_robot_decision)%R. - { rewrite H0. apply threshold_pos. } + { rewrite H0. apply strict_ratio_bounds. } destruct_match; try lra; []. symmetry. split; apply Hlocal_robot_decision. + (* the robot reaches its destination *) change (proj_ratio ratio_0) with 0%R in *. rewrite Rplus_0_l in *. assert (threshold Clocal_robot_decision <= proj_ratio Cchoice)%R. - { transitivity 1; trivial; []. apply Rlt_le. apply threshold_pos. } + { transitivity 1; trivial; []. apply Rlt_le. apply strict_ratio_bounds. } destruct_match; try lra; []. symmetry. hnf. simpl fst. simpl snd. split; apply Hlocal_robot_decision. + (* the robot moves and ends up on the edge *) @@ -735,15 +725,13 @@ simpl activate. destruct_match_eq Hactive. change (Bijection.retraction (iso_E Diso)) with (Bijection.section (iso_E (inverse Diso))). symmetry. etransitivity; try apply (iso_E_compat (inverse Ciso)), He; []. apply E_subrelation. f_equiv. - apply Isomorphism.iso_E_compat. clear -Hiso. now apply inverse_compat. + apply Isomorphism.iso_E_compat. clear -Hiso. apply inverse_compat, Hiso. -- simpl snd. etransitivity; try apply HCisoE; []. change (Bijection.retraction (iso_E Diso)) with (Bijection.section (iso_E (inverse Diso))). symmetry. etransitivity; try apply (iso_E_compat (inverse Ciso)), He; []. - apply E_subrelation. f_equiv. - now apply Isomorphism.iso_E_compat, inverse_compat. - -- unfold equiv. cbn -[equiv Dnew_local_state inverse]. - change (iso_E Diso â»Â¹) with (iso_E (Diso â»Â¹)). - rewrite <- 2 iso_threshold. + apply E_subrelation. f_equiv. apply Isomorphism.iso_E_compat, inverse_compat, Hiso. + -- cbn-[Dnew_local_state IsoInverse]. change ((iso_VE Diso)â»Â¹) with (iso_VE (Diso â»Â¹)). + rewrite <-2 iso_threshold. rewrite (proj2 (projT2 (precondition_satisfied_inv da config g))), (proj2 (projT2 (precondition_satisfied_inv (da_C2D da config) (config_G2V config) g))). now rewrite He. @@ -751,7 +739,7 @@ simpl activate. destruct_match_eq Hactive. destruct Dnew_local_state as [[v' e'] Hvalid]. unfold state_G2V in *. simpl liftG. cbn iota beta. simpl snd. assert (Htest : threshold - (Bijection.section (iso_E (projT1 (precondition_satisfied_inv da config g))) e) = threshold e) + (Bijection.section (iso_E (projT1 (precondition_satisfied_inv da config g))) e) == threshold e) by now rewrite <- iso_threshold, (proj2 (projT2 (precondition_satisfied_inv da config g))). Time destruct (Rle_dec (threshold e) p) as [Hle | Hlt]. ++ destruct Hnew_local_state as [Hv He]. simpl fst in Hv. simpl snd in He. @@ -805,7 +793,7 @@ simpl activate. destruct_match_eq Hactive. cbn -[equiv equiv_dec]. destruct_match. - (* the robot is at the edge src *) unfold add_edge, state_G2V. - assert (He := threshold_pos e). rewrite Hchoose_inactive. + assert (He := strict_ratio_bounds (threshold e)). rewrite Hchoose_inactive. change (proj_ratio ratio_0) with 0. rewrite Rplus_0_l at 2. destruct (Rle_bool (threshold e) (proj_ratio (choose_inactive da config id))) eqn:Htest; rewrite Rle_bool_true_iff in Htest || rewrite Rle_bool_false_iff in Htest; @@ -817,7 +805,7 @@ simpl activate. destruct_match_eq Hactive. + (* the robot is on an edge *) cbn -[equiv]. destruct_match; cbn -[equiv]. - (* the robot is already past the edge threshold *) - assert (He := threshold_pos e). + assert (He := strict_ratio_bounds (threshold e)). assert (Hp := ratio_bounds (choose_inactive da config id)). transitivity (exist valid_stateV (tgt e, e) (or_intror (reflexivity (tgt e)))); [| now repeat destruct_match]. @@ -834,7 +822,7 @@ simpl activate. destruct_match_eq Hactive. destruct (Rle_bool (threshold e) (p + choose_inactive da config id)) eqn:Hcase; rewrite Rle_bool_true_iff in Hcase || rewrite Rle_bool_false_iff in Hcase; repeat destruct_match; try solve [split; reflexivity | simpl in *; contradiction]; []. - contradiction Hcase. transitivity 1; trivial; []. apply Rlt_le. apply threshold_pos. + contradiction Hcase. transitivity 1; trivial; []. apply Rlt_le. apply strict_ratio_bounds. Qed. End GraphEquivalence. diff --git a/Models/RingSSync.v b/Models/RingSSync.v index b262f29d963869e6a0479c428d1f5c6c67b006cf..d3752a1540185656436c7de65550987b6c03be6c 100644 --- a/Models/RingSSync.v +++ b/Models/RingSSync.v @@ -1,12 +1,12 @@ Require Import Utf8. -From Pactole Require Export Setting Util.Fin Util.Bijection Spaces.Graph - Spaces.Ring Observations.MultisetObservation. +From Pactole Require Export Setting Spaces.Graph Util.Bijection + Spaces.Isomorphism Util.Fin Spaces.Ring Observations.MultisetObservation. Set Implicit Arguments. Section RingSSync. -Context {gt2 : greater_than 2} {Robots : Names}. +Context {n : nat} {ltc_2_n : 2 <c n} {Robots : Names}. Global Existing Instance NodesLoc. @@ -22,8 +22,8 @@ Global Instance RC : robot_choice direction := and we can choose the orientation of the ring. *) Global Instance FC : frame_choice (location * bool) := { frame_choice_bijection := - λ nb, if snd nb then comp (asbf (fst nb)â»Â¹) (symbf (fst nb)) - else asbf (fst nb)â»Â¹; + λ nb, if snd nb then (trans (fst nb)) ∘ (sym (fst nb)) + else trans (fst nb); frame_choice_Setoid := eq_setoid _ }. Global Existing Instance NoChoice. diff --git a/Spaces/Graph.v b/Spaces/Graph.v index a3227694ae56eec46100c75d202768e4175b777a..43027f22889437700864a72fb4fb72becfd92e4b 100644 --- a/Spaces/Graph.v +++ b/Spaces/Graph.v @@ -8,11 +8,10 @@ *) (**************************************************************************) -Require Import Rbase SetoidDec. -From Pactole Require Import Util.Coqlib Util.Fin Core.State. +Require Import Utf8 Rbase SetoidDec. +From Pactole Require Import Util.Coqlib Util.Fin Util.Ratio Core.State. Set Implicit Arguments. - Class Graph (V E : Type) := { V_Setoid :> Setoid V; E_Setoid :> Setoid E; @@ -21,31 +20,63 @@ Class Graph (V E : Type) := { src : E -> V; (* source and target of an edge *) tgt : E -> V; - threshold : E -> R; (* TODO: use [strict_ratio] instead? *) - threshold_pos : forall e, (0 < threshold e < 1)%R; src_compat :> Proper (equiv ==> equiv) src; tgt_compat :> Proper (equiv ==> equiv) tgt; - threshold_compat :> Proper (equiv ==> Logic.eq) threshold; find_edge : V -> V -> option E; find_edge_compat :> Proper (equiv ==> equiv ==> opt_eq equiv) find_edge; - find_edge_None : forall a b : V, find_edge a b == None <-> forall e : E, ~(src e == a /\ tgt e == b); - find_edge_Some : forall v1 v2 e, find_edge v1 v2 == Some e <-> v1 == src e /\ v2 == tgt e }. + find_edge_Some : ∀ e v1 v2, find_edge v1 v2 == Some e <-> v1 == src e /\ v2 == tgt e}. -Global Opaque threshold_pos src_compat tgt_compat threshold_compat find_edge_compat find_edge_None find_edge_Some. +Class ThresholdGraph (V E : Type) := { + nothreshold_graph : Graph V E; + threshold : E -> strict_ratio; + threshold_compat :> Proper (equiv ==> equiv) threshold}. +Coercion nothreshold_graph : ThresholdGraph >-> Graph. +Existing Instance nothreshold_graph. +Global Opaque src_compat tgt_compat threshold_compat find_edge_compat find_edge_Some. -(** A finite graph ia a graph where the set [V] of vertices is a prefix of N. *) -(* FIXME: nothing prevents E from being infinite! *) -(* TODO: Maybe we should reuse the type used for robot names *) -Definition finite_node n := fin n. +Section some_lemmas. + +Context {V E : Type} {G : Graph V E}. + +Lemma find_edge_None : ∀ v1 v2 : V, + @find_edge _ _ G v1 v2 == None <-> ∀ e : E, ¬ (v1 == src e /\ v2 == tgt e). +Proof using . + intros. rewrite <- (Decidable.not_not_iff _ (option_decidable _)). + setoid_rewrite not_None_iff. split. + - intros H1 e H2. apply H1. exists e. apply find_edge_Some, H2. + - intros H1 [e H2]. apply (H1 e), find_edge_Some, H2. +Qed. -(* We explictely define the setoid here to avoid using proj1_Setoid instead. *) -Instance finite_node_Setoid n : Setoid (finite_node n) := fin_Setoid n. -Instance finite_node_EqDec n : EqDec (finite_node_Setoid n) := @fin_dec n. +Lemma find_edge_pick : ∀ v1 v2 : V, + pick_spec (λ e : E, v1 == src e /\ v2 == tgt e) (find_edge v1 v2). +Proof using . + intros. apply pick_Some_None. intros e. + apply find_edge_Some. apply find_edge_None. +Qed. -Definition FiniteGraph (n : nat) E := Graph (finite_node n) E. +(* The specifications of find_edge make the graph simple *) +Lemma simple_graph : ∀ e1 e2 : E, + @src _ _ G e1 == src e2 /\ @tgt _ _ G e1 == tgt e2 -> e1 == e2. +Proof using . + intros * [Hs Ht]. apply Some_inj. erewrite <-2 (proj2 (find_edge_Some _ _ _)). + apply find_edge_compat. 3,4: split. 1,2,5,6: symmetry. + 1,3,5: apply Hs. all: apply Ht. +Qed. + +Lemma simple_graph_iff : ∀ e1 e2 : E, + @src _ _ G e1 == src e2 /\ @tgt _ _ G e1 == tgt e2 <-> e1 == e2. +Proof using . + intros. split. apply simple_graph. intros <-. split. all: reflexivity. +Qed. + +End some_lemmas. + +(** A finite graph ia a graph where the set [V] of vertices is a prefix of N. *) +(* FIXME: nothing prevents E from being infinite! *) +(* Definition FiniteGraph (n : nat) E := Graph (fin n) E. Existing Class FiniteGraph. -Global Identity Coercion proj_graph : FiniteGraph >-> Graph. +Global Identity Coercion proj_graph : FiniteGraph >-> Graph. *) Definition NodesLoc {V E : Type} {G : Graph V E} : Location := make_Location V. diff --git a/Spaces/Grid.v b/Spaces/Grid.v index 30fa88429525501df99b02f5fb426a7958a6e62c..a67194c0c67b2b30fe55ca95a3618279c1ec6643 100644 --- a/Spaces/Grid.v +++ b/Spaces/Grid.v @@ -71,35 +71,22 @@ Instance Z2 : Graph node edge. simple refine {| V_EqDec := node_EqDec; E_EqDec := edge_EqDec; src := fst; - tgt := edge_tgt; - threshold := fun _ => (1 / 2)%R |}. + tgt := edge_tgt; |}. Proof using . -* (* threshold_pos *) - intros. lra. * (* find_edge *) exact (fun x y : node => if equiv_dec (EqDec := node_EqDec) y x then Some (x, Self) else if y =?= (fst x + 1, snd x) then Some (x, East) else if y =?= (fst x, snd x + 1) then Some (x, North) else if y =?= (fst x - 1, snd x) then Some (x, West) else if y =?= (fst x, snd x - 1) then Some (x, South) else None). -* (* find_edge_None *) - intros x y. cbn -[equiv]. repeat destruct_match. - + abstract (split; try tauto; []; intro He; apply (He (x, Self)); auto). - + abstract (split; try tauto; []; intro He; apply (He (x, East)); auto). - + abstract (split; try tauto; []; intro He; apply (He (x, North)); auto). - + abstract (split; try tauto; []; intro He; apply (He (x, West)); auto). - + abstract (split; try tauto; []; intro He; apply (He (x, South)); auto). - + split; intros _; auto; []. - abstract (intros [x' d] [Hx He]; simpl in Hx; subst x'; - rewrite <- He in *; destruct d; unfold edge_tgt in *; simpl in *; auto). * (* find_edge_Some *) - intros x y e. cbn -[equiv]. repeat destruct_match; simpl; - try abstract (solve [ split; intro Heq; subst; unfold edge_tgt; simpl in *; try tauto; []; - destruct e as [p d], Heq as [? Heq]; simpl in *; f_equal; trivial; []; - subst; unfold edge_tgt in *; - destruct p, d; simpl in *; reflexivity || (injection Heq; lia) ]); []. - split; try tauto; []. intros []. - abstract (subst; destruct e as [? []]; unfold edge_tgt in *; simpl in *; auto). + intros [p d] x y. cbn-[equiv]. all: repeat destruct_match. 1-5: rewrite e. + 1-5: rewrite (injective_compat_iff Some_inj). all: cbn. 6: split; [apply except|]. + 1-5: rewrite pair_equal_spec; split; intros []; subst. 1-10: split; [reflexivity|]. + all: unfold edge_tgt in *. all: cbn in *. 1,3,5,7,9: reflexivity. 6: intros []; subst. + all: destruct d; cbn in *. 1-25: try reflexivity. 8,12,16,20: symmetry in H0. + all: rewrite (surjective_pairing p) in * at 1. all: cbn in *. + all: rewrite pair_equal_spec in *. all: lia. Defined. (** ** Change of frame of reference in Z² **) diff --git a/Spaces/Isomorphism.v b/Spaces/Isomorphism.v index 1421d198500c68d1ddca90debc95d59e2421902f..74827af60f4b2499f39ee7281ce8e6ddb1ba3229 100644 --- a/Spaces/Isomorphism.v +++ b/Spaces/Isomorphism.v @@ -9,13 +9,8 @@ (**************************************************************************) -Require Import Utf8. -Require Import SetoidDec. -Require Import Rbase Rbasic_fun. -Require Import Psatz. -Require Import Pactole.Util.Coqlib. -Require Import Pactole.Util.Bijection. -Require Import Pactole.Core.Configuration. +Require Import Utf8 SetoidDec Rbase Rbasic_fun Psatz. +From Pactole Require Import Util.Coqlib Util.Bijection. Require Import Pactole.Spaces.Graph. Set Implicit Arguments. @@ -32,22 +27,18 @@ Context {G : Graph V E}. Record isomorphism := { iso_V :> bijection V; iso_E : bijection E; - iso_T : bijection R; - iso_morphism : forall e, iso_V (src e) == src (iso_E e) - /\ iso_V (tgt e) == tgt (iso_E e); - iso_threshold : forall e, iso_T (threshold e) = threshold (iso_E e); - iso_incr : forall a b, (a < b)%R -> (iso_T a < iso_T b)%R; - iso_bound_T : forall r, (0 < iso_T r < 1)%R <-> (0 < r < 1)%R }. + iso_morphism : ∀ e, iso_V (src e) == src (iso_E e) + /\ iso_V (tgt e) == tgt (iso_E e) }. Global Instance isomorphism_Setoid : Setoid isomorphism. -simple refine {| equiv := fun iso1 iso2 => iso1.(iso_V) == iso2.(iso_V) - /\ iso1.(iso_E) == iso2.(iso_E) - /\ iso1.(iso_T) == iso2.(iso_T) |}; autoclass. -Proof. split. -+ intro f. now repeat split. -+ intros f g Hfg; destruct Hfg as [HV [HE HT]]. repeat split; intro; now symmetry. -+ intros f g h Hfg Hgh. destruct Hfg as [? [? ?]], Hgh as [? [? ?]]. - repeat split; intro; etransitivity; eauto. +Proof using . + simple refine {| + equiv := λ iso1 iso2, iso1.(iso_V) == iso2.(iso_V) + /\ iso1.(iso_E) == iso2.(iso_E) |}; autoclass. split. + + intro f. now repeat split. + + intros f g Hfg; destruct Hfg as [HV HE]. split; now symmetry. + + intros f g h Hfg Hgh. destruct Hfg as [? ?], Hgh as [? ?]. + split; etransitivity; eauto. Defined. Instance iso_V_compat : Proper (equiv ==> equiv) iso_V. @@ -56,33 +47,26 @@ Proof using . intros ? ? Heq ?. now apply Heq. Qed. Instance iso_E_compat : Proper (equiv ==> equiv) iso_E. Proof using . intros ? ? Heq ?. now apply Heq. Qed. -Instance iso_T_compat : Proper (equiv ==> equiv) iso_T. -Proof using . intros ? ? Heq ?. now apply Heq. Qed. - +Lemma equiv_iso_V_to_iso_E iso1 iso2: + iso_V iso1 == iso_V iso2 -> iso_E iso1 == iso_E iso2. +Proof using . + intros Heqiso_V e. apply simple_graph. rewrite <-! (proj1 (iso_morphism _ _)), + <-! (proj2 (iso_morphism _ _)), <- Heqiso_V. split. all: reflexivity. +Qed. Definition id : isomorphism. -refine {| iso_V := @id V _; - iso_E := @id E _; - iso_T := @id R _ |}. -Proof. -+ now intros. -+ now intros. -+ now intros. -+ now intros. +Proof using . + refine {| iso_V := id; + iso_E := id |}. + now intros. Defined. - Definition comp (f g : isomorphism) : isomorphism. -refine {| - iso_V := compose f.(iso_V) g.(iso_V); - iso_E := compose f.(iso_E) g.(iso_E); - iso_T := compose f.(iso_T) g.(iso_T) |}. -Proof. -+ intro. simpl. - split; now rewrite <- 2 (proj1 (iso_morphism _ _)) || rewrite <- 2 (proj2 (iso_morphism _ _)). -+ intro. simpl. now rewrite 2 iso_threshold. -+ intros. simpl. now do 2 apply iso_incr. -+ intro. simpl. now rewrite 2 iso_bound_T. +Proof using . + refine {| + iso_V := compose f.(iso_V) g.(iso_V); + iso_E := compose f.(iso_E) g.(iso_E) |}. + intro. simpl. split; now rewrite <- 2 (proj1 (iso_morphism _ _)) || rewrite <- 2 (proj2 (iso_morphism _ _)). Defined. Global Instance IsoComposition : Composition isomorphism. @@ -92,54 +76,42 @@ Proof. intros f1 f2 Hf g1 g2 Hg. repeat split; intro; simpl; now rewrite Hf, Hg. (* Global Instance compose_compat : Proper (equiv ==> equiv ==> equiv) compose. Proof. intros f1 f2 Hf g1 g2 Hg. repeat split; intro; simpl; now rewrite Hf, Hg. Qed. *) -Lemma compose_assoc : forall f g h, f ∘ (g ∘ h) == (f ∘ g) ∘ h. +Lemma compvE : ∀ iso1 iso2 : isomorphism, + iso_V (iso1 ∘ iso2) = iso_V iso1 ∘ iso_V iso2 :> bijection V. +Proof using . reflexivity. Qed. + +Lemma compeE : ∀ iso1 iso2 : isomorphism, + iso_E (iso1 ∘ iso2) = iso_E iso1 ∘ iso_E iso2 :> bijection E. +Proof using . reflexivity. Qed. + +Lemma compose_assoc : ∀ f g h, f ∘ (g ∘ h) == (f ∘ g) ∘ h. Proof using . intros f g h; repeat split; simpl; reflexivity. Qed. Definition inv (iso : isomorphism) : isomorphism. +Proof using . refine {| iso_V := inverse iso.(iso_V); - iso_E := inverse iso.(iso_E); - iso_T := inverse iso.(iso_T) - |}. -Proof. -+ intro. simpl. rewrite <- 2 Inversion, (proj1 (iso_morphism _ _)), (proj2 (iso_morphism _ _)). + iso_E := inverse iso.(iso_E) |}. + intro. simpl. rewrite <- 2 Inversion, (proj1 (iso_morphism _ _)), (proj2 (iso_morphism _ _)). split; apply src_compat || apply tgt_compat; now rewrite Inversion. -+ intro. simpl. change eq with (@equiv R _). rewrite <- Inversion, iso_threshold. - apply threshold_compat. now rewrite Inversion. -+ intros a b Hab. - simpl. - assert (Hincr := iso_incr iso). - assert (forall x y, @retraction R _ (iso_T iso) x < @retraction R _ (iso_T iso) y -> x < y)%R. - { intros. - specialize (Hincr (@retraction R _ (iso_T iso) x) (@retraction R _ (iso_T iso) y) H). - now do 2 rewrite section_retraction in Hincr. } - assert (Hnondecr: - (forall x y, x <= y -> @retraction R _ (iso_T iso) x <= @retraction R _ (iso_T iso) y)%R). - { intros x y Hle. apply Rnot_lt_le. apply Rle_not_lt in Hle. intuition. } - destruct (Hnondecr a b) as [| Heq]; auto; intuition; []. - apply (f_equal (iso_T iso)) in Heq. rewrite 2 section_retraction in Heq. lra. -+ intro. simpl. - assert (Hbound := iso_bound_T iso). specialize (Hbound (@retraction R _ (iso_T iso) r)). - now rewrite section_retraction in Hbound. Defined. Global Instance IsoInverse : Inverse isomorphism. -refine {| inverse := inv |}. -Proof. -intros f g [? [? ?]]. -repeat split; intro; simpl; change eq with (@equiv R _); f_equiv; auto. +Proof using . + refine {| inverse := inv |}. + intros f g [? ?]. repeat split; intro; simpl; change eq with (@equiv R _); f_equiv; auto. Defined. Lemma id_inv : idâ»Â¹ == id. -Proof using . split; [|split]. all: setoid_rewrite id_inv. all: reflexivity. Qed. +Proof using . split. all: setoid_rewrite id_inv. all: reflexivity. Qed. Lemma id_comp_l : ∀ iso : isomorphism, id ∘ iso == iso. -Proof using. intros. split; [|split]. all: setoid_rewrite id_comp_l. all: reflexivity. Qed. +Proof using. intros. split. all: setoid_rewrite id_comp_l. all: reflexivity. Qed. Lemma id_comp_r : ∀ iso : isomorphism, iso ∘ id == iso. -Proof using. intros. split; [|split]. all: setoid_rewrite id_comp_r. all: reflexivity. Qed. +Proof using. intros. split. all: setoid_rewrite id_comp_r. all: reflexivity. Qed. Lemma inv_inv : ∀ iso : isomorphism, isoâ»Â¹â»Â¹ == iso. -Proof using . intros. split; [|split]. all: setoid_rewrite inv_inv. all: reflexivity. Qed. +Proof using . intros. split. all: setoid_rewrite inv_inv. all: reflexivity. Qed. (* Global Instance inverse_compat : Proper (equiv ==> equiv) inverse. Proof. @@ -147,27 +119,20 @@ intros f g [? [? ?]]. repeat split; intro; simpl; change eq with (@equiv R _); f_equiv; auto. Qed. *) -Lemma compose_inverse_l : forall iso : isomorphism, iso â»Â¹ ∘ iso == id. +Lemma compose_inverse_l : ∀ iso : isomorphism, iso â»Â¹ ∘ iso == id. Proof using . intro. repeat split; intro; simpl; try now rewrite retraction_section; autoclass. Qed. -Lemma compose_inverse_r : forall iso : isomorphism, iso ∘ (iso â»Â¹) == id. +Lemma compose_inverse_r : ∀ iso : isomorphism, iso ∘ (iso â»Â¹) == id. Proof using . intro. repeat split; intro; simpl; try now rewrite section_retraction; autoclass. Qed. -Lemma inverse_compose : forall f g : isomorphism, (f ∘ g) â»Â¹ == (g â»Â¹) ∘ (f â»Â¹). +Lemma inverse_compose : ∀ f g : isomorphism, (f ∘ g) â»Â¹ == (g â»Â¹) ∘ (f â»Â¹). Proof using . intros f g; repeat split; intro; simpl; reflexivity. Qed. -Lemma injective : forall iso : isomorphism, Preliminary.injective equiv equiv iso. -Proof using . -intros f x y Heq. transitivity (id x); try reflexivity; []. -rewrite <- (compose_inverse_l f). simpl. rewrite Heq. -now apply compose_inverse_l. -Qed. - End Isomorphism. Arguments isomorphism {V} {E} G. -Lemma find_edge_iso_Some `{G : Graph} : forall (iso : isomorphism G) src tgt e, +Lemma find_edge_iso_Some `{G : Graph} : ∀ (iso : isomorphism G) src tgt e, @find_edge _ _ G (iso src) (iso tgt) == Some (iso.(iso_E) e) <-> @find_edge _ _ G src tgt == Some e. Proof using . @@ -184,7 +149,7 @@ revert iso src tgt e. apply strong_and. simpl in Hstep. now rewrite 3 Bijection.retraction_section in Hstep. Qed. -Lemma find_edge_iso_None `{G : Graph} : forall (iso : isomorphism G) src tgt, +Lemma find_edge_iso_None `{G : Graph} : ∀ (iso : isomorphism G) src tgt, @find_edge _ _ G (iso src) (iso tgt) == None <-> @find_edge _ _ G src tgt == None. Proof using . intros iso src tgt. destruct (find_edge src tgt) eqn:Hcase. diff --git a/Spaces/Ring.v b/Spaces/Ring.v index 228e5ad4de3c0f55d82e39fdf26fd561ff57fe50..894c1163b577aec279b2fdce14cf59149912f422 100644 --- a/Spaces/Ring.v +++ b/Spaces/Ring.v @@ -14,36 +14,49 @@ From Pactole Require Import Util.Coqlib Util.Bijection Spaces.Graph (** ** A ring **) -Section Ring. - -Context {_n : greater_than 2}. -Notation n := (@ub 2 _n). - -Definition ring_node := finite_node ub. Inductive direction := Forward | Backward | SelfLoop. -Definition ring_edge := (ring_node * direction)%type. Global Instance direction_Setoid : Setoid direction := eq_setoid direction. -Global Instance ring_edge_Setoid : Setoid ring_edge := prod_Setoid _ _. - Lemma direction_eq_dec_subproof : ∀ d1 d2 : direction, {d1 = d2} + {d1 <> d2}. Proof using . decide equality. Defined. Global Instance direction_EqDec : EqDec direction_Setoid := direction_eq_dec_subproof. -Global Instance ring_edge_EqDec : EqDec ring_edge_Setoid := prod_EqDec _ _. - -(* Returns the element of nat to give as parameter - to mod2fin to translate in the direction dir *) -Definition dir2nat (d : direction) : nat := +(* Returns the nat to give as parameter to + addm to translate in the direction dir *) +Definition dir2nat {n : nat} {ltc_2_n : 2 <c n} (d : direction) : nat := match d with | SelfLoop => 0 | Forward => 1 - | Backward => Nat.pred ub + | Backward => Nat.pred n end. +(* ''inverse'' of dir2nat *) +Definition nat2Odir {n : nat} {ltc_2_n : 2 <c n} (m : nat) : option direction := + match m with + | 0 => Some SelfLoop + | 1 => Some Forward + | _ => if m =? Nat.pred n then Some Backward else None + end. + +Global Instance ltc_1_n {n : nat} {ltc_2_n : 2 <c n} : 1 <c n + := (lt_s_u 1 ltac:(auto)). + +Definition ring_edge {n : nat} {ltc_2_n : 2 <c n} := (fin n * direction)%type. + +Section Ring. + +Context {n : nat} {ltc_2_n : 2 <c n}. + +Global Instance ring_edge_Setoid : Setoid ring_edge := eq_setoid ring_edge. + +Lemma ring_edge_dec : ∀ e1 e2 : ring_edge, {e1 = e2} + {e1 ≠e2}. +Proof using . apply pair_dec. apply fin_EqDec. apply direction_EqDec. Qed. + +Global Instance ring_edge_EqDec : EqDec ring_edge_Setoid := ring_edge_dec. + Definition dir2nat_compat : Proper (equiv ==> equiv) dir2nat := _. Lemma dir20 : dir2nat SelfLoop = 0. @@ -52,36 +65,22 @@ Proof using . reflexivity. Qed. Lemma dir21 : dir2nat Forward = 1. Proof using . reflexivity. Qed. -Lemma dir2SSn : dir2nat Backward = Nat.pred ub. +Lemma dir2pred_n : dir2nat Backward = Nat.pred n. Proof using . reflexivity. Qed. -Lemma dir2nat_lt : ∀ d : direction, dir2nat d < ub. +Lemma dir2nat_lt : ∀ d : direction, dir2nat d < n. Proof using . - intros. destruct d. all: cbn. 2: apply lt_pred_ub. - all: apply lt_sb_ub. all: auto. + intros. destruct d. all: cbn. 2: apply lt_pred_u. + all: apply lt_s_u. all: auto. Qed. -Lemma dir2nat_inj : Util.Preliminary.injective equiv equiv dir2nat. +Lemma dir2natI : Util.Preliminary.injective equiv equiv dir2nat. Proof using . intros d1 d2 H'. destruct d1, d2. all: inversion H' as [H]. all: try reflexivity. all: exfalso. 1,4: symmetry in H. - all: eapply neq_pred_ub_sb. 2,4,6,8: exact H. all: auto. -Qed. - -Lemma dir2mod2fin : ∀ d : direction, - fin2nat (mod2fin (dir2nat d)) = dir2nat d. -Proof using . - intros. rewrite mod2fin2nat. apply Nat.mod_small. apply dir2nat_lt. + all: eapply neq_pred_u_s. 2,4,6,8: exact H. all: auto. Qed. -(* ''inverse'' of dir2nat *) -Definition nat2Odir (m : nat) : option direction := - match m with - | 0 => Some SelfLoop - | 1 => Some Forward - | _ => if m =? Nat.pred n then Some Backward else None - end. - Definition nat2Odir_compat : Proper (equiv ==> equiv) nat2Odir := _. Lemma nat2Odir_Some : ∀ (d : direction) (m : nat), @@ -89,10 +88,10 @@ Lemma nat2Odir_Some : ∀ (d : direction) (m : nat), Proof using . unfold dir2nat, nat2Odir. intros d m. split; intros H. all: destruct d, m as [|p]. all: try inversion H. all: try reflexivity. - 4:{ exfalso. eapply neq_pred_ub_sb. 2: exact H. auto. } + 4:{ exfalso. eapply neq_pred_u_s. 2: exact H. auto. } 4: rewrite Nat.eqb_refl. all: destruct p. all: try inversion H. - all: try reflexivity. 4:{ exfalso. eapply neq_pred_ub_sb. 2: exact H. - auto. } all: destruct (S (S p) =? Nat.pred ub) eqn:Hd. + all: try reflexivity. 4:{ exfalso. eapply neq_pred_u_s. 2: exact H. + auto. } all: destruct (S (S p) =? Nat.pred n) eqn:Hd. all: try inversion H. symmetry. apply Nat.eqb_eq, Hd. Qed. @@ -109,12 +108,9 @@ Qed. Lemma nat2Odir_pick : ∀ m : nat, pick_spec (λ d : direction, dir2nat d = m) (nat2Odir m). Proof using . - intros. unfold nat2Odir. destruct m as [|m]. apply Pick, dir20. - destruct m as [|m]. apply Pick, dir21. - destruct (S (S m) =? Nat.pred ub) eqn:Hd. - - apply Nat.eqb_eq in Hd. rewrite Hd. apply Pick, dir2SSn. - - apply Nat.eqb_neq in Hd. apply Nopick. intros d H. apply Hd. - destruct d. all: inversion H. subst. reflexivity. + intros. apply pick_Some_None. intros d. + all: rewrite opt_Setoid_eq by reflexivity. + apply nat2Odir_Some. apply nat2Odir_None. Qed. Definition nat2dir (m : nat) (H : ∃ d, dir2nat d = m) : direction. @@ -127,7 +123,7 @@ Lemma dir2natK : ∀ (d1 : direction) (H : ∃ d2, dir2nat d2 = dir2nat d1), nat2dir (dir2nat d1) H = d1. Proof using . intros. unfold nat2dir. destruct (nat2Odir_pick (dir2nat d1)) as [m Hd | Hd]. - eapply dir2nat_inj, Hd. contradiction (Hd d1). reflexivity. + eapply dir2natI, Hd. contradiction (Hd d1). reflexivity. Qed. Lemma nat2dirK : @@ -137,68 +133,137 @@ Proof using . exact Hd. exfalso. destruct H as [d H]. apply (Hd d). exact H. Qed. +(* Returns the fin n to give as parameter to + addf to translate in the direction dir *) +Definition dir2nod (d : direction) : fin n := Fin (dir2nat_lt d). + +(* ''inverse'' of dir2nat *) +Definition nod2Odir (v : fin n) : option direction := nat2Odir v. + +Definition dir2nod_compat : Proper (equiv ==> equiv) dir2nod := _. + +Definition nod2Odir_compat : Proper (equiv ==> equiv) nod2Odir := _. + +Lemma dir2nod2nat : ∀ d : direction, dir2nod d = dir2nat d :> nat. +Proof using . reflexivity. Qed. + +Lemma dir2nodE : ∀ d : direction, dir2nod d = mod2fin (dir2nat d). +Proof using . + intros. apply fin2natI. rewrite dir2nod2nat. + symmetry. apply mod2fin_small, dir2nat_lt. +Qed. + +Lemma nod2OdirE : ∀ v : fin n, nod2Odir v = nat2Odir v. +Proof using . reflexivity. Qed. + +Lemma dir2fin0 : dir2nod SelfLoop = fin0. +Proof using . apply fin2natI. rewrite dir2nod2nat. apply dir20. Qed. + +Lemma dir2fin1 : dir2nod Forward = fin1. +Proof using . + apply fin2natI. symmetry. rewrite dir2nod2nat, dir21. apply fin12nat. +Qed. + +Lemma dir2max : dir2nod Backward = fin_max. +Proof using . apply fin2natI. rewrite dir2nod2nat. apply dir2pred_n. Qed. + +Lemma dir2nodI : Util.Preliminary.injective equiv equiv dir2nod. +Proof using . + intros d1 d2 H. apply dir2natI. rewrite <-2 dir2nod2nat, H. reflexivity. +Qed. + +Lemma nod2Odir_Some : ∀ (d : direction) (v : fin n), + nod2Odir v = Some d <-> dir2nod d = v. +Proof using . + intros. rewrite nod2OdirE. split. all: intros H. apply fin2natI, + nat2Odir_Some, H. apply nat2Odir_Some. rewrite <- dir2nod2nat, H. reflexivity. +Qed. + +Lemma nod2Odir_None : ∀ v : fin n, + nod2Odir v = None <-> (∀ d : direction, dir2nod d ≠v). +Proof using . + intros. rewrite nod2OdirE. split. + - intros H d Habs. eapply (proj1 (nat2Odir_None _)). apply H. + rewrite <- Habs. symmetry. apply dir2nod2nat. + - intros H. apply nat2Odir_None. intros d Habs. apply (H d), fin2natI, Habs. +Qed. + +Lemma nod2Odir_pick : ∀ v : fin n, + pick_spec (λ d : direction, dir2nod d = v) (nod2Odir v). +Proof using . + intros. apply pick_Some_None. intros d. + all: rewrite opt_Setoid_eq by reflexivity. + apply nod2Odir_Some. apply nod2Odir_None. +Qed. + +Definition nod2dir (v : fin n) (H : ∃ d, dir2nod d = v) : direction. +Proof using . + apply (nat2dir v). abstract (destruct H as [d H]; exists d; + rewrite <- H; symmetry; apply dir2nod2nat). +Defined. + +Lemma dir2nodK : ∀ (d1 : direction) (H : ∃ d2, dir2nod d2 = dir2nod d1), + nod2dir (dir2nod d1) H = d1. +Proof using . intros. erewrite <- (dir2natK d1). reflexivity. Qed. + +Lemma nod2dirK : + ∀ (v : fin n) (H : ∃ d, dir2nod d = v), dir2nod (nod2dir v H) = v. +Proof using . + intros. apply fin2natI. erewrite <- (nat2dirK v). reflexivity. +Qed. + (** From a node, if we move in one direction, get get to another node. *) -Definition move_along (v : ring_node) (d : direction) := addm v (dir2nat d). +Definition move_along (v : fin n) (d : direction) + := addf v (dir2nod d). Definition move_along_compat : Proper (equiv ==> equiv ==> equiv) move_along := _. -Lemma move_alongE : ∀ (v : ring_node) (d : direction), - move_along v d = addm v (dir2nat d). +Lemma move_alongE : ∀ (v : fin n) (d : direction), + move_along v d = addf v (dir2nod d). Proof using . intros. reflexivity. Qed. -Definition move_along2nat : ∀ (v : ring_node) (d : direction), - fin2nat (move_along v d) = (v + dir2nat d) mod ub. -Proof using . intros. apply addm2nat. Qed. +Definition move_along2nat : ∀ (v : fin n) (d : direction), + fin2nat (move_along v d) = (v + dir2nat d) mod n. +Proof using . intros. apply addf2nat. Qed. Lemma move_alongI_ : ∀ d : direction, Util.Preliminary.injective equiv equiv (λ v, move_along v d). -Proof using . intros. apply addIm. Qed. +Proof using . intros. apply addIf. Qed. -Lemma move_along_I : ∀ v : ring_node, +Lemma move_along_I : ∀ v : fin n, Util.Preliminary.injective equiv equiv (move_along v). -Proof using . - intros v d1 d2 H. eapply dir2nat_inj, addm_bounded_diff_inj. 3: exact H. - all: split. 1,3: apply Nat.le_0_l. all: apply dir2nat_lt. -Qed. +Proof using . intros v d1 d2 H. eapply dir2nodI, addfI, H. Qed. -Lemma move_along_0 : ∀ v : ring_node, move_along v SelfLoop = v. -Proof using . apply addm0. Qed. +Lemma move_along_0 : ∀ v : fin n, move_along v SelfLoop = v. +Proof using . apply addf0. Qed. -Lemma move_along0_ : ∀ d : direction, move_along fin0 d = mod2fin (dir2nat d). -Proof using . - intros. apply fin2nat_inj. rewrite move_along2nat, mod2fin2nat, fin02nat, - Nat.add_0_l. reflexivity. -Qed. +Lemma move_along0_ : ∀ d : direction, move_along fin0 d = dir2nod d. +Proof using . intros. rewrite move_alongE. apply add0f. Qed. -Lemma move_alongAC : ∀ (v : ring_node) (d1 d2 : direction), +Lemma move_alongAC : ∀ (v : fin n) (d1 d2 : direction), move_along (move_along v d2) d1 = move_along (move_along v d1) d2. -Proof using . intros. apply addmAC. Qed. +Proof using . intros. rewrite 4 move_alongE. apply addfAC. Qed. -Lemma addm_move_along : ∀ (v : ring_node) (m : nat) (d : direction), +Lemma addm_move_along : ∀ (v : fin n) (m : nat) (d : direction), addm (move_along v d) m = move_along (addm v m) d. -Proof using . intros. apply addmAC. Qed. +Proof using . intros. rewrite 2 move_alongE, 2 addm_addf. apply addfAC. Qed. -Lemma subm_move_along : ∀ (v : ring_node) (m : nat) (d : direction), +Lemma subm_move_along : ∀ (v : fin n) (m : nat) (d : direction), subm (move_along v d) m = move_along (subm v m) d. -Proof using . intros. apply addmAC. Qed. +Proof using . intros. rewrite 2 move_alongE, 2 submE. apply addfAC. Qed. -Lemma addf_move_along : ∀ (v1 v2 : ring_node) (d : direction), +Lemma addf_move_along : ∀ (v1 v2 : fin n) (d : direction), addf (move_along v1 d) v2 = move_along (addf v1 v2) d. -Proof using . intros. apply addmAC. Qed. +Proof using . intros. rewrite 2 move_alongE. apply addfAC. Qed. -Lemma subf_move_along : ∀ (v1 v2 : ring_node) (d : direction), - subf v2 (move_along v1 d) = subm (subf v2 v1) (dir2nat d). -Proof using . - intros. rewrite subm_subf, <- subf_addf, addf_addm, dir2mod2fin. reflexivity. -Qed. +Lemma subf_move_along : ∀ (v1 v2 : fin n) (d : direction), + subf v2 (move_along v1 d) = subf (subf v2 v1) (dir2nod d). +Proof using . intros. rewrite move_alongE. apply subf_addf. Qed. -Lemma subf_move_along' : ∀ (v1 v2 : ring_node) (d : direction), +Lemma subf_move_along' : ∀ (v1 v2 : fin n) (d : direction), subf (move_along v1 d) v2 = move_along (subf v1 v2) d. -Proof using . - intros. rewrite <- oppf_subf, subf_move_along, oppf_subm, subfA, addfC, - <- addf_subf, addf_addm, dir2mod2fin. reflexivity. -Qed. +Proof using . intros. rewrite 2 move_alongE, 2 subfE. apply addfAC. Qed. (* returns the ''opposite'' direction *) Definition swap_direction (d : direction) := @@ -224,309 +289,290 @@ Proof using . intros d. unfold swap_direction. destruct d. all: reflexivity. Qed. -Definition swap_direction_bij : bijection direction +Definition swbd : bijection direction := cancel_bijection swap_direction _ swap_directionI swap_directionK swap_directionK. -Lemma dir2nat_swap_direction : ∀ d : direction, - dir2nat (swap_direction d) = oppm (dir2nat d). +Lemma swbdE : swbd = swap_direction :> (direction -> direction). +Proof using . reflexivity. Qed. + +Lemma swbdVE : swbdâ»Â¹ = swap_direction :> (direction -> direction). +Proof using . reflexivity. Qed. + +Lemma swbdV : swbdâ»Â¹ == swbd. +Proof using . intros dir. rewrite swbdE, swbdVE. reflexivity. Qed. + +Lemma swbdK : swbd ∘ swbd == Bijection.id. +Proof using . intros dir. rewrite compE, swbdE. apply swap_directionK. Qed. + +Lemma dir2nod_swap_direction : ∀ d : direction, + dir2nod (swap_direction d) = oppf (dir2nod d). Proof using . - intros. rewrite oppm_oppf, oppf2nat, dir2mod2fin. symmetry. destruct d. - 2: cbn-[Nat.modulo Nat.sub]. 1,3: cbn-[Nat.modulo]. rewrite Nat.sub_1_r. - apply Nat.mod_small, lt_pred_ub. rewrite Nat.sub_0_r. apply Nat.mod_same, - neq_ub_0. rewrite <- S_pred_ub at 1. rewrite Nat.sub_succ_l, - Nat.sub_diag. apply Nat.mod_small. apply lt_sb_ub. auto. reflexivity. + intros. destruct d. all: cbn. all: symmetry. 1,2: rewrite dir2max, dir2fin1. + apply oppf1. apply oppf_max. rewrite dir2fin0. apply oppf0. Qed. -Lemma move_along_swap_direction : ∀ (v : ring_node) (d : direction), - move_along v (swap_direction d) = subm v (dir2nat d). +Lemma move_along_swap_direction : ∀ (v : fin n) (d : direction), + move_along v (swap_direction d) = subf v (dir2nod d). Proof using . - intros. rewrite move_alongE, dir2nat_swap_direction. symmetry. apply submE. + intros. rewrite move_alongE, dir2nod_swap_direction. symmetry. apply subfE. Qed. -Lemma move_alongK' : ∀ (v : ring_node) (d : direction), - subm (move_along v d) (dir2nat d) = v. -Proof using . intros. apply addmVKV. Qed. +Lemma move_alongK' : ∀ (v : fin n) (d : direction), + subf (move_along v d) (dir2nod d) = v. +Proof using . intros. rewrite move_alongE. apply addfVKV. Qed. -Lemma submVKV' : ∀ (v : ring_node) (d : direction), - move_along (subm v (dir2nat d)) d = v. -Proof using . intros. apply submVKV. Qed. +Lemma submVKV' : ∀ (v : fin n) (d : direction), + move_along (subf v (dir2nod d)) d = v. +Proof using . intros. rewrite move_alongE. apply subfVKV. Qed. -Lemma move_alongK : ∀ (v : ring_node) (d : direction), +Lemma move_alongK : ∀ (v : fin n) (d : direction), move_along (move_along v d) (swap_direction d) = v. Proof using . intros. rewrite move_along_swap_direction. apply move_alongK'. Qed. -Lemma move_along_dir2nat_diff : ∀ (v1 v2 : ring_node) (d : direction), - dir2nat d = diff v2 v1 <-> move_along v1 d = v2. +Lemma move_along_dir2nod_subf : ∀ (v1 v2 : fin n) (d : direction), + dir2nod d = subf v2 v1 <-> move_along v1 d = v2. Proof using . - intros. split; intros H. - - rewrite move_alongE, H, diffE, addm_addf, mod2finK. apply subfVK. - - subst. rewrite diffE, subf_move_along', subff, move_along0_. - symmetry. apply dir2mod2fin. + intros. rewrite move_alongE. split; intros H. rewrite H. apply subfVK. + rewrite <- H. symmetry. apply addfKV. Qed. -Lemma symf_move_along : ∀ (c v : ring_node) (d : direction), +Lemma symf_move_along : ∀ (c v : fin n) (d : direction), symf c (move_along v d) = move_along (symf c v) (swap_direction d). Proof using . - intros. rewrite move_along_swap_direction, move_alongE, addm_addf. - apply symf_addf. + intros. rewrite move_along_swap_direction, move_alongE. apply symf_addf. Qed. -Lemma nat2Odir_diff_Some : - ∀ (v1 v2 : ring_node) (d : direction), nat2Odir (diff v1 v2) = Some d - <-> nat2Odir (diff v2 v1) = Some (swap_direction d). +Lemma move_along_symf : ∀ (v1 v2 : fin n) (d : direction), + move_along (symf v1 v2) d = symf v1 (move_along v2 (swap_direction d)). Proof using . - intros. split; intros H. all: apply nat2Odir_Some in H. - all: apply nat2Odir_Some. all: rewrite diffE, <- oppm_diff, <- H, - dir2nat_swap_direction. 2: rewrite oppmK, dir2mod2fin. all: reflexivity. + intros. rewrite symf_move_along, swap_directionK. reflexivity. Qed. -Lemma nat2Odir_diff_None : ∀ v1 v2 : ring_node, - nat2Odir (diff v2 v1) = None <-> ∀ d : direction, move_along v1 d ≠v2. +Lemma nod2Odir_subf_Some : + ∀ (v1 v2 : fin n) (d : direction), nod2Odir (subf v1 v2) = Some d + <-> nod2Odir (subf v2 v1) = Some (swap_direction d). Proof using . - intros. split. - - intros H d Heq. apply move_along_dir2nat_diff in Heq. - rewrite <- Heq, (proj2 (nat2Odir_Some _ _)) in H by reflexivity. - inversion H. - - intros H. apply nat2Odir_None. intros d Heq. - apply (H d), move_along_dir2nat_diff, Heq. + intros. rewrite 2 nod2Odir_Some, dir2nod_swap_direction, <- (oppf_subf v2). + symmetry. apply injective_eq_iff, oppfI. Qed. -Definition find_edge_subterm (v1 v2 : ring_node) : option ring_edge - := match (nat2Odir (diff v2 v1)) with +Lemma nod2Odir_subf_None : ∀ v1 v2 : fin n, + nod2Odir (subf v2 v1) = None <-> ∀ d : direction, move_along v1 d ≠v2. +Proof using . + intros. rewrite nod2Odir_None. split. all: intros H d Heq. all: apply (H d). + all: apply move_along_dir2nod_subf, Heq. +Qed. + +Definition find_edge_subterm (v1 v2 : fin n) : option ring_edge + := match (nod2Odir (subf v2 v1)) with | Some d1 => Some (v1, d1) | None => None end. -Lemma find_edge_None_subproof : - ∀ v1 v2 : ring_node, find_edge_subterm v1 v2 == None - <-> ∀ e : ring_edge, ~(fst e == v1 /\ move_along (fst e) (snd e) == v2). -Proof using . - intros. unfold find_edge_subterm. destruct (nat2Odir (diff v2 v1)) eqn:Hd. - - split; intros H. { inversion H. } apply (H (v1, d)). clear H. split. - { reflexivity. } apply move_along_dir2nat_diff, nat2Odir_Some, Hd. - - split; intros H. 2: reflexivity. clear H. intros e [H1 H2]. cbn in H1, H2. - subst v1. unshelve eapply (proj1 (nat2Odir_diff_None (fst e) v2) _ (snd e)). - 2: exact H2. rewrite <- Hd. reflexivity. -Qed. - -Lemma find_edge_Some_subproof : - ∀ (v1 v2 : ring_node) (e : ring_edge), find_edge_subterm v1 v2 == Some e +Lemma find_edge_subproof : + ∀ (e : ring_edge) (v1 v2 : fin n), find_edge_subterm v1 v2 == Some e <-> v1 == fst e /\ v2 == move_along (fst e) (snd e). Proof using . - intros. unfold find_edge_subterm. destruct (nat2Odir (diff v2 v1)) eqn:Hd. - - split. all: cbn. all: intros [H1 H2]. all: subst. - all: split; [reflexivity|]. all: apply nat2Odir_Some in Hd. - all: apply move_along_dir2nat_diff in Hd. symmetry; exact Hd. - apply move_along_I in Hd. exact Hd. - - split. { intros H. inversion H. } cbn. intros [H1 H2]. subst v1. - unshelve eapply (proj1 (nat2Odir_diff_None (fst e) v2) _ (snd e)). - 2: symmetry; exact H2. rewrite Hd. reflexivity. + intros. rewrite opt_Setoid_eq by reflexivity. unfold find_edge_subterm. + cbn. rewrite (eq_sym_iff v2), <- move_along_dir2nod_subf. + destruct (nod2Odir_pick (subf v2 v1)) as [d Hd|Hd]. + - rewrite (injective_eq_iff Some_eq_inj), <- (pair_eqE (v1, d)). cbn[fst snd]. + split. all: intros [H1 H2]. all: subst. all: split; [reflexivity |]. + apply Hd. apply dir2nodI. rewrite H2. apply Hd. + - split. intros H. inversion H. intros [? H]. subst. exfalso. eapply Hd, H. Qed. -Definition Ring (thd : ring_edge -> R) - (thd_pos : ∀ e, (0 < thd e < 1)%R) - (thd_compat : Proper (equiv ==> eq) thd) - : FiniteGraph n ring_edge. -Proof using . - refine {| - V_EqDec := @finite_node_EqDec ub; - E_EqDec := ring_edge_EqDec; - src := fst; - tgt := λ e, move_along (fst e) (snd e); - threshold := thd; - threshold_pos := thd_pos; - threshold_compat := thd_compat; - find_edge := find_edge_subterm; - find_edge_None := find_edge_None_subproof; - find_edge_Some := find_edge_Some_subproof |}. - abstract (intros ?? H; rewrite H; reflexivity). -Defined. - -(** If we do not care about threshold values, we just take 1/2 everywhere. *) -Global Instance nothresholdRing : FiniteGraph n ring_edge := - Ring (fun _ => 1/2)%R - ltac:(abstract (intro; lra)) - (fun _ _ _ => eq_refl). - -Global Existing Instance proj_graph. +Global Instance Ring : Graph (fin n) ring_edge := {| + V_EqDec := @fin_EqDec n; + E_EqDec := ring_edge_EqDec; + src := fst; + tgt := λ e, move_along (fst e) (snd e); + find_edge := find_edge_subterm; + find_edge_Some := find_edge_subproof |}. Global Instance Ring_isomorphism_Setoid : - Setoid (isomorphism nothresholdRing) := isomorphism_Setoid. + Setoid (isomorphism Ring) := isomorphism_Setoid. -Global Instance Ring_isomorphism_Inverse : Inverse (isomorphism nothresholdRing) - := @IsoInverse _ _ nothresholdRing. +Global Instance Ring_isomorphism_Inverse : Inverse (isomorphism Ring) + := @IsoInverse _ _ Ring. Global Instance Ring_isomorphism_Composition - : Composition (isomorphism nothresholdRing) - := @IsoComposition _ _ nothresholdRing. + : Composition (isomorphism Ring) := @IsoComposition _ _ Ring. (** ** Ring operations **) -Lemma find_edge_move_along : ∀ (v : ring_node) (d : direction), - find_edge v (move_along v d) == Some (v, d). +Lemma srcE : ∀ e : ring_edge, src e = fst e. +Proof using . reflexivity. Qed. + +Lemma tgtE : ∀ e : ring_edge, tgt e = move_along (fst e) (snd e). +Proof using . reflexivity. Qed. + +Lemma find_edgeE : ∀ v1 v2 : fin n, + find_edge v1 v2 = match (nod2Odir (subf v2 v1)) with + | Some d1 => Some (v1, d1) + | None => None + end. +Proof using . reflexivity. Qed. + +Lemma find_edge_move_along : ∀ (v : fin n) (d : direction), + find_edge v (move_along v d) = Some (v, d). Proof using . - intros. setoid_rewrite find_edge_Some. split. all: reflexivity. + intros. rewrite <- opt_Setoid_eq. setoid_rewrite find_edge_Some. + split. all: reflexivity. Qed. -Definition trans (m : nat) : isomorphism nothresholdRing. +Definition trans (v : fin n) : isomorphism Ring. Proof using . - unshelve refine {| - iso_V := (asbm m)â»Â¹; - iso_E := prod_bij ((asbm m)â»Â¹) (@Bijection.id _ direction_Setoid); - iso_T := @Bijection.id _ R_Setoid |}. - - (* iso_morphism *) - abstract (intros; split; [reflexivity | apply addm_move_along]). - - (* iso_threshold *) abstract (reflexivity). - - (* iso_incr *) abstract (intros * H; apply H). - - (* iso_bound_T *) abstract (reflexivity). + refine {| + iso_V := asbf vâ»Â¹; + iso_E := prod_eq_bij (asbf vâ»Â¹) Bijection.id |}. + abstract (intros; split; [reflexivity | apply subf_move_along']). Defined. +Global Opaque trans. Definition trans_compat : Proper (equiv ==> equiv) trans := _. -Lemma trans2nat : ∀ (m : nat) (v : ring_node), fin2nat (trans m v) - = (v + (n - m mod ub)) mod ub. -Proof using . apply asbmV2nat. Qed. +Lemma transvE : ∀ v : fin n, trans v = asbf vâ»Â¹ :> bijection (fin n). +Proof using . reflexivity. Qed. -Lemma transV2nat : ∀ (m : nat) (v : ring_node), - fin2nat ((trans m)â»Â¹ v) = (v + m) mod ub. -Proof using . apply asbm2nat. Qed. +Lemma transvVE : ∀ v : fin n, trans vâ»Â¹ == asbf v :> bijection (fin n). +Proof using . intros. rewrite <- (Bijection.inv_inv (asbf v)). reflexivity. Qed. -Lemma trans_mod : ∀ m : nat, trans (m mod ub) == trans m. +Lemma transeE : + ∀ v : fin n, iso_E (trans v) = prod_eq_bij (asbf vâ»Â¹) Bijection.id. +Proof using . reflexivity. Qed. + +Lemma transeVE : + ∀ v : fin n, iso_E (trans vâ»Â¹) == prod_eq_bij (asbf v) Bijection.id. Proof using . - intros. split; [|split]. all: unfold trans. - all: cbn-[equiv inverse Nat.modulo]. 1,2: rewrite asbm_mod. - all: reflexivity. + intros. rewrite <- (Bijection.inv_inv (asbf v)), + <- Bijection.id_inv, <- prod_eq_bij_inv. reflexivity. Qed. -Lemma trans0 : trans 0 == id. +Lemma trans0 : trans fin0 == id. Proof using . - split; [|split]. all: unfold trans. - all: cbn-[equiv inverse Nat.modulo]. 1,2: rewrite asbm0, Bijection.id_inv. - 2: rewrite prod_bij_id. all: reflexivity. + split. rewrite transvE. 2: rewrite transeE. all: rewrite asbf0, + Bijection.id_inv. reflexivity. apply prod_eq_bij_id. Qed. -Lemma transs : ∀ v : ring_node, trans v v = fin0. -Proof using . apply asbmmV. Qed. - -Lemma transV0 : ∀ v : ring_node, (trans v)â»Â¹ fin0 = v. -Proof using . apply asb0m. Qed. +Lemma transI : ∀ v1 v2 : fin n, trans v1 == trans v2 → v1 = v2. +Proof using . intros * H. apply subfI with (f1:=v1), H. Qed. -Lemma trans_fin_max : trans 1 fin0 = fin_max. -Proof using . apply asbmV_fin_max. Qed. +Lemma transVI : ∀ v1 v2 : fin n, (trans v1)â»Â¹ == (trans v2)â»Â¹ → v1 = v2. +Proof using . intros * H. apply addfI with (f1:=v1), H. Qed. -Lemma transI_fin_max : (trans 1)â»Â¹ fin_max = fin0. -Proof using . apply asbm_fin_max. Qed. - -Lemma transI_fin0_divide : ∀ (v : ring_node) (m : nat), - Nat.divide n (v + m) ↔ (trans m)â»Â¹ v = fin0. -Proof using . apply asbm_fin0_divide. Qed. - -Lemma trans_locally_inj : ∀ (m1 m2 : nat), - m1 ≤ m2 → m2 - m1 < n → trans m1 == trans m2 → m1 = m2. +Lemma transAC : + ∀ v1 v2 : fin n, trans v2 ∘ (trans v1) == trans v1 ∘ (trans v2). Proof using . - intros * ?? H. eapply asbmV_locally_inj. 3: apply H. all: eassumption. + intros. split. rewrite 2 (compvE (trans _)), 2 transvE. apply asbfVAC. + rewrite 2 (compeE (trans _)), 2 transeE. setoid_rewrite prod_eq_bij_comp. + rewrite (asbfVAC v2). reflexivity. Qed. -Lemma transV_locally_inj : ∀ (m1 m2 : nat), - m1 ≤ m2 → m2 - m1 < n → (trans m1)â»Â¹ == (trans m2)â»Â¹ → m1 = m2. +Lemma transVAC : ∀ v1 v2 : fin n, + (trans v2)â»Â¹ ∘ (trans v1)â»Â¹ == (trans v1)â»Â¹ ∘ (trans v2)â»Â¹. Proof using . - intros * ?? H. eapply asbm_locally_inj. 3: apply H. all: eassumption. + intros. split. rewrite 2 (compvE (trans _â»Â¹)), 2 transvVE. apply asbfAC. + rewrite 2 (compeE (trans _â»Â¹)), 2 transeVE. setoid_rewrite prod_eq_bij_comp. + rewrite (asbfAC v2). reflexivity. Qed. -Lemma trans_bounded_diff_inj : ∀ (p : nat) (m1 m2 : nat), - bounded_diff n p m1 → bounded_diff n p m2 - → trans m1 == trans m2 → m1 = m2. +Lemma transCV : ∀ v1 v2 : fin n, + trans v1 ∘ (trans v2)â»Â¹ == (trans v2)â»Â¹ ∘ (trans v1). Proof using . - intros * ?? H. eapply asbmV_bounded_diff_inj. 3: apply H. all: eassumption. + intros. split. rewrite (compvE (trans _)), (compvE (trans _â»Â¹)), transvE, + transvVE. symmetry. apply asbfCV. rewrite (compeE (trans _)), + (compeE (trans _â»Â¹)), transeE, transeVE. setoid_rewrite prod_eq_bij_comp. + rewrite (asbfCV v2). reflexivity. Qed. -Lemma transV_bounded_diff_inj : ∀ (p : nat) (m1 m2 : nat), - bounded_diff n p m1 → bounded_diff n p m2 - → (trans m1)â»Â¹ == (trans m2)â»Â¹ → m1 = m2. +Lemma transA : + ∀ v1 v2 : fin n, trans (trans v1 v2) == (trans v1â»Â¹) ∘ (trans v2). Proof using . - intros * ?? H. eapply asbm_bounded_diff_inj. 3: apply H. all: eassumption. + intros. split. rewrite (compvE (trans _â»Â¹)), 3 transvE, transvVE. + apply asbfVAV. rewrite (compeE (trans _â»Â¹)), transvE, 2 transeE, transeVE. + setoid_rewrite prod_eq_bij_comp. rewrite asbfVAV, Bijection.id_comp_l. + reflexivity. Qed. -Lemma transIC : ∀ v1 v2 : ring_node, (trans v1)â»Â¹ v2 = (trans v2)â»Â¹ v1. -Proof using . apply asbmC. Qed. - -Lemma transAC : ∀ (m1 m2 : nat), trans m2 ∘ (trans m1) == trans m1 ∘ (trans m2). +Lemma transAV : + ∀ v1 v2 : fin n, trans ((trans v1)â»Â¹ v2) == (trans v1) ∘ (trans v2). Proof using . - intros. split; [|split]. all: unfold trans. - all: cbn-[equiv BijectionInverse BijectionComposition]. apply asbmVAC. - setoid_rewrite prod_bij_comp. setoid_rewrite (asbmVAC m2 m1). - all: reflexivity. + intros. split. rewrite (compvE (trans _)), 3 transvE, transvVE. + apply asbfVA. rewrite (compeE (trans _)), transvVE, 3 transeE. + setoid_rewrite prod_eq_bij_comp. rewrite asbfVA, Bijection.id_comp_l. + reflexivity. Qed. -Lemma transVAC : ∀ (m1 m2 : nat), - (trans m2)â»Â¹ ∘ (trans m1)â»Â¹ == (trans m1)â»Â¹ ∘ (trans m2)â»Â¹. +Lemma transVA : + ∀ v1 v2 : fin n, trans (trans v1 v2)â»Â¹ == (trans v1) ∘ (trans v2â»Â¹). Proof using . - intros. split; [|split]. all: unfold trans. - all: cbn-[equiv BijectionInverse BijectionComposition]. - rewrite 2 Bijection.inv_inv. apply asbmAC. setoid_rewrite prod_bij_inv. - setoid_rewrite prod_bij_comp. setoid_rewrite Bijection.inv_inv. - setoid_rewrite (asbmAC m2 m1). all: reflexivity. + intros. split. rewrite (compvE (trans _)), transvE, 2 transvVE. + apply asbfAV. rewrite (compeE (trans _)), transvE, transeE, 2 transeVE. + setoid_rewrite prod_eq_bij_comp. rewrite asbfAV, Bijection.id_comp_l. + reflexivity. Qed. -Lemma trans_comp : ∀ (m2 m1 : nat), trans m2 ∘ (trans m1) == trans (m1 + m2). +Lemma transVAV : + ∀ v1 v2 : fin n, trans ((trans v1)â»Â¹ v2)â»Â¹ == (trans v1â»Â¹) ∘ (trans v2â»Â¹). Proof using . - intros. split; [|split]. all: unfold trans. - all: cbn-[equiv BijectionInverse BijectionComposition]. apply asbmV_comp. - setoid_rewrite prod_bij_comp. setoid_rewrite asbmV_comp. - all: setoid_rewrite Bijection.id_comp_l. all: reflexivity. + intros. split. rewrite (compvE (trans _â»Â¹)), 3 transvVE. + apply asbfA. rewrite (compeE (trans _â»Â¹)), transvVE, 3 transeVE. + setoid_rewrite prod_eq_bij_comp. + rewrite asbfA, Bijection.id_comp_l. reflexivity. Qed. -Lemma transV_comp : ∀ (m2 m1 : nat), - (trans m2)â»Â¹ ∘ (trans m1)â»Â¹ == (trans (m1 + m2))â»Â¹. +Lemma move_along_trans : ∀ (v1 v2 : fin n) (d : direction), + move_along (trans v1 v2) d == trans v1 (move_along v2 d). Proof using . - intros. split; [|split]. all: unfold trans. - all: cbn-[equiv BijectionInverse BijectionComposition]. - rewrite 3 Bijection.inv_inv. apply asbm_comp. - setoid_rewrite prod_bij_inv. setoid_rewrite prod_bij_comp. - setoid_rewrite Bijection.inv_inv. setoid_rewrite asbm_comp. - all: setoid_rewrite Bijection.id_inv. all: setoid_rewrite Bijection.id_comp_l. - all: reflexivity. + intros. rewrite transvE, (asbfVE v1). symmetry. apply subf_move_along'. Qed. -Lemma transCV : ∀ (m1 m2 : nat), - trans m1 ∘ (trans m2)â»Â¹ == (trans m2)â»Â¹ ∘ (trans m1). +Lemma trans_inj : Preliminary.injective equiv equiv trans. Proof using . - intros. split; [|split]. all: unfold trans. - all: cbn-[equiv BijectionComposition BijectionInverse]. - 2: setoid_rewrite prod_bij_inv. 2: setoid_rewrite prod_bij_comp. - 1,2: setoid_rewrite Bijection.inv_inv. symmetry. apply asbmCV. - all: rewrite Bijection.id_inv. setoid_rewrite (asbmCV m2 m1). - all: reflexivity. + intros v1 v2 [H _]. eapply addfI. rewrite (addfC _ v2). + apply inverse_compat in H. rewrite (transvVE v1), (transvVE v2) in H. + specialize (H v2). rewrite (asbfE v1), (asbfE v2) in H. apply H. Qed. -Lemma move_along_trans : ∀ (v : ring_node) (m : nat) (d : direction), - move_along (trans m v) d == trans m (move_along v d). -Proof using . intros. symmetry. apply subm_move_along. Qed. - -Lemma trans_inj : ∀ m : nat, Util.Preliminary.injective equiv equiv (trans m). -Proof using . intros. apply (@injective _ _ nothresholdRing). Qed. - -Definition sym (c : nat) : isomorphism nothresholdRing. +Definition sym (v : fin n) : isomorphism Ring. Proof using . refine {| - iso_V := symbf (mod2fin c); - iso_E := prod_bij (symbf (mod2fin c)) swap_direction_bij; - iso_T := @Bijection.id _ R_Setoid |}. - - (* iso_morphism *) - abstract (intros; split; [reflexivity | apply symf_move_along]). - - (* iso_threshold *) abstract (reflexivity). - - (* iso_incr *) abstract (intros * H; apply H). - - (* iso_bound_T *) abstract (reflexivity). + iso_V := sybf v; + iso_E := prod_eq_bij (sybf v) swbd |}. + (* iso_morphism *) + abstract (intros; split; [reflexivity | apply symf_move_along]). Defined. +Global Opaque sym. + +Definition sym_compat : Proper (equiv ==> equiv) sym := _. + +Lemma symvE : ∀ v : fin n, sym v = sybf v :> bijection (fin n). +Proof using . reflexivity. Qed. + +Lemma symeE : ∀ v : fin n, iso_E (sym v) = prod_eq_bij (sybf v) swbd. +Proof using . reflexivity. Qed. -Global Instance sym_compat : Proper (equiv ==> equiv) sym := _. +Lemma symvVE : ∀ v : fin n, sym vâ»Â¹ == sybf v :> bijection (fin n). +Proof using . intros v1 v2. rewrite <- (sybfV v1). reflexivity. Qed. -Lemma sym_involutive : ∀ c, (sym c) ∘ (sym c) == id. +Lemma symeVE : ∀ v : fin n, iso_E (sym vâ»Â¹) == prod_eq_bij (sybf v) swbd. Proof using . - intros. split; [|split]. all: apply Bijection.compose_inverse_r. + intros. rewrite <- sybfV, <- swbdV, <- prod_eq_bij_inv. reflexivity. Qed. +Lemma symK : ∀ v : fin n, (sym v) ∘ (sym v) == id. +Proof using . + intros. split. rewrite (compvE (sym v)), symvE. apply sybfK. + rewrite (compeE (sym v)), symeE. setoid_rewrite prod_eq_bij_comp. + rewrite (sybfK v), swbdK, prod_eq_bij_id. reflexivity. +Qed. + +Lemma move_along_sym : ∀ (v1 v2 : fin n) (d : direction), + move_along (sym v1 v2) d = sym v1 (move_along v2 (swap_direction d)). +Proof using . intros. rewrite symvE, (sybfE v1). apply move_along_symf. Qed. + End Ring. diff --git a/Spaces/ThresholdIsomorphism.v b/Spaces/ThresholdIsomorphism.v new file mode 100644 index 0000000000000000000000000000000000000000..8011d799cc64af5a841c05f4700b8ebf5dc92a6f --- /dev/null +++ b/Spaces/ThresholdIsomorphism.v @@ -0,0 +1,100 @@ +Require Import Utf8 SetoidDec Rbase Rbasic_fun Psatz. +From Pactole Require Import Util.Coqlib Util.Bijection Util.Ratio. +From Pactole Require Import Spaces.Graph Spaces.Isomorphism. +Set Implicit Arguments. + +Section ThresholdIsomorphism. + +Context {V E : Type}. +Context {G : ThresholdGraph V E}. + +Record threshold_isomorphism := { + iso_VE :> isomorphism G; + iso_T : bijection strict_ratio; + iso_threshold : ∀ e : E, iso_T (threshold e) == threshold (iso_E iso_VE e) :> strict_ratio; + iso_incr : ∀ a b : strict_ratio, (a < b)%R -> (iso_T a < iso_T b)%R }. + +Global Instance threshold_isomorphism_Setoid : Setoid threshold_isomorphism. +Proof using . + simple refine {| + equiv := λ iso1 iso2, iso1.(iso_VE) == iso2.(iso_VE) + /\ iso1.(iso_T) == iso2.(iso_T) |}; autoclass. split. + + intro f. now repeat split. + + intros f g Hfg; destruct Hfg as [HVE HT]. split; now symmetry. + + intros f g h Hfg Hgh. destruct Hfg as [? ?], Hgh as [? ?]. + split; etransitivity; eauto. +Defined. + +Instance iso_T_compat : Proper (equiv ==> equiv) iso_T. +Proof using . intros ? ? Heq ?. now apply Heq. Qed. + +Definition id : threshold_isomorphism. +Proof using . + refine {| iso_VE := id; + iso_T := Bijection.id |}. + + now intros. + + now intros. +Defined. + +Definition comp (f g : threshold_isomorphism) : threshold_isomorphism. +Proof using . + refine {| + iso_VE := compose f.(iso_VE) g.(iso_VE); + iso_T := compose f.(iso_T) g.(iso_T) |}. + + intro. simpl. now rewrite 2 iso_threshold. + + intros. simpl. now do 2 apply iso_incr. +Defined. + +Global Instance TIsoComposition : Composition threshold_isomorphism. +Proof using . + refine {| compose := comp |}. + intros f1 f2 Hf g1 g2 Hg. split. all: apply compose_compat. + 3,4: f_equiv. 1,3: apply Hf. all: apply Hg. +Defined. + +Lemma compose_assoc : ∀ f g h, f ∘ (g ∘ h) == (f ∘ g) ∘ h. +Proof using . intros f g h; repeat split; simpl; reflexivity. Qed. + +Definition inv (tiso : threshold_isomorphism) : threshold_isomorphism. +Proof using . + refine {| iso_VE := inverse tiso.(iso_VE); + iso_T := inverse tiso.(iso_T) |}. + + intros e. cbn-[equiv]. rewrite <- Inversion, iso_threshold, + section_retraction. reflexivity. + + intros a b Hab. apply Rnot_le_lt. intros [Hd | Hd]. + all: apply (Rlt_not_le b a Hab). left. erewrite <- (section_retraction _ b), + <- (section_retraction _ a). apply iso_incr, Hd. right. + apply proj_ratio_compat, proj_strict_ratio_compat, + (injective (iso_T tiso â»Â¹)), proj_strict_ratio_inj, proj_ratio_inj, Hd. +Defined. + +Global Instance TIsoInverse : Inverse threshold_isomorphism. +Proof using . + refine {| inverse := inv |}. + intros f g [? ?]. split. all: apply inverse_compat. all: assumption. +Defined. + +Lemma id_inv : idâ»Â¹ == id. +Proof using . split. apply id_inv. setoid_rewrite Bijection.id_inv. reflexivity. Qed. + +Lemma id_comp_l : ∀ tiso : threshold_isomorphism, id ∘ tiso == tiso. +Proof using. intros. split. apply id_comp_l. setoid_rewrite Bijection.id_comp_l. reflexivity. Qed. + +Lemma id_comp_r : ∀ tiso : threshold_isomorphism, tiso ∘ id == tiso. +Proof using. intros. split. apply id_comp_r. setoid_rewrite Bijection.id_comp_r. reflexivity. Qed. + +Lemma inv_inv : ∀ tiso : threshold_isomorphism, tisoâ»Â¹â»Â¹ == tiso. +Proof using . intros. split. apply inv_inv. setoid_rewrite Bijection.inv_inv. reflexivity. Qed. + +Lemma compose_inverse_l : ∀ tiso : threshold_isomorphism, tiso â»Â¹ ∘ tiso == id. +Proof using . intro. split. apply compose_inverse_l. setoid_rewrite Bijection.compose_inverse_l. reflexivity. Qed. + +Lemma compose_inverse_r : ∀ tiso : threshold_isomorphism, tiso ∘ (tiso â»Â¹) == id. +Proof using . intro. split. apply compose_inverse_r. setoid_rewrite Bijection.compose_inverse_r. reflexivity. Qed. + +Lemma inverse_compose : ∀ f g : threshold_isomorphism, (f ∘ g) â»Â¹ == (g â»Â¹) ∘ (f â»Â¹). +Proof using . intros f g. split. apply inverse_compose. setoid_rewrite Bijection.inverse_compose. reflexivity. Qed. + +End ThresholdIsomorphism. + +Arguments threshold_isomorphism {V} {E} G. diff --git a/Util/Bijection.v b/Util/Bijection.v index afc072664f6728391a250c4efc692bc1036bab51..5d8b1ec72c1a099d17e4789c046b344f85e28195 100644 --- a/Util/Bijection.v +++ b/Util/Bijection.v @@ -22,7 +22,7 @@ Set Implicit Arguments. (** Bijections on a type [T] with an equivalence relation [eqT] *) Section Bijections. -Context (T : Type). +Context {T : Type}. Context {HeqT : Setoid T}. Record bijection := { @@ -33,13 +33,18 @@ Record bijection := { Global Existing Instance section_compat. Global Instance bij_Setoid : Setoid bijection. -simple refine {| equiv := fun f g => forall x, f.(section) x == g x |}; auto; []. +simple refine {| equiv := λ f g, ∀ x, f.(section) x == g x |}; auto; []. Proof. split. + repeat intro. reflexivity. + repeat intro. now symmetry. + repeat intro. etransitivity; eauto. Defined. +(* Use fun_Setoid directly in the definition of bij_Setoid? *) +Global Instance bij_Setoid_eq_compat : Proper ((@equiv bijection bij_Setoid) + ==> (@equiv (T -> T) (fun_Setoid T HeqT))) section. +Proof using . intros ?? H. apply H. Qed. + Global Instance section_full_compat : Proper (equiv ==> (equiv ==> equiv)) section. Proof using . intros f g Hfg x y Hxy. rewrite Hxy. now apply Hfg. Qed. @@ -68,6 +73,10 @@ Proof. intros f1 f2 Hf g1 g2 Hg x. cbn -[equiv]. rewrite (Hf (g1 x)). f_equiv. apply Hg. Defined. + +Lemma compE : ∀ f g : bijection, f ∘ g = (λ t, f (g t)) :> (T -> T). +Proof using . reflexivity. Qed. + (* Global Instance compose_compat : Proper (equiv ==> equiv ==> equiv) compose. Proof. @@ -141,8 +150,11 @@ Arguments bijection T {_}. Arguments section {_} {_} !_ x. Arguments retraction {_} {_} !_ x. -Definition prod_bij {A B : Type} (SA : Setoid A) (SB : Setoid B) - (BA : bijection A) (BB : bijection B) : bijection (A * B). +Section prod_bij. + +Context {A B : Type} (SA : Setoid A) (SB : Setoid B). + +Definition prod_bij (BA : bijection A) (BB : bijection B) : bijection (A * B). Proof using . refine {| section := λ p, (BA (fst p), BB (snd p)); @@ -152,23 +164,94 @@ Proof using . apply Inversion; reflexivity). Defined. -Global Instance prod_bij_compat : ∀ {A B : Type} (SA : Setoid A) (SB : Setoid B), - Proper (equiv ==> equiv ==> equiv) (@prod_bij A B SA SB). +Lemma prod_bijE : ∀ (BA : bijection A) (BB : bijection B), + prod_bij BA BB = (λ p : A * B, (BA (fst p), BB (snd p))) :> (A * B -> A * B). +Proof using . reflexivity. Qed. + +Lemma prod_bijVE : ∀ (BA : bijection A) (BB : bijection B), + prod_bij BA BBâ»Â¹ = (λ p, (BAâ»Â¹ (fst p), BBâ»Â¹ (snd p))) :> (A * B -> A * B). +Proof using . reflexivity. Qed. + +Global Instance prod_bij_compat : Proper (equiv ==> equiv ==> equiv) prod_bij. Proof using . - intros * BA1 BA2 HA BB1 BB2 HB p. cbn. split. rewrite HA. reflexivity. - rewrite HB. reflexivity. + intros * BA1 BA2 HA BB1 BB2 HB p. rewrite 2 prod_bijE. split. + rewrite HA. reflexivity. rewrite HB. reflexivity. Qed. -Lemma prod_bij_id : ∀ {A B : Type} (SA : Setoid A) (SB : Setoid B), - prod_bij (@id A SA) (@id B SB) == (@id _ (prod_Setoid SA SB)). -Proof using . intros * p. unfold prod_bij. cbn. split. all: reflexivity. Qed. +Lemma prod_bij_id : prod_bij (@id A SA) (@id B SB) == id. +Proof using . intros * p. rewrite prod_bijE. split. all: reflexivity. Qed. -Lemma prod_bij_comp : ∀ {A B : Type} (SA : Setoid A) (SB : Setoid B) - (BA1 BA2 : bijection A) (BB1 BB2 : bijection B), +Lemma prod_bij_comp : ∀ (BA1 BA2 : bijection A) (BB1 BB2 : bijection B), prod_bij BA1 BB1 ∘ (prod_bij BA2 BB2) == prod_bij (BA1 ∘ BA2) (BB1 ∘ BB2). -Proof using . intros * p. cbn. split. all: reflexivity. Qed. +Proof using . intros * p. rewrite compE, 3 prod_bijE, 2 compE. reflexivity. Qed. -Lemma prod_bij_inv : ∀ {A B : Type} (SA : Setoid A) (SB : Setoid B) - (BA : bijection A) (BB : bijection B), +Lemma prod_bij_inv : ∀ (BA : bijection A) (BB : bijection B), (prod_bij BA BB)â»Â¹ == prod_bij (BAâ»Â¹) (BBâ»Â¹). -Proof using . intros * p. cbn. split. all: reflexivity. Qed. +Proof using . intros * p. rewrite prod_bijVE, prod_bijE. reflexivity. Qed. + +Definition prod_eq_bij (BA : @bijection A (eq_setoid A)) + (BB : @bijection B (eq_setoid B)) : @bijection (A * B) (eq_setoid (A * B)). +Proof using . + refine {| + section := λ p, (BA (fst p), BB (snd p)); + retraction := λ p, (BAâ»Â¹ (fst p), BBâ»Â¹ (snd p)) |}. + abstract (intros; split; cbn; intros H; rewrite <- H, <- pair_eqE; + cbn[fst snd]; [rewrite 2 retraction_section | rewrite 2 section_retraction]; + split; reflexivity). +Defined. + +Lemma prod_eq_bijE : ∀ (BA : bijection A) (BB : bijection B), prod_eq_bij BA BB + = (λ p : A * B, (BA (fst p), BB (snd p))) :> (A * B -> A * B). +Proof using . reflexivity. Qed. + +Lemma prod_eq_bijVE : ∀ (BA : bijection A) (BB : bijection B), + prod_eq_bij BA BBâ»Â¹ = (λ p, (BAâ»Â¹ (fst p), BBâ»Â¹ (snd p))) :> (A * B -> A * B). +Proof using . reflexivity. Qed. + +Global Instance prod_eq_bij_compat : + Proper (equiv ==> equiv ==> equiv) prod_eq_bij. +Proof using . + cbn. intros * BA1 BA2 HA BB1 BB2 HB p. rewrite 2 prod_eq_bijE, <- pair_eqE. + split. rewrite HA. reflexivity. rewrite HB. reflexivity. +Qed. + +Lemma prod_eq_bij_id : prod_eq_bij id id == id. +Proof using . + intros * p. rewrite prod_eq_bijE. apply pair_eqE. split. all: reflexivity. +Qed. + +Lemma prod_eq_bij_comp : ∀ (BA1 BA2 : bijection A) (BB1 BB2 : bijection B), + prod_eq_bij BA1 BB1 ∘ (prod_eq_bij BA2 BB2) + == prod_eq_bij (BA1 ∘ BA2) (BB1 ∘ BB2). +Proof using . + intros * p. rewrite compE, 3 prod_eq_bijE, 2 compE. reflexivity. +Qed. + +Lemma prod_eq_bij_inv : ∀ (BA : bijection A) (BB : bijection B), + (prod_eq_bij BA BB)â»Â¹ == prod_eq_bij (BAâ»Â¹) (BBâ»Â¹). +Proof using . intros * p. rewrite prod_eq_bijVE, prod_eq_bijE. reflexivity. Qed. + +End prod_bij. + +Section equiv_bij. + +Context {T : Type} [S1 S2 : Setoid T]. +Context (H : ∀ t1 t2 : T, @equiv T S1 t1 t2 <-> @equiv T S2 t1 t2). + +Definition equiv_bij : @bijection T S1 -> @bijection T S2. +Proof using H. + intros [s r c p]. refine {| section := s; retraction := r |}. + abstract (intros ?? Heq; apply H, c, H, Heq). + abstract (intros; rewrite <-2 H; apply p). +Defined. + +Lemma equiv_bijE : ∀ b : @bijection T S1, equiv_bij b = b :> (T -> T). +Proof using . intros []. reflexivity. Qed. + +Lemma equiv_bijVE : ∀ b : @bijection T S1, equiv_bij bâ»Â¹ = bâ»Â¹ :> (T -> T). +Proof using . intros []. reflexivity. Qed. + +Lemma equiv_bij_id : equiv_bij (@id T S1) == (@id T S2). +Proof using . intros t. rewrite equiv_bijE. reflexivity. Qed. + +End equiv_bij. diff --git a/Util/Fin.v b/Util/Fin.v index 65f0b6a4a43e4630f8e819d9a705954e18a1c960..c1fcbaafb9eaf54b01764dcea07579b90c8966a5 100644 --- a/Util/Fin.v +++ b/Util/Fin.v @@ -5,19 +5,26 @@ Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. +(* + * + * The type fin u with u any nat ≠0 + * consisting of the nats smaller than u + * + *) + Section Fin. -Context (ub : nat). +Context (u : nat). -Inductive fin : Type := Fin : ∀ m : nat, m < ub -> fin. +Inductive fin : Type := Fin : ∀ j : nat, j < u -> fin. Global Instance fin_Setoid : Setoid fin := eq_setoid fin. Lemma fin_dec : ∀ f1 f2 : fin, {f1 = f2} + {f1 <> f2}. Proof using . - intros [m1 H1] [m2 H2]. destruct (Nat.eq_dec m1 m2). - + subst. left. f_equal. apply le_unique. - + right. intro Habs. inv Habs. auto. + intros [j1 H1] [j2 H2]. destruct (Nat.eq_dec j1 j2). + - subst. left. f_equal. apply le_unique. + - right. intro Habs. inv Habs. auto. Qed. Global Instance fin_EqDec : EqDec fin_Setoid := fin_dec. @@ -26,932 +33,1667 @@ Global Coercion fin2nat (f : fin) : nat := match f with @Fin m _ => m end. Definition fin2nat_compat : Proper (equiv ==> equiv) fin2nat := _. -Lemma fin_lt : ∀ f : fin, f < ub. -Proof using . intros. destruct f as [m p]. exact p. Qed. +Lemma fin_lt : ∀ f : fin, f < u. +Proof using . intros. destruct f as [j p]. exact p. Qed. -Lemma fin_le : ∀ f : fin, f <= Nat.pred ub. +Lemma fin_le : ∀ f : fin, f <= Nat.pred u. Proof using . intros. apply Nat.lt_le_pred, fin_lt. Qed. -Lemma fin2nat_inj : Preliminary.injective equiv equiv fin2nat. +Lemma fin_between : ∀ f : fin, 0 <= f < u. +Proof using . intros. split. apply Nat.le_0_l. apply fin_lt. Qed. + +Lemma fin2natI : Preliminary.injective equiv equiv fin2nat. Proof using . - intros [m1 H1] [m2 H2] Heq. cbn in Heq. subst. cbn. f_equal. apply le_unique. + intros [j1 H1] [j2 H2] Heq. cbn in Heq. subst. cbn. f_equal. apply le_unique. Qed. -End Fin. +Lemma lt_gt_cases : ∀ f1 f2 : fin, f1 ≠f2 <-> f1 < f2 ∨ f2 < f1. +Proof using . + intros. erewrite not_iff_compat. apply Nat.lt_gt_cases. + symmetry. apply (injective_eq_iff fin2natI). +Qed. -Class greater_than (lb : nat) := { - ub : nat; - lt_lb_ub : lb < ub; -}. +Lemma le_lt_eq_dec : ∀ f1 f2 : fin, f1 <= f2 -> {f1 < f2} + {f1 = f2}. +Proof using . + intros * H. destruct (le_lt_eq_dec f1 f2 H) as [Hd | Hd]. + left. apply Hd. right. apply fin2natI, Hd. +Qed. -Section greater_than. +Lemma lt_eq_lt_dec : ∀ f1 f2 : fin, {f1 < f2} + {f1 = f2} + {f2 < f1}. +Proof using . + intros. destruct (lt_eq_lt_dec f1 f2) as [[Hd | Hd] | Hd]. + left. left. apply Hd. left. right. apply fin2natI, Hd. right. apply Hd. +Qed. -Context {lb : nat} {gtlb : greater_than lb}. +End Fin. -Lemma neq_ub_lb : ub ≠lb. -Proof using gtlb. - intros * H. apply (lt_irrefl lb). rewrite <- H at 2. exact lt_lb_ub. -Qed. +Section mod2fin. -Lemma le_lb_pred_ub : lb <= Nat.pred ub. -Proof using gtlb. intros. apply Nat.lt_le_pred, lt_lb_ub. Qed. +Context {l u : nat} {ltc_l_u : l <c u}. -Lemma lt_pred_ub : Nat.pred ub < ub. -Proof using gtlb. - intros. eapply lt_pred_n_n, Nat.le_lt_trans. apply Nat.le_0_l. apply lt_lb_ub. -Qed. +(* + * + * The fin u whose underlying nat is 0 + * + *) -Lemma S_pred_ub : S (Nat.pred ub) = ub. -Proof using gtlb. intros. eapply Nat.lt_succ_pred, lt_lb_ub. Qed. +Definition fin0 : fin u := Fin lt_0_u. -Lemma lt_sb_ub : ∀ sb : nat, sb < lb -> sb < ub. -Proof using gtlb. intros * H. eapply Nat.lt_trans. exact H. exact lt_lb_ub. Qed. +Lemma fin02nat : fin0 = 0 :> nat. +Proof using . reflexivity. Qed. -Lemma lt_sb_pred_ub : ∀ sb : nat, sb < lb -> sb < Nat.pred ub. -Proof using gtlb. - intros * H. eapply Nat.lt_le_trans. exact H. exact le_lb_pred_ub. +Lemma all_fin0 : ∀ f : fin u, u = 1 -> f = fin0. +Proof using . + intros f H. apply fin2natI. rewrite fin02nat. subst. + pose proof (fin_lt f) as H. inversion H as [H0 | n H0 H1]. + apply H0. inversion H0. Qed. -Lemma neq_pred_ub_sb : ∀ sb : nat, sb < lb -> Nat.pred ub ≠sb. -Proof using gtlb. - intros * H Habs. eapply lt_not_le. 2: eapply le_lb_pred_ub. - rewrite Habs. exact H. +Lemma all_eq : ∀ f1 f2 : fin u, u = 1 -> f1 = f2. +Proof using ltc_l_u. + intros * H. rewrite (all_fin0 f1), (all_fin0 f2). reflexivity. all: apply H. Qed. -Definition smaller_bound : ∀ sb : nat, sb < lb -> greater_than sb. -Proof using gtlb. - intros * H. constructor 1 with ub. apply lt_sb_ub, H. -Defined. +Lemma fin0_le : ∀ f : fin u, fin0 <= f. +Proof using . intros. rewrite fin02nat. apply Nat.le_0_l. Qed. -Lemma smaller_bound_ub : ∀ (sb : nat) (H : sb < lb), - @ub sb (smaller_bound H) = @ub lb gtlb. -Proof using . reflexivity. Qed. +Lemma neq_fin0_lt_fin0 : ∀ f : fin u, f ≠fin0 <-> fin0 < f. +Proof using . + intros. rewrite <- Nat.neq_0_lt_0, <- fin02nat. symmetry. + apply not_iff_compat, injective_eq_iff, fin2natI. +Qed. -End greater_than. +Lemma eq_fin0_lt_dec : ∀ f : fin u, {f = fin0} + {fin0 < f}. +Proof using . + intros. destruct (eq_0_lt_dec f) as [Hd | Hd]. + left. apply fin2natI, Hd. right. apply Hd. +Qed. -Section prod_bound. +Lemma le_fin0 : ∀ f : fin u, f <= fin0 <-> f = fin0. +Proof using . + intros. rewrite fin02nat, Nat.le_0_r, <- fin02nat. + apply injective_eq_iff, fin2natI. +Qed. -Context {lb1 lb2 : nat} (gtlb1 : greater_than lb1) (gtlb2 : greater_than lb2). +Lemma lt_fin0 : ∀ f1 f2 : fin u, f1 < f2 -> fin0 < f2. +Proof using . intros * H. eapply Nat.lt_lt_0, H. Qed. -Definition prod_bound : greater_than (Nat.pred ((S lb1) * (S lb2))). -Proof using gtlb1 gtlb2. - constructor 1 with (@ub lb1 gtlb1 * @ub lb2 gtlb2). - abstract (erewrite <- Nat.le_succ_l, Nat.lt_succ_pred; - [ apply Nat.mul_le_mono; apply Nat.le_succ_l, lt_lb_ub - | apply Nat.mul_pos_pos; apply Nat.lt_0_succ]). -Defined. +(* + * + * The fin u whose underlying nat is Nat.pred u + * + *) -Lemma prod_bound_ub : @ub _ prod_bound = @ub lb1 gtlb1 * @ub lb2 gtlb2. -Proof using . reflexivity. Qed. +Definition fin_max : fin u := Fin lt_pred_u. -End prod_bound. +Lemma fin_max2nat : fin_max = Nat.pred u :> nat. +Proof using . reflexivity. Qed. -Section mod2fin. +Lemma le_max : ∀ f : fin u, f <= fin_max. +Proof using . + intros. apply lt_n_Sm_le. rewrite fin_max2nat, S_pred_u. apply fin_lt. +Qed. -Context {lb : nat} {gtlb : greater_than lb}. +Lemma max_le : ∀ f : fin u, fin_max <= f <-> f = fin_max. +Proof using . + intros. rewrite Nat.le_lteq, Nat.lt_nge. etransitivity. split. + - intros [H | H]. contradict H. apply le_max. apply H. + - intros H. right. apply H. + - rewrite Nat.eq_sym_iff. apply injective_eq_iff, fin2natI. +Qed. -Definition greater_than_0 : greater_than 0. -Proof using gtlb. - destruct (Nat.eq_dec 0 lb) as [Hd | Hd]. { subst. exact gtlb. } - apply smaller_bound, neq_0_lt, Hd. -Defined. +Lemma neq_max_lt_max : ∀ f : fin u, f ≠fin_max <-> f < fin_max. +Proof using . intros. rewrite <- max_le. apply Nat.nle_gt. Qed. -Lemma greater_than_0_ub : @ub 0 greater_than_0 = @ub lb gtlb. +Lemma eq_max_lt_dec : ∀ f : fin u, {f = fin_max} + {f < fin_max}. Proof using . - unfold greater_than_0. destruct (Nat.eq_dec 0 lb) as [Hd | Hd]. - { subst. reflexivity. } apply smaller_bound_ub. + intros. destruct (fin_dec f fin_max) as [Hd | Hd]. + left. apply Hd. right. apply neq_max_lt_max, Hd. Qed. -Lemma lt_0_ub : 0 < ub. -Proof using . rewrite <- greater_than_0_ub. apply lt_lb_ub. Qed. +Lemma max_lt : ∀ f1 f2 : fin u, f2 < f1 -> f2 < fin_max. +Proof using . + intros * H. rewrite Nat.lt_nge, max_le. intros Ha. + subst. eapply Nat.lt_nge. apply H. apply le_max. +Qed. -Lemma neq_ub_0 : ub ≠0. -Proof using gtlb. rewrite <- greater_than_0_ub. apply neq_ub_lb. Qed. +(* + * + * Building the fin u whose underlying nat + * is j mod u with j the input of the function + * + *) -Lemma le_0_pred_ub : 0 <= Nat.pred ub. -Proof using gtlb. rewrite <- greater_than_0_ub. apply le_lb_pred_ub. Qed. +Lemma mod2fin_lt : ∀ j : nat, j mod u < u. +Proof using ltc_l_u. intros. apply Nat.mod_upper_bound, neq_u_0. Qed. -Lemma mod2fin_lt : ∀ m : nat, m mod ub < ub. -Proof using . intros. apply Nat.mod_upper_bound, neq_ub_0. Qed. +Lemma mod2fin_le : ∀ j : nat, j mod u <= Nat.pred u. +Proof using ltc_l_u. intros. apply Nat.lt_le_pred, mod2fin_lt. Qed. -Definition mod2fin (m : nat) : fin ub := Fin (mod2fin_lt m). +Definition mod2fin (j : nat) : fin u := Fin (mod2fin_lt j). Definition mod2fin_compat : Proper (equiv ==> equiv) mod2fin := _. -Lemma mod2fin_mod : ∀ m : nat, mod2fin (m mod ub) = mod2fin m. +Lemma mod2fin2nat : ∀ j : nat, mod2fin j = j mod u :> nat. +Proof using . intros. reflexivity. Qed. + +Lemma mod2fin_mod : ∀ j : nat, mod2fin (j mod u) = mod2fin j. Proof using . - intros. apply fin2nat_inj. cbn-[Nat.modulo]. apply Nat.mod_mod, neq_ub_0. + intros. apply fin2natI. rewrite 2 mod2fin2nat. apply Nat.mod_mod, neq_u_0. Qed. -Lemma mod2fin2nat : ∀ m : nat, mod2fin m = m mod ub :>nat. -Proof using . intros. reflexivity. Qed. +Lemma mod2fin_mod2fin : ∀ j : nat, mod2fin (mod2fin j) = mod2fin j. +Proof using . intros. rewrite mod2fin2nat. apply mod2fin_mod. Qed. + +Lemma mod2finK : ∀ f : fin u, mod2fin f = f. +Proof using . intros. apply fin2natI, Nat.mod_small, fin_lt. Qed. -Lemma mod2finK : ∀ f : fin ub, mod2fin f = f. -Proof using . intros. apply fin2nat_inj, Nat.mod_small, fin_lt. Qed. +Lemma mod2fin_small : ∀ j : nat, j < u -> mod2fin j = j :> nat. +Proof using . intros * H. rewrite mod2fin2nat. apply Nat.mod_small, H. Qed. -Lemma mod2fin_locally_inj : ∀ m1 m2 : nat, m1 <= m2 - -> m2 - m1 < ub -> mod2fin m1 = mod2fin m2 -> m1 = m2. +Lemma mod2fin_bounded_diffI : ∀ j1 j2 : nat, j1 <= j2 + -> j2 - j1 < u -> mod2fin j1 = mod2fin j2 -> j1 = j2. Proof using . - intros * Hleq Hsub Heq. eapply mod_locally_inj. apply neq_ub_0. + intros * Hleq Hsub Heq. eapply mod_bounded_diffI. apply neq_u_0. exact Hleq. exact Hsub. rewrite <-2 mod2fin2nat, Heq. reflexivity. Qed. -Lemma mod2fin_bounded_diff_inj : ∀ p m1 m2 : nat, bounded_diff ub p m1 - -> bounded_diff ub p m2 -> mod2fin m1 = mod2fin m2 -> m1 = m2. +Lemma mod2fin_betweenI : ∀ p j1 j2 : nat, p <= j1 < p + u + -> p <= j2 < p + u -> mod2fin j1 = mod2fin j2 -> j1 = j2. Proof using . - intros p. eapply (locally_bounded_diff _ - (inj_sym _ _ nat_Setoid (eq_setoid (fin ub)) mod2fin)). - intros * H1 H2. apply mod2fin_locally_inj. all: assumption. + intros p. eapply (bounded_diff_between _ + (inj_sym _ _ nat_Setoid (eq_setoid (fin u)) mod2fin)). + intros * H1 H2. apply mod2fin_bounded_diffI. all: assumption. Qed. -Definition fin0 : fin ub := Fin lt_0_ub. +Lemma mod2fin_muln : ∀ j : nat, mod2fin (u * j) = fin0. +Proof using . + intros. apply fin2natI. rewrite mod2fin2nat, fin02nat, Nat.mul_comm. + apply Nat.mod_mul, neq_u_0. +Qed. -Lemma fin02nat : fin0 = 0 :>nat. -Proof using . reflexivity. Qed. +Lemma mod2fin_le_between_compat : ∀ p j1 j2 : nat, u * p <= j1 < u * p + u + -> u * p <= j2 < u * p + u -> j1 <= j2 -> mod2fin j1 <= mod2fin j2. +Proof using . + intros * Hb1 Hb2 Hle. rewrite 2 mod2fin2nat. + eapply mod_le_between_compat. all: eassumption. +Qed. -Lemma mod2fin0 : mod2fin 0 = fin0. +Lemma mod2fin_lt_between_compat : ∀ p j1 j2 : nat, u * p <= j1 < u * p + u + -> u * p <= j2 < u * p + u -> j1 < j2 -> mod2fin j1 < mod2fin j2. Proof using . - apply fin2nat_inj. rewrite mod2fin2nat, fin02nat. - apply Nat.mod_0_l, neq_ub_0. + intros * Hb1 Hb2 Hlt. rewrite 2 mod2fin2nat. + eapply mod_lt_between_compat. all: eassumption. Qed. -Definition addm (f : fin ub) (m : nat) : fin ub := mod2fin (f + m). +Lemma addn_mod2fin_idemp_l : ∀ j1 j2 : nat, + mod2fin (mod2fin j1 + j2) = mod2fin (j1 + j2). +Proof using . + intros. apply fin2natI. rewrite 3 mod2fin2nat. + apply Nat.add_mod_idemp_l, neq_u_0. +Qed. + +Lemma addn_mod2fin_idemp_r : ∀ j1 j2 : nat, + mod2fin (j1 + mod2fin j2) = mod2fin (j1 + j2). +Proof using . + intros. apply fin2natI. rewrite 3 mod2fin2nat. + apply Nat.add_mod_idemp_r, neq_u_0. +Qed. + +Lemma divide_fin0_mod2fin : ∀ j : nat, Nat.divide u j -> mod2fin j = fin0. +Proof using . + intros * H. apply fin2natI. rewrite mod2fin2nat. + apply Nat.mod_divide, H. apply neq_u_0. +Qed. + +Lemma mod2fin0 : mod2fin 0 = fin0. +Proof using . apply divide_fin0_mod2fin, Nat.divide_0_r. Qed. + +Lemma mod2fin_max : mod2fin (Nat.pred u) = fin_max. +Proof using . apply fin2natI, mod2fin_small, fin_lt. Qed. + +Lemma mod2fin_u : mod2fin u = fin0. +Proof using . apply divide_fin0_mod2fin, Nat.divide_refl. Qed. + +(* + * + * The successor of fin0 + * + *) + +Definition fin1 : fin u := mod2fin 1. + +Lemma fin1E : fin1 = mod2fin 1. +Proof using . reflexivity. Qed. + +Lemma fin12nat {ltc_1_u : 1 <c u} : fin1 = 1 :> nat. +Proof using . rewrite fin1E. apply mod2fin_small, ltc_1_u. Qed. + +(* + * + * Adding either a nat (addm) or + * a fin u (addf) to a fin u + * + *) + +Definition addm (f : fin u) (j : nat) : fin u := mod2fin (f + j). Definition addm_compat : Proper (equiv ==> equiv ==> equiv) addm := _. -Lemma addm2nat : ∀ (f : fin ub) (m : nat), - addm f m = (f + m) mod ub :>nat. -Proof using . unfold addm. intros. apply mod2fin2nat. Qed. +Lemma addmE : ∀ (f : fin u) (j : nat), addm f j = mod2fin (f + j). +Proof using . reflexivity. Qed. + +Lemma addm2nat : ∀ (f : fin u) (j : nat), addm f j = (f + j) mod u :> nat. +Proof using . intros. rewrite addmE. apply mod2fin2nat. Qed. -Lemma addm_mod : ∀ (f : fin ub) (m : nat), addm f (m mod ub) = addm f m. +Lemma addm_mod : ∀ (f : fin u) (j : nat), addm f (j mod u) = addm f j. Proof using . - intros. apply fin2nat_inj. rewrite 2 addm2nat. - apply Nat.add_mod_idemp_r, neq_ub_0. + intros. rewrite 2 addmE, <- mod2fin2nat. apply addn_mod2fin_idemp_r. Qed. -Lemma addIm : ∀ m : nat, - Preliminary.injective equiv equiv (λ f : fin ub, addm f m). +Lemma addm_mod2fin : ∀ (f : fin u) (j : nat), addm f (mod2fin j) = addm f j. +Proof using . intros. rewrite mod2fin2nat. apply addm_mod. Qed. + +Definition addf (f1 f2 : fin u) : fin u := addm f1 f2. + +Definition addf_compat : Proper (equiv ==> equiv ==> equiv) addf := _. + +Lemma addfE : ∀ f1 f2 : fin u, addf f1 f2 = mod2fin (f1 + f2). +Proof using . reflexivity. Qed. + +Lemma addf2nat : ∀ f1 f2 : fin u, addf f1 f2 = (f1 + f2) mod u :> nat. +Proof using . intros. apply addm2nat. Qed. + +Lemma addf_addm : ∀ f1 f2 : fin u, addf f1 f2 = addm f1 f2. +Proof using . reflexivity. Qed. + +Lemma addm_addf : ∀ (f : fin u) (j : nat), addm f j = addf f (mod2fin j). +Proof using . intros. rewrite addf_addm. symmetry. apply addm_mod2fin. Qed. + +Lemma addfC : ∀ f1 f2 : fin u, addf f1 f2 = addf f2 f1. +Proof using . intros. rewrite 2 addfE, Nat.add_comm. reflexivity. Qed. + +Lemma addmC : ∀ (f : fin u) (j : nat), addm f j = addm (mod2fin j) f. +Proof using . intros. rewrite 2 addm_addf, mod2finK. apply addfC. Qed. + +Lemma addmA : ∀ (f1 f2 : fin u) (j : nat), + addm f1 (addm f2 j) = addm (addm f1 f2) j. Proof using . - intros * f1 f2 H. eapply fin2nat_inj, (Nat.add_cancel_r _ _ m). - refine (@mod2fin_bounded_diff_inj m _ _ _ _ H). - all: rewrite plus_comm; split. - 1,3: rewrite (plus_n_O m) at 1. - 1,2: rewrite <-Nat.add_le_mono_l; apply Nat.le_0_l. - all: rewrite <-Nat.add_lt_mono_l; apply fin_lt. + intros. rewrite 4 addmE, addn_mod2fin_idemp_l, + addn_mod2fin_idemp_r, Nat.add_assoc. reflexivity. Qed. -Lemma addm_locally_inj : - ∀ (f : fin ub) (m1 m2 : nat), m1 <= m2 -> m2 - m1 < ub - -> addm f m1 = addm f m2 -> m1 = m2. +Lemma addfA : ∀ f1 f2 f3 : fin u, + addf f1 (addf f2 f3) = addf (addf f1 f2) f3. +Proof using . intros. rewrite 2 (addf_addm _ f3). apply addmA. Qed. + +Lemma addmAC : ∀ (f : fin u) (j1 j2 : nat), + addm (addm f j1) j2 = addm (addm f j2) j1. Proof using . - unfold addm. intros * Hle Hsu Heq. eapply plus_reg_l, mod2fin_locally_inj. - 3: exact Heq. apply Nat.add_le_mono_l, Hle. - rewrite <- minus_plus_simpl_l_reverse. exact Hsu. + intros. rewrite 4 addm_addf, <-2 addfA, (addfC (mod2fin _)). reflexivity. Qed. -Lemma addm_bounded_diff_inj : ∀ (p : nat) (f : fin ub) (m1 m2 : nat), - bounded_diff ub p m1 -> bounded_diff ub p m2 - -> addm f m1 = addm f m2 -> m1 = m2. +Lemma addfAC : ∀ f1 f2 f3 : fin u, + addf (addf f1 f2) f3 = addf (addf f1 f3) f2. +Proof using . intros. rewrite 4 addf_addm. apply addmAC. Qed. + +Lemma addfCA : ∀ f1 f2 f3 : fin u, + addf f1 (addf f2 f3) = addf f2 (addf f1 f3). Proof using . - intros p f. eapply (locally_bounded_diff _ - (inj_sym _ _ nat_Setoid (eq_setoid (fin ub)) (λ x, addm f x))). - intros *. apply addm_locally_inj. + intros. rewrite (addfC f2 f3), addfA, (addfC (addf _ _) _). reflexivity. Qed. -Lemma addm_comp : ∀ (f : fin ub) (m2 m1 : nat), - addm (addm f m1) m2 = addm f (m1 + m2). +Lemma addmCA : ∀ (f1 f2 : fin u) (j : nat), + addm f1 (addm f2 j) = addm f2 (addm f1 j). Proof using . - intros. apply fin2nat_inj. rewrite 3 addm2nat, Nat.add_mod_idemp_l, - Nat.add_assoc by apply neq_ub_0. reflexivity. + intros. rewrite 2 (addm_addf _ j), <-2 addf_addm. apply addfCA. Qed. -Lemma addmAC : ∀ (f : fin ub) (m1 m2 : nat), - addm (addm f m1) m2 = addm (addm f m2) m1. -Proof using . intros. rewrite 2 addm_comp, Nat.add_comm. reflexivity. Qed. +Lemma addfACA : ∀ f1 f2 f3 f4 : fin u, + addf (addf f1 f2) (addf f3 f4) = addf (addf f1 f3) (addf f2 f4). +Proof using . intros. rewrite 2 addfA, (addfAC f1 f3 f2). reflexivity. Qed. -Lemma addm0 : ∀ f : fin ub, addm f 0 = f. +Lemma addmACA : ∀ (f1 f2 f3 : fin u) (j : nat), + addm (addm f1 f2) (addm f3 j) = addm (addm f1 f3) (addm f2 j). Proof using . - intros. apply fin2nat_inj. rewrite addm2nat, Nat.add_0_r, - Nat.mod_small by apply fin_lt. reflexivity. + intros. rewrite 2 (addm_addf _ j), <- 4 addf_addm. apply addfACA. Qed. -Lemma add0m : ∀ f : fin ub, addm fin0 f = f. +Lemma addIm : ∀ j : nat, + Preliminary.injective equiv equiv (λ f : fin u, addm f j). Proof using . - intros. apply fin2nat_inj. rewrite addm2nat, Nat.add_0_l, Nat.mod_small. - reflexivity. apply fin_lt. + intros * f1 f2 H. eapply fin2natI, (Nat.add_cancel_r _ _ j), mod2fin_betweenI. + 1,2: rewrite plus_comm. 1,2: split. 1,3: erewrite <- Nat.add_le_mono_l. + 1,2: apply Nat.le_0_l. 1,2: rewrite Nat.add_0_r, <- Nat.add_lt_mono_l. + 1,2: apply fin_lt. apply H. Qed. -Lemma addm_fin0_divide : ∀ (f : fin ub) (m : nat), - Nat.divide ub (f + m) <-> addm f m = fin0. +Lemma addm_bounded_diffI : + ∀ (f : fin u) (j1 j2 : nat), j1 <= j2 -> j2 - j1 < u + -> addm f j1 = addm f j2 -> j1 = j2. Proof using . - intros. rewrite <- Nat.mod_divide by apply neq_ub_0. split; intros H. - - apply fin2nat_inj. rewrite addm2nat, fin02nat. apply H. - - rewrite <- fin02nat, <- addm2nat. f_equal. apply H. + intros * Hle Hsu H. rewrite 2 addmE in H. eapply plus_reg_l, + mod2fin_bounded_diffI. 3: exact H. apply Nat.add_le_mono_l, Hle. + rewrite <- minus_plus_simpl_l_reverse. exact Hsu. Qed. -Lemma addm_addm2nat : ∀ (f1 f2 : fin ub) (m : nat), - addm f2 (addm f1 m) = addm (addm f2 f1) m. +Lemma addm_betweenI : ∀ (p : nat) (f : fin u) (j1 j2 : nat), + p <= j1 < p + u -> p <= j2 < p + u -> addm f j1 = addm f j2 -> j1 = j2. Proof using . - intros. rewrite addm_comp, addm2nat, addm_mod. reflexivity. + intros p f. eapply (bounded_diff_between _ + (inj_sym _ _ nat_Setoid (eq_setoid (fin u)) (λ x, addm f x))). + intros *. apply addm_bounded_diffI. Qed. -Lemma addmC : ∀ f1 f2 : fin ub, addm f1 f2 = addm f2 f1. +Lemma addIf : ∀ f1 : fin u, + Preliminary.injective equiv equiv (λ f2, addf f2 f1). +Proof using . intros f1 f2 f3. rewrite 2 addf_addm. apply addIm. Qed. + +Lemma addfI : ∀ f1 : fin u, Preliminary.injective equiv equiv (addf f1). Proof using . - intros. apply fin2nat_inj. rewrite 2 addm2nat, Nat.add_comm. reflexivity. + intros f1 f2 f3 H. eapply addIf. setoid_rewrite addfC. apply H. Qed. -Lemma mod2fin_addm : ∀ m1 m2 : nat, mod2fin (m1 + m2) = addm (mod2fin m1) m2. +Lemma addm0 : ∀ f : fin u, addm f 0 = f. +Proof using . intros. rewrite addmE, Nat.add_0_r. apply mod2finK. Qed. + +Lemma add0m : ∀ j : nat, addm fin0 j = mod2fin j. +Proof using . intros. rewrite addmE, Nat.add_0_l. reflexivity. Qed. + +Lemma add0f : ∀ f : fin u, addf fin0 f = f. +Proof using . intros. rewrite addf_addm, add0m. apply mod2finK. Qed. + +Lemma addf0 : ∀ f : fin u, addf f fin0 = f. +Proof using . intros. rewrite addf_addm. apply addm0. Qed. + +(* + * + * The successor of a fin u + * + *) + +Definition sucf (f : fin u) : fin u := addf f fin1. + +Definition sucf_compat : Proper (equiv ==> equiv) sucf := _. + +Lemma sucfE : ∀ f : fin u, sucf f = addf f fin1. +Proof using . reflexivity. Qed. + +Lemma sucfEmod : ∀ f : fin u, sucf f = mod2fin (S f). Proof using . - intros. apply fin2nat_inj. rewrite addm2nat, 2 mod2fin2nat, - Nat.add_mod_idemp_l. reflexivity. apply neq_ub_0. + intros. destruct (@eq_S_l_lt_dec 0 u _) as [Hd | Hd]. apply all_eq, Hd. + unshelve erewrite sucfE, addfE, fin12nat, Nat.add_1_r. apply Hd. reflexivity. Qed. -Definition fin_max : fin ub := Fin lt_pred_ub. +Lemma sucfI : Preliminary.injective equiv equiv sucf. +Proof using . apply addIf. Qed. -Lemma addm_fin_max : addm fin_max 1 = fin0. +Lemma mod2fin_S_sucf : ∀ j : nat, mod2fin (S j) = sucf (mod2fin j). Proof using . - apply fin2nat_inj. rewrite addm2nat, Nat.add_1_r. cbn. - rewrite S_pred_ub. apply Nat.mod_same, neq_ub_0. + intros. rewrite sucfE, addfE, fin1E, addn_mod2fin_idemp_l, + addn_mod2fin_idemp_r, Nat.add_1_r. reflexivity. Qed. -Definition addf (f1 f2 : fin ub) : fin ub := addm f1 f2. +Lemma sucf_max : sucf fin_max = fin0. +Proof using . rewrite sucfEmod, fin_max2nat, S_pred_u. apply mod2fin_u. Qed. -Definition addf_compat : Proper (equiv ==> equiv ==> equiv) addf := _. +Lemma sucf_addf : ∀ f1 f2 : fin u, sucf (addf f1 f2) = addf (sucf f1) f2. +Proof using . intros. rewrite 2 sucfE. apply addfAC. Qed. -Lemma addf2nat : ∀ f1 f2 : fin ub, - addf f1 f2 = (f1 + f2) mod ub :>nat. -Proof using . intros. apply addm2nat. Qed. +Lemma sucf_addm : ∀ (f : fin u) (j : nat), sucf (addm f j) = addm (sucf f) j. +Proof using . intros. rewrite 2 addm_addf. apply sucf_addf. Qed. -Lemma addf_fin0_divide : ∀ f1 f2 : fin ub, - Nat.divide ub (f1 + f2) <-> addf f1 f2 = fin0. -Proof using . intros. apply addm_fin0_divide. Qed. +Lemma addf_sucf : ∀ f1 f2 : fin u, addf (sucf f1) f2 = addf f1 (sucf f2). +Proof using . + intros. rewrite (addfC f1), <-2 sucf_addf, addfC. reflexivity. +Qed. -Lemma addfC : ∀ f1 f2 : fin ub, addf f1 f2 = addf f2 f1. +Lemma addf_mod2fin_S : ∀ j1 j2 : nat, + addf (mod2fin (S j1)) (mod2fin j2) = addf (mod2fin j1) (mod2fin (S j2)). Proof using . - intros. apply fin2nat_inj. rewrite 2 addf2nat, Nat.add_comm. reflexivity. + intros. rewrite mod2fin_S_sucf, addf_sucf, mod2fin_S_sucf. reflexivity. Qed. -Lemma add0f : ∀ f : fin ub, addf fin0 f = f. -Proof using . apply add0m. Qed. +Lemma S_sucf : ∀ f : fin u, f < fin_max -> S f = sucf f. +Proof using . + intros * H. symmetry. rewrite sucfEmod. + apply mod2fin_small, Nat.lt_succ_lt_pred, H. +Qed. -Lemma addf0 : ∀ f : fin ub, addf f fin0 = f. -Proof using . apply addm0. Qed. +Lemma lt_sucf : ∀ f : fin u, f < fin_max -> f < sucf f. +Proof using . + intros * H. rewrite <- S_sucf. apply Nat.lt_succ_diag_r. apply H. +Qed. -Lemma addIf : ∀ f1 : fin ub, - Preliminary.injective equiv equiv (λ f2, addf f2 f1). -Proof using . intros f1 f2 f3. apply addIm. Qed. +Lemma fin1_sucf_fin0 : fin1 = sucf fin0. +Proof using . rewrite sucfEmod, fin02nat. apply fin1E. Qed. -Lemma addfI : ∀ f1 : fin ub, - Preliminary.injective equiv equiv (λ f2, addf f1 f2). +Lemma lt_sucf_le : ∀ f1 f2 : fin u, f1 < sucf f2 -> f1 <= f2. Proof using . - intros f1 f2 f3 H. eapply addIf. setoid_rewrite addfC. apply H. + intros * H1. apply lt_n_Sm_le. rewrite S_sucf. + apply H1. apply neq_max_lt_max. intros H2. subst. + eapply Nat.lt_irrefl, lt_fin0. rewrite <- sucf_max. apply H1. Qed. -Lemma addfA : ∀ f1 f2 f3 : fin ub, - addf f1 (addf f2 f3) = addf (addf f1 f2) f3. +(* + * + * The complement to fin_max of either j mod u + * (revm whose input is a nat j) or of a fin u (revf) + * + *) + +Lemma revf_subproof : ∀ f : fin u, Nat.pred u - f < u. +Proof using ltc_l_u. + intros. eapply Nat.le_lt_trans. apply Nat.le_sub_l. apply lt_pred_u. +Qed. + +Definition revf (f : fin u) : fin u := Fin (revf_subproof f). + +Definition revf_compat : Proper (equiv ==> equiv) revf := _. + +Lemma revf2nat : ∀ f : fin u, revf f = Nat.pred u - f :> nat. +Proof using . reflexivity. Qed. + +Lemma revfK : ∀ f : fin u, revf (revf f) = f. Proof using . - intros. apply fin2nat_inj. rewrite 4 addf2nat, Nat.add_mod_idemp_l, - Nat.add_mod_idemp_r, Nat.add_assoc. reflexivity. all: apply neq_ub_0. + intros. apply fin2natI. rewrite 2 revf2nat, sub_sub, minus_plus. + reflexivity. apply fin_le. Qed. -Lemma addfCA : ∀ f1 f2 f3 : fin ub, - addf f1 (addf f2 f3) = addf f2 (addf f1 f3). +Lemma revfI : Preliminary.injective equiv equiv revf. Proof using . - intros. rewrite (addfC f2 f3), addfA, (addfC (addf _ _) _). reflexivity. + intros f1 f2 H. setoid_rewrite <- revfK. rewrite H. reflexivity. Qed. -Lemma addfAC : ∀ f1 f2 f3 : fin ub, - addf (addf f1 f2) f3 = addf (addf f1 f3) f2. -Proof using . intros. apply addmAC. Qed. +Definition revm (j : nat) : fin u := revf (mod2fin j). -Lemma addfACA : ∀ f1 f2 f3 f4 : fin ub, - addf (addf f1 f2) (addf f3 f4) = addf (addf f1 f3) (addf f2 f4). -Proof using . intros. rewrite 2 addfA, (addfAC f1 f3 f2). reflexivity. Qed. +Definition revm_compat : Proper (equiv ==> equiv) revm := _. -Lemma addf_addm : ∀ f1 f2 : fin ub, addf f1 f2 = addm f1 f2. +Lemma revm_revf : ∀ j : nat, revm j = revf (mod2fin j). Proof using . reflexivity. Qed. -Lemma addm_addf : ∀ (f : fin ub) (m : nat), addm f m = addf f (mod2fin m). +Lemma revf_revm : ∀ f : fin u, revf f = revm f. +Proof using . intros. rewrite revm_revf, mod2finK. reflexivity. Qed. + +Lemma revm2nat : ∀ j : nat, revm j = Nat.pred u - (j mod u) :> nat. Proof using . - intros. rewrite addf_addm, mod2fin2nat, addm_mod. reflexivity. + intros. rewrite revm_revf, revf2nat, mod2fin2nat. reflexivity. Qed. -Definition divn (m : nat) : nat := m / ub. +Lemma revm_mod : ∀ j : nat, revm (j mod u) = revm j. +Proof using . intros. rewrite 2 revm_revf, mod2fin_mod. reflexivity. Qed. -Lemma divnE : ∀ m : nat, divn m = m / ub. -Proof using . reflexivity. Qed. +Lemma revm_mod2fin : ∀ j : nat, revm (mod2fin j) = revm j. +Proof using . intros. rewrite mod2fin2nat. apply revm_mod. Qed. -Lemma addn_divn_mod2fin : ∀ m : nat, ub * divn m + mod2fin m = m. +Lemma revmK : ∀ j : nat, revm (revm j) = mod2fin j. +Proof using . intros. rewrite 2 revm_revf, mod2finK. apply revfK. Qed. + +Lemma revm_bounded_diffI : + ∀ j1 j2 : nat, j1 <= j2 -> j2 - j1 < u -> revm j1 = revm j2 -> j1 = j2. Proof using . - intros. symmetry. rewrite mod2fin2nat. apply Nat.div_mod, neq_ub_0. + intros. eapply mod_bounded_diffI, sub_cancel_l. 4: apply mod2fin_lt. + all: cycle 3. rewrite (S_pred u) at 1 3. rewrite <- Nat.add_1_r, + 2 Nat.add_sub_swap, <-2 revm2nat, H1. reflexivity. 1,2: apply Nat.lt_le_pred, + mod2fin_lt. apply lt_l_u. apply neq_u_0. all: assumption. Qed. -Lemma divn_mod2fin : ∀ m m' : nat, - mod2fin m = mod2fin m' -> divn m = divn m' -> m = m'. +Lemma revm_betweenI : ∀ p j1 j2 : nat, + p <= j1 < p + u -> p <= j2 < p + u -> revm j1 = revm j2 -> j1 = j2. Proof using . - intros * Hm Hd. rewrite <- (addn_divn_mod2fin m), <- (addn_divn_mod2fin m'), - Hm, Hd. reflexivity. + rewrite <- bounded_diff_between. apply revm_bounded_diffI. + apply (@inj_sym _ _ (eq_setoid _) (eq_setoid _)). Qed. -Lemma rev_fin_subproof : ∀ f : fin ub, Nat.pred ub - f < ub. +Lemma revm0 : revm 0 = fin_max. Proof using . - intros. eapply Nat.le_lt_trans. apply Nat.le_sub_l. apply lt_pred_ub. + apply fin2natI. rewrite revm2nat, fin_max2nat, Nat.mod_0_l. + apply Nat.sub_0_r. apply neq_u_0. Qed. -Definition rev_fin (f : fin ub) : fin ub := Fin (rev_fin_subproof f). +Lemma revf0 : revf fin0 = fin_max. +Proof using . rewrite revf_revm, fin02nat. apply revm0. Qed. -Definition rev_fin_compat : Proper (equiv ==> equiv) rev_fin := _. +Lemma revf_le_compat : ∀ f1 f2 : fin u, f1 <= f2 -> revf f2 <= revf f1. +Proof using . intros * H. rewrite 2 revf2nat. apply Nat.sub_le_mono_l, H. Qed. -Lemma rev_fin2nat : ∀ f : fin ub, rev_fin f = Nat.pred ub - f :>nat. -Proof using . reflexivity. Qed. +Lemma revf_lt_compat : ∀ f1 f2 : fin u, f1 < f2 -> revf f2 < revf f1. +Proof using . + intros * H. rewrite 2 revf2nat. apply sub_lt_mono_l. apply fin_le. apply H. +Qed. -Lemma rev_finK : ∀ f : fin ub, rev_fin (rev_fin f) = f. +Lemma revm_le_between_compat : ∀ p j1 j2 : nat, u * p <= j1 < u * p + u + -> u * p <= j2 < u * p + u -> j1 <= j2 -> revm j2 <= revm j1. Proof using . - intros. apply fin2nat_inj. rewrite 2 rev_fin2nat, sub_sub, minus_plus. - reflexivity. apply fin_le. + intros * Hb1 Hb2 Hle. rewrite 2 revm2nat. eapply Nat.sub_le_mono_l, + mod_le_between_compat. all: eassumption. Qed. -Lemma rev_fin_inj : Preliminary.injective equiv equiv rev_fin. +Lemma revm_lt_between_compat : ∀ p j1 j2 : nat, u * p <= j1 < u * p + u + -> u * p <= j2 < u * p + u -> j1 < j2 -> revm j2 < revm j1. Proof using . - intros f1 f2 H. setoid_rewrite <- rev_finK. rewrite H. reflexivity. + intros * Hb1 Hb2 Hle. rewrite 2 revm2nat. apply sub_lt_mono_l. + apply mod2fin_le. eapply mod_lt_between_compat. all: eassumption. Qed. -Definition oppf (f : fin ub) : fin ub := addm (rev_fin f) 1. +Lemma revf_fin_max : revf fin_max = fin0. +Proof using . symmetry. apply revfI. rewrite revfK. apply revf0. Qed. -Definition oppf_compat : Proper (equiv ==> equiv) oppf := _. +Lemma revm_fin_max : revm fin_max = fin0. +Proof using . rewrite <- revf_revm. apply revf_fin_max. Qed. -Lemma oppf2nat : ∀ f : fin ub, oppf f = (ub - f) mod ub :>nat. +Lemma addf_revf : ∀ f : fin u, addf (revf f) f = fin_max. Proof using . - intros. unfold oppf. rewrite addm2nat. f_equal. erewrite rev_fin2nat, - <- Nat.add_sub_swap, Nat.add_1_r, S_pred_ub. reflexivity. apply fin_le. + intros. rewrite addfE, revf2nat, Nat.sub_add. apply mod2fin_max. apply fin_le. Qed. -Definition oppm (m : nat) : fin ub := oppf (mod2fin m). +(* + * + * The complement to fin0 of either j mod u + * (oppm whose input is a nat j) or of a fin u (oppf) + * + *) -Definition oppm_compat : Proper (equiv ==> equiv) oppm := _. +Definition oppf (f : fin u) : fin u := sucf (revf f). -Lemma oppm2nat : ∀ m : nat, oppm m = (ub - (m mod ub)) mod ub :>nat. -Proof using . - intros. unfold oppm. rewrite oppf2nat, mod2fin2nat. reflexivity. -Qed. +Definition oppf_compat : Proper (equiv ==> equiv) oppf := _. + +Lemma oppfE : ∀ f : fin u, oppf f = sucf (revf f). +Proof using . reflexivity. Qed. -Lemma oppm_mod : ∀ m : nat, oppm (m mod ub) = oppm m. +Lemma oppfEmod : ∀ f : fin u, oppf f = mod2fin (u - f). Proof using . - intros. apply fin2nat_inj. rewrite 2 oppm2nat, Nat.mod_mod. - reflexivity. apply neq_ub_0. + intros. rewrite oppfE, sucfEmod, revf2nat, + minus_Sn_m, S_pred_u. reflexivity. apply fin_le. Qed. -Lemma oppm_oppf : ∀ m : nat, oppm m = oppf (mod2fin m). +Lemma oppf2nat : ∀ f : fin u, oppf f = (u - f) mod u :> nat. +Proof using . intros. rewrite oppfEmod. apply mod2fin2nat. Qed. + +Definition oppm (j : nat) : fin u := oppf (mod2fin j). + +Definition oppm_compat : Proper (equiv ==> equiv) oppm := _. + +Lemma oppm_oppf : ∀ j : nat, oppm j = oppf (mod2fin j). Proof using . reflexivity. Qed. -Lemma oppf_oppm : ∀ f : fin ub, oppf f = oppm f. -Proof using . intros. unfold oppm. rewrite mod2finK. reflexivity. Qed. +Lemma oppf_oppm : ∀ f : fin u, oppf f = oppm f. +Proof using . intros. rewrite oppm_oppf, mod2finK. reflexivity. Qed. + +Lemma oppmE : ∀ j : nat, oppm j = sucf (revm j). +Proof using . intros. rewrite oppm_oppf, oppfE, revm_revf. reflexivity. Qed. -Lemma oppm_divide : ∀ m : nat, Nat.divide ub (m + (oppm m)). +Lemma oppmEmod : ∀ j : nat, oppm j = mod2fin (u - (j mod u)). Proof using . - intros. rewrite <- Nat.mod_divide, oppm2nat, Nat.add_mod_idemp_r, - Nat.add_comm, <- Nat.add_mod_idemp_r, Nat.sub_add. apply Nat.mod_same. - 2: apply Nat.lt_le_incl, mod2fin_lt. all: apply neq_ub_0. + intros. rewrite oppm_oppf, oppfEmod, mod2fin2nat. reflexivity. Qed. -Lemma oppf_divide : ∀ f : fin ub, Nat.divide ub (f + (oppf f)). -Proof using . intros. rewrite oppf_oppm. apply oppm_divide. Qed. +Lemma oppm2nat : ∀ j : nat, oppm j = (u - (j mod u)) mod u :> nat. +Proof using . intros. rewrite oppmEmod. apply mod2fin2nat. Qed. + +Lemma oppm_mod : ∀ j : nat, oppm (j mod u) = oppm j. +Proof using . intros. rewrite 2 oppm_oppf, mod2fin_mod. reflexivity. Qed. -Lemma addfKoppf : ∀ f : fin ub, addf (oppf f) f = fin0. +Lemma addfKoppf : ∀ f : fin u, addf (oppf f) f = fin0. Proof using . - intros. apply fin2nat_inj. rewrite addf2nat, Nat.add_comm, fin02nat. - apply Nat.mod_divide. apply neq_ub_0. apply oppf_divide. + intros. rewrite oppfE, <- sucf_addf, addf_revf. apply sucf_max. Qed. -Lemma oppfKaddf : ∀ f : fin ub, addf f (oppf f) = fin0. +Lemma addmKoppm : ∀ j : nat, addm (oppm j) j = fin0. +Proof using . intros. rewrite oppm_oppf, addm_addf. apply addfKoppf. Qed. + +Lemma oppfKaddf : ∀ f : fin u, addf f (oppf f) = fin0. Proof using . intros. rewrite addfC. apply addfKoppf. Qed. -Lemma addfOoppf : ∀ f1 f2 : fin ub, addf (oppf f1) (addf f1 f2) = f2. +Lemma addfOoppf : ∀ f1 f2 : fin u, addf (oppf f1) (addf f1 f2) = f2. Proof using . intros. rewrite addfA, addfKoppf, add0f. reflexivity. Qed. -Lemma oppfOaddf : ∀ f1 f2 : fin ub, addf (addf f2 f1) (oppf f1) = f2. +Lemma oppfOaddf : ∀ f1 f2 : fin u, addf (addf f2 f1) (oppf f1) = f2. Proof using . intros. rewrite addfC, (addfC f2 f1). apply addfOoppf. Qed. -Lemma addfOVoppf : ∀ f1 f2 : fin ub, addf f1 (addf (oppf f1) f2) = f2. +Lemma addfOVoppf : ∀ f1 f2 : fin u, addf f1 (addf (oppf f1) f2) = f2. Proof using . intros. rewrite (addfC _ f2), addfCA, oppfKaddf. apply addf0. Qed. -Lemma oppfOVaddf : ∀ f1 f2 : fin ub, addf (addf f2 (oppf f1)) f1 = f2. +Lemma oppfOVaddf : ∀ f1 f2 : fin u, addf (addf f2 (oppf f1)) f1 = f2. Proof using . intros. rewrite addfC, (addfC f2 _). apply addfOVoppf. Qed. +Lemma oppmOVaddm : ∀ (f : fin u) (j : nat), addm (addf f (oppm j)) j = f. +Proof using . intros. rewrite oppm_oppf, addm_addf. apply oppfOVaddf. Qed. + Lemma oppfI : Preliminary.injective equiv equiv oppf. -Proof using . intros f1 f2 H. eapply rev_fin_inj, addIm, H. Qed. +Proof using . intros f1 f2 H. eapply revfI, sucfI, H. Qed. -Lemma oppm_locally_inj : ∀ m1 m2 : nat, m1 <= m2 - -> m2 - m1 < ub -> oppm m1 = oppm m2 -> m1 = m2. +Lemma oppm_bounded_diffI : ∀ j1 j2 : nat, j1 <= j2 + -> j2 - j1 < u -> oppm j1 = oppm j2 -> j1 = j2. Proof using . - intros * Hleq Hsub Heq. eapply mod2fin_locally_inj, oppfI, Heq. - apply Hleq. apply Hsub. + intros * Hleq Hsub H. rewrite 2 oppmE in H. + eapply revm_bounded_diffI, sucfI. all: eassumption. Qed. -Lemma oppm_bounded_diff_inj : ∀ (p : nat) (m1 m2 : nat), - bounded_diff ub p m1 -> bounded_diff ub p m2 - -> oppm m1 = oppm m2 -> m1 = m2. +Lemma oppm_betweenI : ∀ (p : nat) (j1 j2 : nat), + p <= j1 < p + u -> p <= j2 < p + u -> oppm j1 = oppm j2 -> j1 = j2. Proof using . - rewrite <- locally_bounded_diff. apply oppm_locally_inj. + rewrite <- bounded_diff_between. apply oppm_bounded_diffI. apply (@inj_sym _ _ (eq_setoid _) (eq_setoid _)). Qed. Lemma oppf0 : oppf fin0 = fin0. -Proof using . - apply fin2nat_inj. rewrite oppf2nat, fin02nat, Nat.sub_0_r. - apply Nat.mod_same. apply neq_ub_0. -Qed. +Proof using . rewrite oppfE, revf0. apply sucf_max. Qed. Lemma oppm0 : oppm 0 = fin0. Proof using . rewrite oppm_oppf, mod2fin0. apply oppf0. Qed. -Definition subm (f : fin ub) (m : nat) : fin ub := addm f (oppm m). - -Definition subf (f1 f2 : fin ub) : fin ub := addf f1 (oppf f2). - -Definition subm_compat : Proper (equiv ==> equiv ==> equiv) subm := _. - -Definition subf_compat : Proper (equiv ==> equiv ==> equiv) subf := _. - -Lemma subf_subm : ∀ f1 f2 : fin ub, subf f1 f2 = subm f1 f2. -Proof using . intros. unfold subf, subm. rewrite oppf_oppm. reflexivity. Qed. - -Lemma subm_subf : ∀ (f : fin ub) (m : nat), subm f m = subf f (mod2fin m). +Lemma addn_oppf : ∀ f : fin u, fin0 < f -> f + oppf f = u. Proof using . - intros. unfold subf, subm. rewrite oppm_oppf, addf_addm. reflexivity. + intros * H. rewrite oppf2nat, Nat.mod_small, <- le_plus_minus. + reflexivity. apply Nat.lt_le_incl, fin_lt. apply lt_sub_u, H. Qed. -Lemma subm2nat : ∀ (f : fin ub) (m : nat), - subm f m = (f + (ub - (m mod ub))) mod ub :>nat. +Lemma divide_addn_oppf : ∀ f : fin u, Nat.divide u (f + oppf f). Proof using . - intros. unfold subm. rewrite addm2nat, oppm2nat, Nat.add_mod_idemp_r. - reflexivity. apply neq_ub_0. + intros. destruct (eq_fin0_lt_dec f) as [Hd | Hd]. subst. rewrite oppf0, + Nat.add_0_r. apply Nat.divide_0_r. rewrite addn_oppf. reflexivity. apply Hd. Qed. -Lemma subf2nat : ∀ f1 f2 : fin ub, - subf f1 f2 = (f1 + (ub - f2)) mod ub :>nat. +Lemma divide_addn_oppm : ∀ j : nat, Nat.divide u (j + oppm j). Proof using . - intros. rewrite subf_subm, subm2nat, (Nat.mod_small f2). reflexivity. - apply fin_lt. + intros. rewrite (Nat.div_mod j) at 1. rewrite <- Nat.add_assoc, + <- mod2fin2nat, oppm_oppf. apply Nat.divide_add_r. apply Nat.divide_factor_l. + apply divide_addn_oppf. apply neq_u_0. Qed. -Lemma subfE : ∀ f1 f2 : fin ub, subf f1 f2 = addf f1 (oppf f2). -Proof using . reflexivity. Qed. +(* + * + * Substracting either a nat (subm) or + * a fin u (subf) to a fin u + * + *) + +Definition subm (f : fin u) (j : nat) : fin u := addm f (oppm j). -Lemma submE : ∀ (f : fin ub) (m : nat), subm f m = addm f (oppm m). +Definition subm_compat : Proper (equiv ==> equiv ==> equiv) subm := _. + +Definition subf (f1 f2 : fin u) : fin u := addf f1 (oppf f2). + +Definition subf_compat : Proper (equiv ==> equiv ==> equiv) subf := _. + +Lemma subfE : ∀ f1 f2 : fin u, subf f1 f2 = addf f1 (oppf f2). Proof using . reflexivity. Qed. -Lemma oppf_rev_fin : ∀ f : fin ub, oppf f = addm (rev_fin f) 1. +Lemma submE : ∀ (f : fin u) (j : nat), subm f j = addf f (oppm j). Proof using . reflexivity. Qed. -Lemma rev_fin_oppf : ∀ f : fin ub, rev_fin f = subm (oppf f) 1. +Lemma subf_subm : ∀ f1 f2 : fin u, subf f1 f2 = subm f1 f2. +Proof using . intros. rewrite submE, subfE, oppf_oppm. reflexivity. Qed. + +Lemma subm_subf : ∀ (f : fin u) (j : nat), subm f j = subf f (mod2fin j). +Proof using . intros. rewrite submE, subfE, oppm_oppf. reflexivity. Qed. + +Lemma submEmod : ∀ (f : fin u) (j : nat), + subm f j = mod2fin (f + (u - (j mod u))). Proof using . - intros. rewrite submE, oppf_rev_fin, oppm_oppf, 2 addm_addf, mod2finK, - oppfOaddf. reflexivity. + intros. rewrite submE, addfE, oppm2nat, + <- mod2fin2nat. apply addn_mod2fin_idemp_r. Qed. -Lemma subm_mod : ∀ (f : fin ub) (m : nat), subm f (m mod ub) = subm f m. +Lemma subm2nat : ∀ (f : fin u) (j : nat), + subm f j = (f + (u - (j mod u))) mod u :> nat. +Proof using . intros. rewrite submEmod. apply mod2fin2nat. Qed. + +Lemma subfEmod : ∀ f1 f2 : fin u, subf f1 f2 = mod2fin (f1 + (u - f2)). +Proof using . + intros. rewrite subf_subm, submEmod, Nat.mod_small. reflexivity. apply fin_lt. +Qed. + +Lemma subf2nat : ∀ f1 f2 : fin u, + subf f1 f2 = (f1 + (u - f2)) mod u :> nat. +Proof using . intros. rewrite subfEmod. apply mod2fin2nat. Qed. + +Lemma subm_mod : ∀ (f : fin u) (j : nat), subm f (j mod u) = subm f j. Proof using . intros. rewrite 2 subm_subf, mod2fin_mod. reflexivity. Qed. -Lemma subIf : ∀ f1 : fin ub, +Lemma subm_mod2fin : ∀ (f : fin u) (j : nat), subm f (mod2fin j) = subm f j. +Proof using . intros. rewrite mod2fin2nat. apply subm_mod. Qed. + +Lemma subIf : ∀ f1 : fin u, Preliminary.injective equiv equiv (λ f2, subf f2 f1). Proof using . intros f1 f2 f3 H. eapply addIf, H. Qed. -Lemma subfI : ∀ f1 : fin ub, - Preliminary.injective equiv equiv (λ f2, subf f1 f2). +Lemma subfI : ∀ f1 : fin u, + Preliminary.injective equiv equiv (subf f1). Proof using . intros f1 f2 f3 H. eapply oppfI, addfI, H. Qed. -Lemma subIm : ∀ m : nat, - Preliminary.injective equiv equiv (λ f, subm f m). -Proof using . intros m f1 f2. rewrite 2 subm_subf. apply subIf. Qed. +Lemma subIm : ∀ j : nat, + Preliminary.injective equiv equiv (λ f, subm f j). +Proof using . intros j f1 f2. rewrite 2 subm_subf. apply subIf. Qed. -Lemma subm_locally_inj : ∀ (f : fin ub) (m1 m2 : nat), - m1 <= m2 -> m2 - m1 < ub -> subm f m1 = subm f m2 -> m1 = m2. +Lemma subm_bounded_diffI : ∀ (f : fin u) (j1 j2 : nat), + j1 <= j2 -> j2 - j1 < u -> subm f j1 = subm f j2 -> j1 = j2. Proof using . - intros * Hle Hsu. rewrite 2 subm_subf. intros Heq. eapply mod2fin_locally_inj, - subfI. all: eassumption. + intros * Hle Hsu H. rewrite 2 submE in H. + eapply oppm_bounded_diffI, addfI. all: eassumption. Qed. -Lemma subm_bounded_diff_inj : ∀ (p : nat) (f : fin ub) (m1 m2 : nat), - bounded_diff ub p m1 -> bounded_diff ub p m2 - -> subm f m1 = subm f m2 -> m1 = m2. +Lemma subm_betweenI : ∀ (p : nat) (f : fin u) (j1 j2 : nat), + p <= j1 < p + u -> p <= j2 < p + u -> subm f j1 = subm f j2 -> j1 = j2. Proof using . - intros p f. eapply (locally_bounded_diff _ - (inj_sym _ _ nat_Setoid (eq_setoid (fin ub)) (λ x, subm f x))). - intros *. apply subm_locally_inj. + intros p f. eapply (bounded_diff_between _ + (inj_sym _ _ nat_Setoid (eq_setoid (fin u)) (λ x, subm f x))). + intros *. apply subm_bounded_diffI. Qed. -Lemma sub0f : ∀ f : fin ub, subf fin0 f = oppf f. -Proof using . intros. unfold subf. apply add0f. Qed. +Lemma sub0f : ∀ f : fin u, subf fin0 f = oppf f. +Proof using . intros. rewrite subfE. apply add0f. Qed. -Lemma sub0m : ∀ m : nat, subm fin0 m = oppm m. -Proof using . intros. unfold subm. apply add0m. Qed. +Lemma sub0m : ∀ j : nat, subm fin0 j = oppm j. +Proof using . intros. rewrite submE. apply add0f. Qed. -Lemma subf0 : ∀ f : fin ub, subf f fin0 = f. -Proof using . intros. unfold subf. rewrite oppf0. apply addf0. Qed. +Lemma subf0 : ∀ f : fin u, subf f fin0 = f. +Proof using . intros. rewrite subfE, oppf0. apply addf0. Qed. -Lemma subm0 : ∀ f : fin ub, subm f 0 = f. +Lemma subm0 : ∀ f : fin u, subm f 0 = f. Proof using . intros. rewrite subm_subf, mod2fin0. apply subf0. Qed. -Lemma subfAC : ∀ f1 f2 f3 : fin ub, - subf (subf f1 f2) f3 = subf (subf f1 f3) f2. -Proof using . intros. unfold subf. rewrite addfAC. reflexivity. Qed. +Lemma subff : ∀ f : fin u, subf f f = fin0. +Proof using . intros. rewrite subfE. apply oppfKaddf. Qed. -Lemma submAC : ∀ (f : fin ub) (m1 m2 : nat), - subm (subm f m1) m2 = subm (subm f m2) m1. -Proof using . intros f m1 m2. unfold subm. rewrite addmAC. reflexivity. Qed. +Lemma submm : ∀ j : nat, subm (mod2fin j) (mod2fin j) = fin0. +Proof using . intros. rewrite subm_subf, mod2finK. apply subff. Qed. -Lemma subff : ∀ f : fin ub, subf f f = fin0. -Proof using . intros. unfold subf. apply oppfKaddf. Qed. +Lemma subfAC : ∀ f1 f2 f3 : fin u, + subf (subf f1 f2) f3 = subf (subf f1 f3) f2. +Proof using . intros. rewrite 4 subfE. apply addfAC. Qed. -Lemma submm : ∀ f : fin ub, subm f f = fin0. -Proof using . intros. rewrite subm_subf, mod2finK. apply subff. Qed. +Lemma submAC : ∀ (f : fin u) (j1 j2 : nat), + subm (subm f j1) j2 = subm (subm f j2) j1. +Proof using . intros. rewrite 4 submE. apply addmAC. Qed. -Lemma submKmod2fin : ∀ m : nat, subm (mod2fin m) m = fin0. -Proof using . intros. rewrite subm_subf. apply subff. Qed. +Lemma addf_subf : ∀ f1 f2 f3 : fin u, + addf (subf f1 f2) f3 = subf (addf f1 f3) f2. +Proof using . intros. rewrite 2 subfE. apply addfAC. Qed. -Lemma fin2natKsubm : ∀ f : fin ub, subm f f = fin0. -Proof using . intros. rewrite subm_subf, mod2finK. apply subff. Qed. +Lemma addm_subm : ∀ (f : fin u) (j1 j2 : nat), + addm (subm f j2) j1 = subm (addm f j1) j2. +Proof using . intros. rewrite 2 addm_addf, 2 subm_subf. apply addf_subf. Qed. -Lemma addfKV : ∀ f1 f2 : fin ub, subf (addf f1 f2) f1 = f2. +Lemma addfKV : ∀ f1 f2 : fin u, subf (addf f1 f2) f1 = f2. Proof using . intros. rewrite addfC, subfE. apply oppfOaddf. Qed. -Lemma addfVKV : ∀ f1 f2 : fin ub, subf (addf f2 f1) f1 = f2. +Lemma addfVKV : ∀ f1 f2 : fin u, subf (addf f2 f1) f1 = f2. Proof using . intros. rewrite addfC. apply addfKV. Qed. -Lemma subfVK : ∀ f1 f2 : fin ub, addf f1 (subf f2 f1) = f2. +Lemma addmVKV : ∀ (j : nat) (f : fin u), subm (addm f j) j = f. +Proof using . intros. rewrite subm_subf, addm_addf. apply addfVKV. Qed. + +Lemma subfVK : ∀ f1 f2 : fin u, addf f1 (subf f2 f1) = f2. Proof using . intros. eapply subIf. rewrite addfKV. reflexivity. Qed. -Lemma subfVKV : ∀ f1 f2 : fin ub, addf (subf f2 f1) f1 = f2. +Lemma subfVKV : ∀ f1 f2 : fin u, addf (subf f2 f1) f1 = f2. Proof using . intros. rewrite addfC. apply subfVK. Qed. -Lemma addmVKV : ∀ (m : nat) (f : fin ub), subm (addm f m) m = f. -Proof using . intros. rewrite subm_subf, addm_addf. apply addfVKV. Qed. +Lemma submVKV : ∀ (j : nat) (f : fin u), addm (subm f j) j = f. +Proof using . intros. rewrite addm_addf, subm_subf. apply subfVKV. Qed. -Lemma submVKV : ∀ (m : nat) (f : fin ub), addm (subm f m) m = f. -Proof using . intros. rewrite subm_subf, addm_addf. apply subfVKV. Qed. +(* + * + * The predecessor of a fin u + * + *) -Lemma subm_fin_max : subm fin0 1 = fin_max. -Proof using . rewrite <- addm_fin_max, subm_subf, addm_addf. apply addfVKV. Qed. +Definition pref (f : fin u) : fin u := subf f fin1. + +Definition pref_compat : Proper (equiv ==> equiv) pref := _. + +Lemma prefE : ∀ f : fin u, pref f = subf f fin1. +Proof using . reflexivity. Qed. -Lemma oppf_addf : ∀ f1 f2 : fin ub, oppf (addf f1 f2) = subf (oppf f1) f2. +Lemma prefI : Preliminary.injective equiv equiv pref. +Proof using . apply subIf. Qed. + +Lemma sucfK : ∀ f : fin u, pref (sucf f) = f. +Proof using . intros. rewrite prefE, sucfE. apply addfVKV. Qed. + +Lemma prefK : ∀ f : fin u, sucf (pref f) = f. +Proof using . intros. rewrite sucfE, prefE. apply subfVKV. Qed. + +Lemma prefEmod : ∀ f : fin u, pref f = mod2fin (f + Nat.pred u). Proof using . - intros. apply fin2nat_inj. rewrite subf2nat, 2 oppf2nat, addf2nat, - Nat.add_mod_idemp_l, Nat.add_sub_assoc, <- Nat.add_sub_swap, + intros. apply sucfI. symmetry. rewrite prefK, <- mod2fin_S_sucf, plus_n_Sm, + S_pred_u, <- addn_mod2fin_idemp_r, mod2fin_u, Nat.add_0_r. apply mod2finK. +Qed. + +Lemma revfE : ∀ f : fin u, revf f = pref (oppf f). +Proof using . intros. rewrite oppfE. symmetry. apply sucfK. Qed. + +Lemma pref0 : pref fin0 = fin_max. +Proof using . rewrite <- sucf_max. apply sucfK. Qed. + +Lemma pred_pref : ∀ f : fin u, fin0 < f -> Nat.pred f = pref f. +Proof using . + intros * H. symmetry. apply Nat.succ_inj. erewrite Nat.lt_succ_pred, S_sucf. + apply fin2nat_compat, prefK. rewrite <- neq_max_lt_max, <- pref0. + eapply not_iff_compat. 2: apply neq_fin0_lt_fin0. 2,3: apply H. + split. apply prefI. apply pref_compat. +Qed. + +Lemma pref_subf : ∀ f1 f2 : fin u, pref (subf f1 f2) = subf (pref f1) f2. +Proof using . intros. rewrite 2 prefE. apply subfAC. Qed. + +Lemma pref_subm : ∀ (f : fin u) (j : nat), pref (subm f j) = subm (pref f) j. +Proof using . intros. rewrite 2 subm_subf. apply pref_subf. Qed. + +Lemma sucf_subf : ∀ f1 f2 : fin u, sucf (subf f1 f2) = subf (sucf f1) f2. +Proof using . intros. rewrite 2 sucfE. apply addf_subf. Qed. + +Lemma sucf_subm : ∀ (f : fin u) (j : nat), sucf (subm f j) = subm (sucf f) j. +Proof using . intros. rewrite 2 subm_subf. apply sucf_subf. Qed. + +Lemma pref_addf : ∀ f1 f2 : fin u, pref (addf f1 f2) = addf (pref f1) f2. +Proof using . intros. rewrite 2 prefE. symmetry. apply addf_subf. Qed. + +Lemma pref_addm : ∀ (f : fin u) (j : nat), pref (addm f j) = addm (pref f) j. +Proof using . intros. rewrite 2 addm_addf. apply pref_addf. Qed. + +Lemma addf_pref : ∀ f1 f2 : fin u, addf (pref f1) f2 = addf f1 (pref f2). +Proof using . + intros. rewrite (addfC f1), <- 2 pref_addf, addfC. reflexivity. +Qed. + +Lemma lt_pref : ∀ f : fin u, fin0 < f -> pref f < f. +Proof using . + intros * H. rewrite <- pred_pref. eapply lt_pred. all: apply H. +Qed. + +Lemma fin0_pref_fin1 : fin0 = pref fin1. +Proof using. symmetry. apply sucfI. rewrite prefK. apply fin1_sucf_fin0. Qed. + +Lemma lt_pref_le : ∀ f1 f2 : fin u, pref f1 < f2 -> f1 <= f2. +Proof using . + intros * H1. apply Nat.lt_pred_le. rewrite pred_pref. + apply H1. apply neq_fin0_lt_fin0. intros H2. subst. + eapply Nat.lt_irrefl, max_lt. rewrite <- pref0. apply H1. +Qed. + +(* + * + * The complementaries (either to fin_max or fin0) + * or the output of several other functions + * (then used to prove other lemmas) + * + *) + +Lemma revf_addf : ∀ f1 f2 : fin u, revf (addf f1 f2) = subf (revf f1) f2. +Proof using . + intros. apply fin2natI. erewrite subf2nat, 2 revf2nat, addf2nat, + Nat.add_sub_assoc, <- Nat.add_sub_swap, <- (Nat.mod_small (Nat.pred _ - _)), <- Nat.sub_add_distr, <- (Nat.mod_add (_ - ((_ + _) mod _)) 1), Nat.mul_1_l, <- Nat.add_sub_swap, (Nat.mod_eq (_ + _)), sub_sub, (Nat.add_sub_swap _ (_ * _)), Nat.mul_comm, Nat.mod_add. reflexivity. - 3: apply Nat.mul_div_le. 2: apply Nat.add_le_mono. - 1,4,5,7,10: apply neq_ub_0. all: apply Nat.lt_le_incl. - 3: apply mod2fin_lt. all: apply fin_lt. + 3: apply Nat.mul_div_le. 2: apply Nat.add_le_mono. 1,4,5,7: apply neq_u_0. + 4: apply lt_sub_lt_add_l. 1,3,4,6: apply Nat.lt_le_pred. + 6: apply Nat.lt_lt_add_l. 5,7: apply Nat.lt_le_incl. 1,4-6: apply fin_lt. + 1,2: apply mod2fin_lt. all: apply lt_pred_u. Qed. -Lemma oppf_subf : ∀ f1 f2 : fin ub, oppf (subf f1 f2) = subf f2 f1. -Proof using . - intros. eapply subIf. rewrite <- oppf_addf, subfVKV, subfAC, subff, sub0f. - reflexivity. -Qed. +Lemma revf_addm : ∀ (f : fin u) (j : nat), revf (addm f j) = subm (revf f) j. +Proof using . intros. rewrite addm_addf. apply revf_addf. Qed. -Lemma oppf_subm : ∀ (f : fin ub) (m : nat), - oppf (subm f m) = subf (mod2fin m) f. -Proof using . intros. rewrite subm_subf. apply oppf_subf. Qed. +Lemma revf_sucf : ∀ f : fin u, revf (sucf f) = pref (revf f). +Proof using . intros. rewrite sucfE, prefE. apply revf_addf. Qed. + +Lemma oppf_addf : ∀ f1 f2 : fin u, oppf (addf f1 f2) = subf (oppf f1) f2. +Proof using . intros. rewrite 2 oppfE, revf_addf. apply sucf_subf. Qed. -Lemma oppfK : ∀ f : fin ub, oppf (oppf f) = f. -Proof using . intros. rewrite <- (sub0f f), oppf_subf. apply subf0. Qed. +Lemma oppf_addm : ∀ (f : fin u) (j : nat), oppf (addm f j) = subm (oppf f) j. +Proof using . intros. rewrite addm_addf, subm_subf. apply oppf_addf. Qed. -Lemma oppmK : ∀ m : nat, oppm (oppm m) = mod2fin m. -Proof using . intros. rewrite 2 oppm_oppf, mod2finK. apply oppfK. Qed. +Lemma oppf_sucf : ∀ f : fin u, oppf (sucf f) = pref (oppf f). +Proof using . intros. rewrite sucfE, prefE. apply oppf_addf. Qed. -Lemma oppf_add : ∀ f : fin ub, f ≠fin0 -> f + oppf f = ub. +Lemma revf_subf : ∀ f1 f2 : fin u, revf (subf f1 f2) = addf (revf f1) f2. Proof using . - intros f H. rewrite oppf2nat, Nat.mod_small. symmetry. apply le_plus_minus. - 2: apply Nat.sub_lt. 1,2: apply Nat.lt_le_incl, fin_lt. apply neq_0_lt. - intros Habs. apply H. apply fin2nat_inj. rewrite fin02nat, Habs. reflexivity. + intros. eapply subIf. rewrite <- revf_addf, addfVKV, subfVKV. reflexivity. Qed. -Lemma addfE : ∀ f1 f2 : fin ub, addf f1 f2 = subf f1 (oppf f2). -Proof using . intros. rewrite subfE, oppfK. reflexivity. Qed. +Lemma revf_subm : ∀ (f : fin u) (j : nat), revf (subm f j) = addm (revf f) j. +Proof using . intros. rewrite subm_subf, addm_addf. apply revf_subf. Qed. + +Lemma revf_pref : ∀ f : fin u, revf (pref f) = sucf (revf f). +Proof using . intros. rewrite sucfE, prefE. apply revf_subf. Qed. -Lemma addmE : ∀ (f : fin ub) (m : nat), addm f m = subm f (oppm m). +Lemma oppf_subf : ∀ f1 f2 : fin u, oppf (subf f2 f1) = subf f1 f2. Proof using . - intros. rewrite submE, oppmK, mod2fin2nat. symmetry. apply addm_mod. + intros. rewrite oppfE, revf_subf, sucf_addf, + <- oppfE, addfC. symmetry. apply subfE. Qed. -Lemma subfA : ∀ f1 f2 f3 : fin ub, +Lemma oppf_subm : ∀ (f : fin u) (j : nat), + oppf (subm f j) = subf (mod2fin j) f. +Proof using . intros. rewrite subm_subf. apply oppf_subf. Qed. + +Lemma oppf_pref : ∀ f : fin u, oppf (pref f) = sucf (oppf f). +Proof using . intros. rewrite 2 oppfE, revf_pref. reflexivity. Qed. + +Lemma revf_oppf : ∀ f : fin u, revf (oppf f) = pref f. +Proof using . intros. rewrite oppfE, revf_sucf, revfK. reflexivity. Qed. + +Lemma oppfK : ∀ f : fin u, oppf (oppf f) = f. +Proof using . intros. rewrite oppfE, revf_oppf. apply prefK. Qed. + +Lemma oppmK : ∀ j : nat, oppf (oppm j) = mod2fin j. +Proof using . intros. rewrite oppm_oppf. apply oppfK. Qed. + +Lemma oppf_revf : ∀ f : fin u, oppf (revf f) = sucf f. +Proof using . intros. rewrite oppfE, revfK. reflexivity. Qed. + +Lemma addf_subf_oppf : ∀ f1 f2 : fin u, addf f1 f2 = subf f1 (oppf f2). +Proof using . intros. rewrite subfE, oppfK. reflexivity. Qed. + +Lemma addm_subf_oppm : ∀ (f : fin u) (j : nat), addm f j = subf f (oppm j). +Proof using . intros. rewrite subfE, oppmK. apply addm_addf. Qed. + +Lemma subf_subf : ∀ f1 f2 f3 : fin u, subf f1 (subf f2 f3) = subf (addf f1 f3) f2. Proof using . intros. rewrite 3subfE, oppf_addf, subfE, oppfK, addfA. apply addfAC. Qed. -Lemma addf_subf : ∀ f1 f2 f3 : fin ub, - addf (subf f1 f2) f3 = subf (addf f1 f3) f2. -Proof using . intros. rewrite addfE, subfAC, <- addfE. reflexivity. Qed. - -Lemma addm_subm : ∀ (f : fin ub) (m1 m2 : nat), - addm (subm f m2) m1 = subm (addm f m1) m2. -Proof using . intros. rewrite 2 addm_addf, 2 subm_subf. apply addf_subf. Qed. +Lemma subm_subm : ∀ (f1 f2 : fin u) (j : nat), + subm f1 (subm f2 j) = subm (addm f1 j) f2. +Proof using . + intros. rewrite <-2 subf_subm, subm_subf, addm_addf. apply subf_subf. +Qed. -Lemma subf_addf : ∀ f1 f2 f3 : fin ub, +Lemma subf_addf : ∀ f1 f2 f3 : fin u, subf f1 (addf f2 f3) = subf (subf f1 f2) f3. Proof using . intros. rewrite 3 subfE, <- addfA, oppf_addf, subfE. reflexivity. Qed. -Lemma subf_addf_addf : ∀ f1 f2 f3 : fin ub, +Lemma subf_addf_addf : ∀ f1 f2 f3 : fin u, subf (addf f1 f3) (addf f2 f3) = subf f1 f2. -Proof using . intros. rewrite <- subfA, addfVKV. reflexivity. Qed. +Proof using . intros. rewrite <- subf_subf, addfVKV. reflexivity. Qed. -Lemma subf_addm_addm : ∀ (f1 f2 : fin ub) (m : nat), - subf (addm f1 m) (addm f2 m) = subf f1 f2. +Lemma subf_addm_addm : ∀ (f1 f2 : fin u) (j : nat), + subf (addm f1 j) (addm f2 j) = subf f1 f2. Proof using . intros. rewrite 2 addm_addf. apply subf_addf_addf. Qed. -Lemma subm_comp : ∀ (f : fin ub) (m2 m1 : nat), - subm (subm f m1) m2 = subm f (m1 + m2). -Proof using . - intros. rewrite 3 subm_subf, <- subf_addf, mod2fin_addm, addm_addf. - reflexivity. -Qed. +Lemma subf_addf_addfC : ∀ f1 f2 f3 : fin u, + subf (addf f1 f2) (addf f1 f3) = subf f2 f3. +Proof using . intros. rewrite 2 (addfC f1). apply subf_addf_addf. Qed. -Lemma subfCAC : ∀ f1 f2 f3 f4 : fin ub, +Lemma subfCAC : ∀ f1 f2 f3 f4 : fin u, subf (subf f1 f2) (subf f3 f4) = subf (subf f1 f3) (subf f2 f4). Proof using . intros. rewrite <- subf_addf, addfC, addf_subf, addfC, <- addf_subf, addfC, subf_addf. reflexivity. Qed. -Definition asbf (f : fin ub) : bijection (fin ub) := - cancel_bijection (λ f2 : fin ub, subf f2 f) _ (@addIf f) - (addfVKV f) (subfVKV f). +Lemma subf_sucf : ∀ f1 f2 : fin u, subf (sucf f1) f2 = subf f1 (pref f2). +Proof using . intros. rewrite sucfE, prefE. symmetry. apply subf_subf. Qed. -Definition asbm (m : nat) : bijection (fin ub) := - cancel_bijection (λ f : fin ub, subm f m) _ (@addIm m) - (addmVKV m) (submVKV m). +Lemma subf_pref : ∀ f1 f2 : fin u, subf (pref f1) f2 = subf f1 (sucf f2). +Proof using . + intros. rewrite sucfE, prefE, subfAC. symmetry. apply subf_addf. +Qed. -Lemma asbmE : ∀ m : nat, asbm m == asbf (mod2fin m). -Proof using . intros m f. apply addm_addf. Qed. +Lemma oppf1 : oppf fin1 = fin_max. +Proof using . rewrite fin1_sucf_fin0, oppf_sucf, oppf0. apply pref0. Qed. -Lemma asbmVE : ∀ m : nat, asbm mâ»Â¹ == asbf (mod2fin m)â»Â¹. -Proof using . intros m f. apply subm_subf. Qed. +Lemma oppm1 : oppm 1 = fin_max. +Proof using . rewrite oppm_oppf, <- fin1E. apply oppf1. Qed. -Lemma asbfE : ∀ f : fin ub, asbf f == asbm f. -Proof using . intros f1 f2. apply addf_addm. Qed. +Lemma oppf_max : oppf fin_max = fin1. +Proof using . rewrite <- oppf1. apply oppfK. Qed. -Lemma asbfVE : ∀ f : fin ub, asbf fâ»Â¹ == asbm fâ»Â¹. -Proof using . intros f1 f2. apply subf_subm. Qed. +(* + * + * The symmetric of f by the center c + * + *) -Lemma asbm2nat : ∀ (m : nat) (f : fin ub), asbm m f = (f + m) mod ub :>nat. -Proof using . intros. apply addm2nat. Qed. +Definition symf (c f : fin u) : fin u := addf c (subf c f). -Lemma asbf2nat : ∀ (f1 f2 : fin ub), - asbf f1 f2 = (f2 + f1) mod ub :>nat. -Proof using . intros. apply addf2nat. Qed. +Definition symf_compat : Proper (equiv ==> equiv ==> equiv) symf := _. -Lemma asbmV2nat : ∀ (m : nat) (f : fin ub), - (asbm m)â»Â¹ f = (f + (ub - (m mod ub))) mod ub :>nat. -Proof using . intros. apply subm2nat. Qed. +Lemma symfE : ∀ c f : fin u, symf c f = addf c (subf c f). +Proof using . reflexivity. Qed. -Lemma asbfV2nat : ∀ (f1 f2 : fin ub), - (asbf f1)â»Â¹ f2 = (f2 + (ub - f1)) mod ub :>nat. -Proof using . intros. apply subf2nat. Qed. +Lemma symfEmod : ∀ c f : fin u, symf c f = mod2fin (u - f + c + c). +Proof using . + intros. rewrite symfE, addfE, subfEmod, addn_mod2fin_idemp_r, Nat.add_assoc, + (Nat.add_comm _ (_ - _)), Nat.add_assoc. reflexivity. +Qed. -Lemma asbm_mod : ∀ m : nat, asbm (m mod ub) == asbm m. -Proof using . intros m f. apply addm_mod. Qed. +Lemma symf2nat : ∀ c f : fin u, symf c f = (u - f + c + c) mod u :> nat. +Proof using . intros. rewrite symfEmod. apply mod2fin2nat. Qed. -Lemma asbmV_mod : ∀ m : nat, (asbm (m mod ub))â»Â¹ == (asbm m)â»Â¹. -Proof using . intros m f. apply subm_mod. Qed. +Lemma symfI : ∀ c : fin u, Preliminary.injective equiv equiv (symf c). +Proof using . intros c f1 f2 H. eapply subfI, addfI, H. Qed. -Lemma asbmI : ∀ m : nat, Preliminary.injective equiv equiv (asbm m). -Proof using . intros. apply addIm. Qed. +Lemma symfK : ∀ c f : fin u, symf c (symf c f) = f. +Proof using . + intros. rewrite 2 symfE, subf_addf, subff, sub0f, oppf_subf. apply subfVK. +Qed. -Lemma asbmVI : ∀ m : nat, Preliminary.injective equiv equiv ((asbm m)â»Â¹). -Proof using . intros. apply subIm. Qed. +Lemma sym0f : ∀ f : fin u, symf fin0 f = oppf f. +Proof using . intros. rewrite symfE, add0f. apply sub0f. Qed. -Lemma asbfI : ∀ f : fin ub, Preliminary.injective equiv equiv (asbf f). -Proof using . intros. apply addIf. Qed. +Lemma symf0 : ∀ c : fin u, symf c fin0 = addf c c. +Proof using . intros. rewrite symfE, subf0. reflexivity. Qed. -Lemma asbfVI : ∀ f : fin ub, Preliminary.injective equiv equiv ((asbm f)â»Â¹). -Proof using . intros. apply subIf. Qed. +Lemma symff : ∀ f : fin u, symf f f = f. +Proof using . intros. rewrite symfE, subff. apply addf0. Qed. -Lemma asbm0 : asbm 0 == @id _ _. -Proof using . intros f. apply addm0. Qed. +Lemma symf_addf : ∀ c f1 f2 : fin u, symf c (addf f1 f2) = subf (symf c f1) f2. +Proof using . + intros. rewrite 2 symfE, subf_addf, addfC, addf_subf, addfC. reflexivity. +Qed. -Lemma asbmV0 : (asbm 0)â»Â¹ == (@id _ _)â»Â¹. -Proof using . intros f. apply subm0. Qed. +Lemma symf_subf : ∀ c f1 f2 : fin u, symf c (subf f1 f2) = addf (symf c f1) f2. +Proof using . + intros. rewrite 2 symfE, subf_subf, <- addf_subf, addfA. reflexivity. +Qed. -Lemma asbf0 : asbf fin0 == @id _ _. -Proof using . intros f. apply addf0. Qed. +Lemma symf_sucf : ∀ c f : fin u, symf c (sucf f) = pref (symf c f). +Proof using . intros. rewrite sucfE, prefE. apply symf_addf. Qed. -Lemma asbfV0 : (asbf fin0)â»Â¹ == (@id _ _)â»Â¹. -Proof using . intros f. apply subf0. Qed. +Lemma symf_pref : ∀ c f : fin u, symf c (pref f) = sucf (symf c f). +Proof using . intros. rewrite prefE, sucfE. apply symf_subf. Qed. -Lemma asb0f : ∀ f : fin ub, asbf f fin0 = f. -Proof using . intros. apply add0f. Qed. +(* + * + * The compatibility between various functions + * on fin u and both < and <= + * + *) -Lemma asb0m : ∀ f : fin ub, asbm f fin0 = f. -Proof using . intros. apply add0f. Qed. +Lemma addm_lt_large : ∀ (f : fin u) (j : nat), + fin0 < f -> oppf f <= j mod u -> addm f j < f. +Proof using . + intros * Hlt Hle. erewrite <- (Nat.mod_small f), <- Nat.mod_add, <- addm_mod, + addm2nat, Nat.mul_1_l. eapply mod_lt_between_compat. 1,2: rewrite Nat.mul_1_r. + 3: apply plus_lt_compat_l. 1,2: split. 4: apply plus_lt_compat_r. + 3: apply le_plus_r. 2: apply plus_lt_compat. 2,4,7: apply fin_lt. + 4: apply neq_u_0. 2,3: apply mod2fin_lt. eapply le_trans. + 2: apply plus_le_compat_l, Hle. rewrite addn_oppf. reflexivity. apply Hlt. +Qed. -Lemma asb0fV : ∀ f : fin ub, (asbf f)â»Â¹ fin0 = oppf f. -Proof using . intros. apply sub0f. Qed. +Lemma addm_le_small : + ∀ (f : fin u) (j : nat), j mod u < oppf f -> f <= addm f j. +Proof using . + intros * H. destruct (eq_fin0_lt_dec f) as [| Hd]. subst. apply fin0_le. + erewrite <- (Nat.mod_small f), <- addm_mod, addm2nat. + eapply mod_le_between_compat. 1,2: rewrite Nat.mul_0_r. apply fin_between. + 2: apply le_plus_l. 2: apply fin_lt. split. apply Nat.le_0_l. + eapply Nat.lt_le_trans. apply plus_lt_compat_l, H. + rewrite addn_oppf. reflexivity. apply Hd. +Qed. -Lemma asbffV : ∀ f : fin ub, (asbf f)â»Â¹ f = fin0. -Proof using . intros. apply subff. Qed. +Lemma addm_lt_small : + ∀ (f : fin u) (j : nat), fin0 < j mod u < oppf f -> f < addm f j. +Proof using . + intros * [H1 H2]. apply Nat.le_neq. split. apply addm_le_small, H2. + intros Habs. apply <- Nat.neq_0_lt_0. apply H1. eapply addm_betweenI. + 1,2: rewrite Nat.add_0_l. 1,2: split. 1,3: apply Nat.le_0_l. apply mod2fin_lt. + apply lt_0_u. symmetry. rewrite addm0, addm_mod. apply fin2natI, Habs. +Qed. -Lemma asbmmV : ∀ f : fin ub, (asbm f)â»Â¹ f = fin0. -Proof using . intros. apply submm. Qed. +Lemma addf_lt_large : + ∀ f1 f2 : fin u, fin0 < f1 -> oppf f1 <= f2 -> addf f1 f2 < f1. +Proof using . + intros * Hlt Hle. rewrite addf_addm. apply addm_lt_large. apply Hlt. + rewrite Nat.mod_small by apply fin_lt. apply Hle. +Qed. -Lemma asbm_fin_max : asbm 1 fin_max = fin0. -Proof using . apply addm_fin_max. Qed. +Lemma addf_le_small : + ∀ f1 f2 : fin u, f2 < oppf f1 -> f1 <= addf f1 f2. +Proof using . + intros * H. rewrite addf_addm. apply addm_le_small. + rewrite Nat.mod_small by apply fin_lt. apply H. +Qed. -Lemma asbmV_fin_max : (asbm 1)â»Â¹ fin0 = fin_max. -Proof using . apply subm_fin_max. Qed. +Lemma addf_lt_small : ∀ f1 f2 : fin u, fin0 < f2 < oppf f1 -> f1 < addf f1 f2. +Proof using . + intros * H. rewrite addf_addm. apply addm_lt_small. + rewrite Nat.mod_small. apply H. apply fin_lt. +Qed. -Lemma asbfC : ∀ f1 f2 : fin ub, asbf f1 f2 = asbf f2 f1. -Proof using . intros. apply addfC. Qed. +Lemma addm_le_compat_large : ∀ (f : fin u) (j1 j2 : nat), + oppf f <= j1 mod u <= j2 mod u -> addm f j1 <= addm f j2. +Proof using . + intros * [H1 H2]. destruct (eq_fin0_lt_dec f) as [| Hd]. + - subst. rewrite (add0m j1), (add0m j2), 2 mod2fin2nat. apply H2. + - rewrite <- (addm_mod f j1), <- (addm_mod f j2), 2 addmE. + eapply mod2fin_le_between_compat. 3: apply plus_le_compat_l, H2. + all: rewrite Nat.mul_1_r. all: split. 2,4: apply plus_lt_compat, mod2fin_lt. + 2,3: apply fin_lt. all: eapply le_trans. 2,4: apply plus_le_compat_l. + 3: eapply le_trans. 2,3: apply H1. 2: apply H2. all: rewrite addn_oppf. + 1,3: reflexivity. all: apply Hd. +Qed. -Lemma asbmC : ∀ f1 f2 : fin ub, asbm f1 f2 = asbm f2 f1. -Proof using . intros. apply addfC. Qed. +Lemma addm_le_compat_small : ∀ (f : fin u) (j1 j2 : nat), + j1 mod u <= j2 mod u < oppf f -> addm f j1 <= addm f j2. +Proof using . + intros * [H1 H2]. destruct (eq_fin0_lt_dec f) as [| Hd]. + - subst. rewrite (add0m j1), (add0m j2), 2 mod2fin2nat. apply H1. + - rewrite <- (addm_mod f j1), <- (addm_mod f j2), 2 addmE. + eapply mod2fin_le_between_compat. 3: apply plus_le_compat_l, H1. + all: rewrite Nat.mul_0_r, Nat.add_0_l. all: split. 1,3: apply Nat.le_0_l. + eapply le_lt_trans. apply plus_le_compat_l, H1. all: eapply lt_le_trans. + 1,3: apply plus_lt_compat_l, H2. all: rewrite addn_oppf. + 1,3: reflexivity. all: apply Hd. +Qed. -Lemma asbm_comp : ∀ (m2 m1 : nat), asbm m2 ∘ (asbm m1) == asbm (m1 + m2). -Proof using . intros * f. cbn-[addm]. apply addm_comp. Qed. +Lemma addf_le_compat_large : ∀ f1 f2 f3 : fin u, + oppf f1 <= f2 <= f3 -> addf f1 f2 <= addf f1 f3. +Proof using . + intros * H. rewrite 2 addf_addm. apply addm_le_compat_large. + rewrite 2 Nat.mod_small by apply fin_lt. apply H. +Qed. -Lemma asbmV_comp : ∀ (m2 m1 : nat), - (asbm m2)â»Â¹ ∘ (asbm m1)â»Â¹ == (asbm (m1 + m2))â»Â¹. -Proof using . intros * f. cbn-[subm]. apply subm_comp. Qed. +Lemma addf_le_compat_small : ∀ f1 f2 f3 : fin u, + f2 <= f3 < oppf f1 -> addf f1 f2 <= addf f1 f3. +Proof using . + intros * H. rewrite 2 addf_addm. apply addm_le_compat_small. + rewrite 2 Nat.mod_small by apply fin_lt. apply H. +Qed. -Lemma asbmAC : ∀ (m1 m2 : nat), asbm m2 ∘ (asbm m1) == asbm m1 ∘ (asbm m2). -Proof using . intros * f. cbn-[addm]. apply addmAC. Qed. +Lemma addm_lt_compat_large : ∀ (f : fin u) (j1 j2 : nat), + oppf f <= j1 mod u < j2 mod u -> addm f j1 < addm f j2. +Proof using . + intros * [H1 H2]. apply Nat.le_neq. split. apply addm_le_compat_large. + split. apply H1. apply Nat.lt_le_incl, H2. rewrite <- (addm_mod _ j1), + <- (addm_mod _ j2). intros Habs. eapply Nat.lt_irrefl. + erewrite addm_betweenI. apply H2. 3: apply fin2natI, Habs. + all: split. 1,3: apply Nat.le_0_l. all: apply mod2fin_lt. +Qed. -Lemma asbmVAC : ∀ (m1 m2 : nat), - (asbm m2)â»Â¹ ∘ (asbm m1)â»Â¹ == (asbm m1)â»Â¹ ∘ (asbm m2)â»Â¹. -Proof using . intros * f. cbn-[subm]. apply submAC. Qed. +Lemma addm_lt_compat_small : ∀ (f : fin u) (j1 j2 : nat), + j1 mod u < j2 mod u < oppf f -> addm f j1 < addm f j2. +Proof using . + intros * [H1 H2]. apply Nat.le_neq. split. apply addm_le_compat_small. + split. apply Nat.lt_le_incl, H1. apply H2. rewrite <- (addm_mod _ j1), + <- (addm_mod _ j2). intros Habs. eapply Nat.lt_irrefl. + erewrite addm_betweenI. apply H1. 3: apply fin2natI, Habs. + all: split. 1,3: apply Nat.le_0_l. all: apply mod2fin_lt. +Qed. -Lemma asbm_fin0_divide : ∀ (f : fin ub) (m : nat), - Nat.divide ub (f + m) ↔ asbm m f = fin0. -Proof using . apply addm_fin0_divide. Qed. +Lemma addf_lt_compat_large : + ∀ f1 f2 f3 : fin u, oppf f1 <= f2 < f3 -> addf f1 f2 < addf f1 f3. +Proof using . + intros * H. rewrite 2 addf_addm. apply addm_lt_compat_large. + rewrite 2 Nat.mod_small. apply H. all: apply fin_lt. +Qed. -Lemma asbm_locally_inj : ∀ m1 m2 : nat, - m1 ≤ m2 → m2 - m1 < ub → asbm m1 == asbm m2 → m1 = m2. +Lemma addf_lt_compat_small : + ∀ f1 f2 f3 : fin u, f2 < f3 < oppf f1 -> addf f1 f2 < addf f1 f3. Proof using . - intros * ?? H. apply addm_locally_inj with fin0. 3: apply H. all: assumption. + intros * H. rewrite 2 addf_addm. apply addm_lt_compat_small. + rewrite 2 Nat.mod_small. apply H. all: apply fin_lt. Qed. -Lemma asbmV_locally_inj : ∀ m1 m2 : nat, - m1 ≤ m2 → m2 - m1 < ub → (asbm m1)â»Â¹ == (asbm m2)â»Â¹ → m1 = m2. +Lemma lt_fin0_oppf : ∀ f : fin u, fin0 < f -> fin0 < oppf f. Proof using . - intros * ?? H. apply subm_locally_inj with fin0. 3: apply H. all: assumption. + intros * H1. apply neq_fin0_lt_fin0 in H1. apply neq_fin0_lt_fin0. intros H2. + apply H1. apply oppfI. rewrite oppf0. apply H2. Qed. -Lemma asbm_bounded_diff_inj : ∀ (p : nat) (m1 m2 : nat), - bounded_diff ub p m1 → bounded_diff ub p m2 - → asbm m1 == asbm m2 → m1 = m2. +Lemma oppf_le_compat : ∀ f1 f2 : fin u, fin0 < f1 <= f2 -> oppf f2 <= oppf f1. Proof using . - intros * ?? H. apply addm_bounded_diff_inj with (p:=p) (f:=fin0). - 3: apply H. all: eassumption. + intros * [H1 H2]. rewrite 2 oppfEmod. eapply mod2fin_le_between_compat. + 1,2: rewrite Nat.mul_0_r, Nat.add_0_l. 1,2: split. 1,3: apply Nat.le_0_l. + 3: apply Nat.sub_le_mono_l, H2. all: apply lt_sub_u. + eapply Nat.lt_le_trans. 1,3: apply H1. apply H2. Qed. -Lemma asbmV_bounded_diff_inj : ∀ (p : nat) (m1 m2 : nat), - bounded_diff ub p m1 → bounded_diff ub p m2 - → (asbm m1)â»Â¹ == (asbm m2)â»Â¹ → m1 = m2. +Lemma oppf_lt_compat : ∀ f1 f2 : fin u, fin0 < f1 < f2 -> oppf f2 < oppf f1. Proof using . - intros * ?? H. apply subm_bounded_diff_inj with (p:=p) (f:=fin0). - 3: apply H. all: eassumption. + intros * [H1 H2]. rewrite (S_pred_pos (oppf f1)). apply le_lt_n_Sm. + rewrite pred_pref, <- oppf_sucf. apply oppf_le_compat. rewrite <- S_sucf. + split. apply Nat.lt_0_succ. 2: eapply Nat.lt_le_trans. 1,2: apply H2. + apply le_max. all: apply lt_fin0_oppf, H1. Qed. -Lemma asbmCV : ∀ m1 m2 : nat, - asbm m1 ∘ (asbm m2)â»Â¹ == (asbm m2)â»Â¹ ∘ (asbm m1). -Proof using . intros * f. cbn-[addm subm]. apply addm_subm. Qed. +Lemma oppf_le_inj : ∀ f1 f2 : fin u, fin0 < oppf f2 <= oppf f1 -> f1 <= f2. +Proof using . + intros * H. rewrite <- (oppfK f1), <- (oppfK f2). apply oppf_le_compat, H. +Qed. -Lemma asbfCV : ∀ f1 f2 : fin ub, - asbf f1 ∘ (asbf f2)â»Â¹ == (asbf f2)â»Â¹ ∘ (asbf f1). -Proof using . intros * f3. cbn-[addf subf]. apply addf_subf. Qed. +Lemma oppf_lt_inj : ∀ f1 f2 : fin u, fin0 < oppf f2 < oppf f1 -> f1 < f2. +Proof using . + intros * H. rewrite <- (oppfK f1), <- (oppfK f2). apply oppf_lt_compat, H. +Qed. -Lemma asbf_asbm : ∀ (m : nat) (f : fin ub), - asbf (asbm m f) == asbm m ∘ (asbf f). +Lemma opmm_le_compat : + ∀ j1 j2 : nat, fin0 < j1 mod u <= j2 mod u -> oppm j2 <= oppm j1. Proof using . - intros m f1 f2. cbn-[addm addf]. rewrite 2 addm_addf. apply addfA. + intros * H. rewrite 2 oppm_oppf. apply oppf_le_compat. + rewrite 2 mod2fin2nat. apply H. Qed. -Lemma asbf_asbf : ∀ f1 f2 : fin ub, - asbf (asbf f1 f2) == asbf f1 ∘ (asbf f2). -Proof using . intros f1 f2 f3. cbn-[addf]. apply addfA. Qed. +Lemma oppm_lt_compat : + ∀ j1 j2 : nat, fin0 < j1 mod u < j2 mod u -> oppm j2 < oppm j1. +Proof using . + intros * H. rewrite 2 oppm_oppf. apply oppf_lt_compat. + rewrite 2 mod2fin2nat. apply H. +Qed. -Lemma asbf_asbmV : ∀ (m : nat) (f : fin ub), - asbf ((asbm m)â»Â¹ f) == (asbm m)â»Â¹ ∘ (asbf f). +Lemma oppm_le_inj : + ∀ j1 j2 : nat, fin0 < oppm j2 <= oppm j1 -> j1 mod u <= j2 mod u. Proof using . - intros m f1 f2. cbn-[subm addf]. rewrite 2 subm_subf, addfC, (addfC f2 f1). - apply addf_subf. + intros * H. rewrite <-2 mod2fin2nat. apply oppf_le_inj. + rewrite <-2 oppm_oppf. apply H. Qed. -Lemma asbf_asbfV : ∀ f1 f2 : fin ub, - asbf ((asbf f1)â»Â¹ f2) == (asbf f1)â»Â¹ ∘ (asbf f2). +Lemma oppm_lt_inj : + ∀ j1 j2 : nat, fin0 < oppm j2 < oppm j1 -> j1 mod u < j2 mod u. Proof using . - intros f1 f2 f3. cbn-[subf addf]. rewrite addfC, (addfC f3 f2). - apply addf_subf. + intros * H. rewrite <-2 mod2fin2nat. apply oppf_lt_inj. + rewrite <-2 oppm_oppf. apply H. Qed. -Lemma asbfV_asbm : ∀ (m : nat) (f : fin ub), - (asbf (asbm m f))â»Â¹ == (asbm m)â»Â¹ ∘ (asbf f)â»Â¹. +Lemma lt_fin0_subf : ∀ f1 f2 : fin u, f1 ≠f2 -> fin0 < subf f1 f2. Proof using . - intros m f1 f2. cbn-[addm subf]. rewrite addm_addf, subm_subf. - apply subf_addf. + intros * H1. apply neq_fin0_lt_fin0. intros H2. apply H1. + eapply subIf. rewrite (subff f2). apply H2. Qed. -Lemma asbfV_asbf : ∀ f1 f2 : fin ub, - (asbf (asbf f1 f2))â»Â¹ == (asbf f1)â»Â¹ ∘ (asbf f2)â»Â¹. -Proof using . intros f1 f2 f3. cbn-[addf subf]. apply subf_addf. Qed. +Lemma subf_lt_small : ∀ f1 f2 : fin u, fin0 < f2 <= f1 -> subf f1 f2 < f1. +Proof using . + intros * H. rewrite subfE. apply addf_lt_large. + 2: apply oppf_le_compat, H. eapply Nat.lt_le_trans. all: apply H. +Qed. -Lemma asbfV_asbmV : ∀ (m : nat) (f : fin ub), - (asbf ((asbm m)â»Â¹ f))â»Â¹ == asbm m ∘ (asbf f)â»Â¹. +Lemma subf_le_small : + ∀ f1 f2 : fin u, f2 <= f1 -> subf f1 f2 <= f1. Proof using . - intros m f1 f2. cbn-[subm addf]. rewrite addm_addf, subm_subf, addf_subf. - apply subfA. + intros * H. destruct (eq_fin0_lt_dec f1) as [| H1]. subst. + rewrite sub0f, le_fin0, <- oppf0. f_equal. apply le_fin0, H. + destruct (eq_fin0_lt_dec f2) as [| H2]. subst. rewrite subf0. + reflexivity. rewrite subfE. apply Nat.lt_le_incl, addf_lt_large. + apply H1. apply oppf_le_compat. split. apply H2. apply H. Qed. -Lemma asbfV_asbfV : ∀ f1 f2 : fin ub, - (asbf ((asbf f1)â»Â¹ f2))â»Â¹ == asbf f1 ∘ (asbf f2)â»Â¹. +Lemma subf_lt_large : ∀ f1 f2 : fin u, fin0 < f1 < f2 -> f1 < subf f1 f2. Proof using . - intros f1 f2 f3. cbn-[subf addf]. rewrite addf_subf. apply subfA. + intros * H. rewrite subfE. apply addf_lt_small. split. + 2: apply oppf_lt_compat, H. eapply lt_fin0_oppf, lt_trans. all: apply H. Qed. -(* builds the symmetric of f by the center c *) -Definition symf (c f : fin ub) : fin ub := addf c (subf c f). +Lemma subf_le_large : ∀ f1 f2 : fin u, f1 < f2 -> f1 <= subf f1 f2. +Proof using . + intros * H. destruct (eq_fin0_lt_dec f1) as [| Hd]. subst. apply fin0_le. + rewrite subfE. apply addf_le_small, oppf_lt_compat. split. apply Hd. apply H. +Qed. -Definition symf_compat : Proper (equiv ==> equiv ==> equiv) symf := _. +Lemma subm_lt_small : + ∀ (f : fin u) (j : nat), fin0 < j mod u <= f -> subm f j < f. +Proof using . + intros * H. rewrite subm_subf. apply subf_lt_small. + rewrite mod2fin2nat. apply H. +Qed. -Lemma symfE : ∀ c f : fin ub, symf c f = addf c (subf c f). -Proof using . reflexivity. Qed. +Lemma subm_le_small : + ∀ (f : fin u) (j : nat), j mod u <= f -> subm f j <= f. +Proof using . + intros * H. rewrite subm_subf. apply subf_le_small. + rewrite mod2fin2nat. apply H. +Qed. -Lemma symf2nat : ∀ c f : fin ub, symf c f = (ub - f + c + c) mod ub :>nat. +Lemma subm_lt_large : + ∀ (f : fin u) (j : nat), fin0 < f < j mod u -> f < subm f j. Proof using . - intros. rewrite symfE, addf2nat, subf2nat, Nat.add_mod_idemp_r, Nat.add_assoc, - (Nat.add_comm _ (_ - _)), Nat.add_assoc. reflexivity. apply neq_ub_0. + intros * H. rewrite subm_subf. apply subf_lt_large. + rewrite mod2fin2nat. apply H. Qed. -Lemma symfI : ∀ c : fin ub, Preliminary.injective equiv equiv (symf c). -Proof using . intros c f1 f2 H. eapply subfI, addfI, H. Qed. +Lemma subm_le_large : ∀ (f : fin u) (j : nat), f < j mod u -> f <= subm f j. +Proof using . + intros * H. rewrite subm_subf. apply subf_le_large. + rewrite mod2fin2nat. apply H. +Qed. -Lemma symfK : ∀ c f : fin ub, symf c (symf c f) = f. +Lemma subf_le_compat_small : + ∀ f1 f2 f3 : fin u, f3 <= f2 <= f1 -> subf f1 f2 <= subf f1 f3. Proof using . - intros. rewrite 2 symfE, subf_addf, subff, sub0f, oppf_subf. apply subfVK. + intros * H. destruct (eq_fin0_lt_dec f3) as [| Hd]. subst. + rewrite subf0. apply subf_le_small, H. rewrite 2 subfE. + apply addf_le_compat_large. split. all: apply oppf_le_compat. + all: split. eapply Nat.lt_le_trans. 1,4: apply Hd. all: apply H. Qed. -Definition symbf (c : fin ub) : bijection (fin ub) := - cancel_bijection (symf c) _ (@symfI c) (symfK c) (symfK c). +Lemma subf_le_compat_large : + ∀ f1 f2 f3 : fin u, f1 < f3 <= f2 -> subf f1 f2 <= subf f1 f3. +Proof using . + intros * H. destruct (eq_fin0_lt_dec f1) as [| Hd]. subst. + rewrite 2 sub0f. apply oppf_le_compat, H. rewrite 2 subfE. + apply addf_le_compat_small. split. apply oppf_le_compat. split. + eapply lt_fin0. 1,2: apply H. apply oppf_lt_compat. split. apply Hd. apply H. +Qed. -Lemma symbfV : ∀ c : fin ub, (symbf c)â»Â¹ == symbf c. -Proof using . intros c f. reflexivity. Qed. +Lemma subf_lt_compat_small : + ∀ f1 f2 f3 : fin u, f3 < f2 <= f1 -> subf f1 f2 < subf f1 f3. +Proof using . + intros * H. destruct (eq_fin0_lt_dec f3) as [| Hd]. subst. + rewrite subf0. apply subf_lt_small, H. rewrite 2 subfE. + apply addf_lt_compat_large. split. apply oppf_le_compat. split. + eapply lt_fin0. 1,2: apply H. apply oppf_lt_compat. split. apply Hd. apply H. +Qed. -Lemma symf_addf : ∀ (c f1 f2 : fin ub), - symf c (addf f1 f2) = subf (symf c f1) f2. +Lemma subf_lt_compat_large : + ∀ f1 f2 f3 : fin u, f1 < f3 < f2 -> subf f1 f2 < subf f1 f3. Proof using . - intros. rewrite 2 symfE, subf_addf, addfC, addf_subf, addfC. reflexivity. + intros * H. destruct (eq_fin0_lt_dec f1) as [| Hd]. subst. + rewrite 2 sub0f. apply oppf_lt_compat, H. rewrite 2 subfE. + apply addf_lt_compat_small. split. all: apply oppf_lt_compat. + all: split. eapply lt_fin0. 1,2,4: apply H. apply Hd. Qed. -Lemma sym0f : ∀ f : fin ub, symf fin0 f = oppf f. -Proof using . intros. rewrite symfE, add0f. apply sub0f. Qed. +Lemma subm_le_compat_small : ∀ (f : fin u) (j1 j2 : nat), + j2 mod u <= j1 mod u <= f -> subm f j1 <= subm f j2. +Proof using . + intros * H. rewrite 2 subm_subf. apply subf_le_compat_small. + rewrite 2 mod2fin2nat. apply H. +Qed. -Lemma symf0 : ∀ c : fin ub, symf c fin0 = addf c c. -Proof using . intros. rewrite symfE, subf0. reflexivity. Qed. +Lemma subm_le_compat_large : ∀ (f : fin u) (j1 j2 : nat), + f < j2 mod u <= j1 mod u -> subm f j1 <= subm f j2. +Proof using . + intros * H. rewrite 2 subm_subf. apply subf_le_compat_large. + rewrite 2 mod2fin2nat. apply H. +Qed. -Lemma symff : ∀ f : fin ub, symf f f = f. -Proof using . intros. rewrite symfE, subff. apply addf0. Qed. +Lemma subm_lt_compat_small : ∀ (f : fin u) (j1 j2 : nat), + j2 mod u < j1 mod u <= f -> subm f j1 < subm f j2. +Proof using . + intros * H. rewrite 2 subm_subf. apply subf_lt_compat_small. + rewrite 2 mod2fin2nat. apply H. +Qed. -Definition diff (f1 f2 : fin ub) : nat := subf f1 f2. +Lemma subm_lt_compat_large : ∀ (f : fin u) (j1 j2 : nat), + f < j2 mod u < j1 mod u -> subm f j1 < subm f j2. +Proof using . + intros * H. rewrite 2 subm_subf. apply subf_lt_compat_large. + rewrite 2 mod2fin2nat. apply H. +Qed. -Definition diff_compat : Proper (equiv ==> equiv ==> equiv) diff := _. +Lemma subf_lt_smallV : ∀ f1 f2 : fin u, fin0 < f2 < f1 -> subf f1 f2 < oppf f2. +Proof using . + intros * H. apply oppf_lt_inj. rewrite oppfK, oppf_subf. + split. apply H. apply subf_lt_large, H. +Qed. -Lemma diffE : ∀ f1 f2 : fin ub, diff f1 f2 = subf f1 f2. -Proof using . intros. reflexivity. Qed. +Lemma subf_le_smallV : + ∀ f1 f2 : fin u, fin0 < f2 < f1 -> subf f1 f2 <= oppf f2. +Proof using . + intros * H. apply oppf_le_inj. rewrite oppfK, oppf_subf. + split. apply H. apply subf_le_large, H. +Qed. + +Lemma subf_lt_largeV : ∀ f1 f2 : fin u, fin0 < f1 < f2 -> oppf f2 < subf f1 f2. +Proof using . + intros * H. apply oppf_lt_inj. rewrite oppfK, oppf_subf. split. + apply lt_fin0_subf, lt_gt_cases. right. 2: apply subf_lt_small. + 2: split. 3: apply Nat.lt_le_incl. all: apply H. +Qed. + +Lemma subf_le_largeV : ∀ f1 f2 : fin u, f1 < f2 -> oppf f2 <= subf f1 f2. +Proof using . + intros * H. apply oppf_le_inj. rewrite oppfK, oppf_subf. split. + apply lt_fin0_subf, lt_gt_cases. right. 2: apply subf_le_small. + 2: apply Nat.lt_le_incl. all: apply H. +Qed. + +Lemma subf_le_compat_largeV : + ∀ f1 f2 f3 : fin u, f1 <= f3 <= f2 -> subf f3 f1 <= subf f2 f1. +Proof using . + intros * [H1 H2]. destruct (le_lt_eq_dec H1) as [Hd |]. 2:{ subst. + rewrite subff. apply fin0_le. } apply oppf_le_inj. rewrite 2 oppf_subf. + split. apply lt_fin0_subf, lt_gt_cases. left. eapply Nat.lt_le_trans. + 3: apply subf_le_compat_large. 3: split. 1,3: apply Hd. all: apply H2. +Qed. + +Lemma subf_le_compat_smallV : + ∀ f1 f2 f3 : fin u, f3 <= f2 < f1 -> subf f3 f1 <= subf f2 f1. +Proof using . + intros * H. apply oppf_le_inj. rewrite 2 oppf_subf. split. + apply lt_fin0_subf, lt_gt_cases. right. 2: apply subf_le_compat_small. + 2: split. 3: apply Nat.lt_le_incl. all: apply H. +Qed. + +Lemma subf_lt_compat_largeV : + ∀ f1 f2 f3 : fin u, f1 <= f3 < f2 -> subf f3 f1 < subf f2 f1. +Proof using . + intros * [H1 H2]. destruct (le_lt_eq_dec H1) as [Hd |]. 2: subst. + 2: rewrite subff. apply oppf_lt_inj. rewrite 2 oppf_subf. split. + 1,3: apply lt_fin0_subf, lt_gt_cases. 3: apply subf_lt_compat_large. + left. 2: right. 3: split. eapply lt_trans. 1,4: apply Hd. all: apply H2. +Qed. + +Lemma subf_lt_compat_smallV : + ∀ f1 f2 f3 : fin u, f3 < f2 < f1 -> subf f3 f1 < subf f2 f1. +Proof using . + intros * H. apply oppf_lt_inj. rewrite 2 oppf_subf. split. + apply lt_fin0_subf, lt_gt_cases. right. 2: apply subf_lt_compat_small. + 2: split. 3: apply Nat.lt_le_incl. all: apply H. +Qed. + +Lemma sucf_le_compat : + ∀ f1 f2 : fin u, f1 <= f2 < fin_max -> sucf f1 <= sucf f2. +Proof using . + intros * H. rewrite 2 sucfE, (addfC f1), (addfC f2). + apply addf_le_compat_small. rewrite oppf1. apply H. +Qed. + +Lemma sucf_lt_compat : + ∀ f1 f2 : fin u, f1 < f2 < fin_max -> sucf f1 < sucf f2. +Proof using . + intros * H. rewrite 2 sucfE, (addfC f1), (addfC f2). + apply addf_lt_compat_small. rewrite oppf1. apply H. +Qed. + +Lemma pref_le_compat : + ∀ f1 f2 : fin u, fin0 < f1 <= f2 -> pref f1 <= pref f2. +Proof using . + intros * H. rewrite 2 prefE. apply subf_le_compat_largeV. split. + 2: apply H. apply lt_pref_le. rewrite <- fin0_pref_fin1. apply H. +Qed. + +Lemma pref_lt_compat : ∀ f1 f2 : fin u, fin0 < f1 < f2 -> pref f1 < pref f2. +Proof using . + intros * H. rewrite 2 prefE. apply subf_lt_compat_largeV. split. + 2: apply H. apply lt_pref_le. rewrite <- fin0_pref_fin1. apply H. +Qed. + +(* + * + * The bijections on fin u whose section adds + * and whose retraction substracts + * + *) + +Definition asbf (f : fin u) : bijection (fin u) := + cancel_bijection (λ f2 : fin u, subf f2 f) _ (@addIf f) + (addfVKV f) (subfVKV f). +Global Opaque asbf. + +Definition asbm (j : nat) : bijection (fin u) := + cancel_bijection (λ f : fin u, subm f j) _ (@addIm j) + (addmVKV j) (submVKV j). +Global Opaque asbm. + +Lemma asbfE : ∀ f1 : fin u, asbf f1 = (λ f2, addf f2 f1) :> (fin u -> fin u). +Proof using . reflexivity. Qed. + +Lemma asbfVE : + ∀ f1 : fin u, (asbf f1)â»Â¹ = (λ f2, subf f2 f1) :> (fin u -> fin u). +Proof using . reflexivity. Qed. + +Lemma asbmE : ∀ j : nat, asbm j = (λ f, addm f j) :> (fin u -> fin u). +Proof using . reflexivity. Qed. -Lemma diffI : ∀ f : fin ub, Preliminary.injective equiv equiv (diff f). -Proof using . intros f1 f2 f3 H. eapply subfI, fin2nat_inj, H. Qed. +Lemma asbmVE : ∀ j : nat, (asbm j)â»Â¹ = (λ f, subm f j) :> (fin u -> fin u). +Proof using . reflexivity. Qed. + +Lemma asbm_asbf : ∀ j : nat, asbm j == asbf (mod2fin j). +Proof using . intros j f. rewrite asbmE, asbfE. apply addm_addf. Qed. + +Lemma asbf_asbm : ∀ f1 : fin u, asbf f1 == asbm f1. +Proof using . intros f1 f2. rewrite asbmE, asbfE. apply addf_addm. Qed. + +Lemma asbmV_asbfV : ∀ j : nat, (asbm j)â»Â¹ == (asbf (mod2fin j))â»Â¹. +Proof using . intros j f. rewrite asbmVE, asbfVE. apply subm_subf. Qed. + +Lemma asbfV_asbmV : ∀ f1 : fin u, (asbf f1)â»Â¹ == (asbm f1)â»Â¹. +Proof using . intros f1 f2. rewrite asbmVE, asbfVE. apply subf_subm. Qed. + +Lemma asbm_mod : ∀ j : nat, asbm (j mod u) == asbm j. +Proof using . intros j f. rewrite 2 asbmE. apply addm_mod. Qed. + +Lemma asbmV_mod : ∀ j : nat, (asbm (j mod u))â»Â¹ == (asbm j)â»Â¹. +Proof using . intros j f. rewrite 2 asbmVE. apply subm_mod. Qed. + +Lemma asbm0 : asbm 0 == id. +Proof using . intros f. rewrite asbmE. apply addm0. Qed. + +Lemma asbf0 : asbf fin0 == id. +Proof using . intros f. rewrite asbfE. apply addf0. Qed. + +Lemma asbfAC : ∀ f1 f2 : fin u, asbf f2 ∘ (asbf f1) == asbf f1 ∘ (asbf f2). +Proof using . intros * f3. rewrite 2 compE, 2 asbfE. apply addfAC. Qed. + +Lemma asbmAC : ∀ j1 j2 : nat, asbm j2 ∘ (asbm j1) == asbm j1 ∘ (asbm j2). +Proof using . intros * f. rewrite 2 compE, 2 asbmE. apply addmAC. Qed. + +Lemma asbmVAC : + ∀ j1 j2 : nat, (asbm j2)â»Â¹ ∘ (asbm j1)â»Â¹ == (asbm j1)â»Â¹ ∘ (asbm j2)â»Â¹. +Proof using . intros * f. rewrite 2 compE, 2 asbmVE. apply submAC. Qed. + +Lemma asbfVAC : + ∀ f1 f2 : fin u, (asbf f2)â»Â¹ ∘ (asbf f1)â»Â¹ == (asbf f1)â»Â¹ ∘ (asbf f2)â»Â¹. +Proof using . intros * f3. rewrite 2 compE, 2 asbfVE. apply subfAC. Qed. -Lemma difIf : ∀ f1 : fin ub, - Preliminary.injective equiv equiv (λ f2, diff f2 f1). -Proof using . intros f1 f2 f3 H. eapply subIf, fin2nat_inj, H. Qed. +Lemma asbmCV : ∀ j1 j2 : nat, asbm j1 ∘ (asbm j2)â»Â¹ == (asbm j2)â»Â¹ ∘ (asbm j1). +Proof using . intros * f. rewrite 2 compE, asbmVE, asbmE. apply addm_subm. Qed. -Lemma diff_lt : ∀ f1 f2 : fin ub, diff f1 f2 < ub. -Proof using . intros. apply fin_lt. Qed. +Lemma asbfCV : + ∀ f1 f2 : fin u, asbf f1 ∘ (asbf f2)â»Â¹ == (asbf f2)â»Â¹ ∘ (asbf f1). +Proof using . intros * f3. rewrite 2 compE, asbfVE, asbfE. apply addf_subf. Qed. -Lemma difff : ∀ f : fin ub, diff f f = 0. -Proof using . intros. rewrite diffE, subff. apply fin02nat. Qed. +Lemma asbfA : ∀ f1 f2 : fin u, asbf (asbf f1 f2) == asbf f1 ∘ (asbf f2). +Proof using . intros * f3. rewrite compE, 3 asbfE. apply addfA. Qed. -Lemma dif0f : ∀ f : fin ub, diff fin0 f = oppf f. -Proof using . intros. rewrite diffE, sub0f. reflexivity. Qed. +Lemma asbfAV : + ∀ f1 f2 : fin u, asbf ((asbf f1)â»Â¹ f2) == (asbf f1)â»Â¹ ∘ (asbf f2). +Proof using . + intros * f3. rewrite compE, 2 asbfE, + asbfVE, addfC, (addfC f3). apply addf_subf. +Qed. -Lemma diff0 : ∀ f : fin ub, diff f fin0 = f. -Proof using . intros. rewrite diffE, subf0. reflexivity. Qed. +Lemma asbfVA : + ∀ f1 f2 : fin u, (asbf (asbf f1 f2))â»Â¹ == (asbf f1)â»Â¹ ∘ (asbf f2)â»Â¹. +Proof using . + intros * f3. rewrite compE, 3 asbfVE, asbfE. apply subf_addf. +Qed. -Lemma diff_neq : ∀ f1 f2 : fin ub, f1 ≠f2 -> 0 < diff f1 f2. +Lemma asbfVAV : + ∀ f1 f2 : fin u, (asbf ((asbf f1)â»Â¹ f2))â»Â¹ == asbf f1 ∘ (asbf f2)â»Â¹. Proof using . - intros * H. apply neq_0_lt. intros Habs. apply H. eapply diffI. - rewrite (difff f1). apply Habs. + intros * f3. rewrite compE, 3 asbfVE, asbfE, addf_subf. apply subf_subf. Qed. -Lemma symf_addm_diff : ∀ c f : fin ub, symf c f = addm c (diff c f). -Proof using . intros. rewrite diffE, addm_addf, mod2finK. reflexivity. Qed. +(* + * + * The bijection on fin u whose section is sucf + * and whose retraction is pref + * + *) + +Definition spbf : bijection (fin u) + := cancel_bijection pref _ sucfI sucfK prefK. +Global Opaque spbf. + +Lemma spbfE : spbf = sucf :> (fin u -> fin u). +Proof using . reflexivity. Qed. + +Lemma spbfVE : spbfâ»Â¹ = pref :> (fin u -> fin u). +Proof using . reflexivity. Qed. + +Lemma asbf1 : asbf fin1 == spbf. +Proof using . intros f. rewrite asbfE, spbfE. symmetry. apply sucfE. Qed. + +Lemma asbm1 : asbm 1 == spbf. +Proof using . rewrite asbm_asbf, <- fin1E. apply asbf1. Qed. + +(* + * + * The bijection on fin u whose section and retraction are revf + * + *) + +Definition rebf : bijection (fin u) + := cancel_bijection revf _ revfI revfK revfK. +Global Opaque rebf. + +Lemma rebfE : rebf = revf :> (fin u -> fin u). +Proof using . reflexivity. Qed. + +Lemma rebfVE : rebfâ»Â¹ = revf :> (fin u -> fin u). +Proof using . reflexivity. Qed. + +Lemma rebfV : rebfâ»Â¹ == rebf. +Proof using . intros f. rewrite rebfE, rebfVE. reflexivity. Qed. + +(* + * + * The bijection on fin u whose section and retraction are oppf + * + *) + +Definition opbf : bijection (fin u) + := cancel_bijection oppf _ oppfI oppfK oppfK. +Global Opaque opbf. + +Lemma opbfE : opbf = oppf :> (fin u -> fin u). +Proof using . reflexivity. Qed. + +Lemma opbfVE : opbfâ»Â¹ = oppf :> (fin u -> fin u). +Proof using . reflexivity. Qed. + +Lemma opbfV : opbfâ»Â¹ == opbf. +Proof using . intros f. rewrite opbfE, opbfVE. reflexivity. Qed. + +Lemma spbf_rebf : spbf ∘ rebf == opbf. +Proof using . + intros f. rewrite compE, spbfE, rebfE, opbfE. symmetry. apply oppfE. +Qed. + +Lemma spbfV_opbf : spbfâ»Â¹ ∘ opbf == rebf. +Proof using. + rewrite <- spbf_rebf, compose_assoc, compose_inverse_l. apply id_comp_l. +Qed. + +(* + * + * The bijection on fin u whose section and retraction are symf + * + *) + +Definition sybf (c : fin u) : bijection (fin u) := + cancel_bijection (symf c) _ (@symfI c) (symfK c) (symfK c). +Global Opaque sybf. + +Lemma sybfE : ∀ c : fin u, sybf c = symf c :> (fin u -> fin u). +Proof using . reflexivity. Qed. + +Lemma sybfVE : ∀ c : fin u, sybf câ»Â¹ = symf c :> (fin u -> fin u). +Proof using . reflexivity. Qed. + +Lemma sybfV : ∀ c : fin u, sybf câ»Â¹ == sybf c. +Proof using . intros c f. rewrite sybfE, sybfVE. reflexivity. Qed. -Lemma oppm_diff : ∀ f1 f2 : fin ub, oppm (diff f1 f2) = subf f2 f1. -Proof using . intros. rewrite diffE, oppm_oppf, mod2finK. apply oppf_subf. Qed. +Lemma sybfK : ∀ c : fin u, sybf c ∘ sybf c == id. +Proof using . intros c f. rewrite compE, sybfE. apply symfK. Qed. End mod2fin. diff --git a/Util/NumberComplements.v b/Util/NumberComplements.v index 60f83cfa6a0f51152a0121d5ec6215f0c0c2c2a9..bdc8998bac253b7d2480d02169484abee7581bd3 100644 --- a/Util/NumberComplements.v +++ b/Util/NumberComplements.v @@ -244,9 +244,88 @@ intros n Hn. replace (Nat.div2 n + Nat.div2 n) with (2 * Nat.div2 n) by lia. rewrite <- Nat.double_twice. symmetry. apply Div2.even_double. now rewrite Even.even_equiv. Qed. -Lemma le_neq_lt : forall m n : nat, n <= m -> n <> m -> n < m. +Lemma eq_0_lt_dec : ∀ n : nat, {n = 0} + {0 < n}. +Proof using . + intros. destruct (le_lt_dec n 0) as [Hd | Hd]. + left. apply Nat.le_0_r, Hd. right. apply Hd. +Qed. + +Lemma sub_sub : ∀ m n p : nat, p <= n → m - (n - p) = m + p - n. +Proof using . + intros * Hle. rewrite <- (Nat.sub_add _ _ Hle) at 2. + rewrite 2 (Nat.add_comm _ p), <- minus_plus_simpl_l_reverse. reflexivity. +Qed. + +Lemma sub_cancel_l : ∀ p m n : nat, m < p -> p - m = p - n <-> m = n. +Proof using . + intros * Hlt. split; intros H. 2:{ subst. reflexivity. } symmetry. + rewrite <- (Nat.add_sub m n), <- (Nat.add_sub n m), Nat.add_comm, + <- (Nat.add_cancel_l _ _ p), 2 Nat.add_sub_assoc, 2 (Nat.add_comm p), + <-2 Nat.add_sub_assoc, Nat.add_cancel_l at 1. exact H. + 3: rewrite Nat.add_comm. 3,4: apply Nat.le_add_r. 1,2: apply Nat.lt_le_incl. + apply lt_O_minus_lt. rewrite <- H. apply lt_minus_O_lt. all: exact Hlt. +Qed. + +Lemma add_mul_lt : ∀ a b c d : nat, a < d -> b < c -> d * b + a < d * c. +Proof using . + intros * lt_a_d lt_b_c. eapply Nat.lt_le_trans. + 2: eapply mult_le_compat_l, lt_b_c. rewrite Nat.mul_succ_r. + apply Nat.add_le_lt_mono. reflexivity. apply lt_a_d. +Qed. + +Lemma le_neq_lt : forall m n : nat, n <= m -> n ≠m -> n < m. Proof. intros n m Hle Hneq. now destruct (le_lt_or_eq _ _ Hle). Qed. +Lemma lt_sub_lt_add_l : ∀ n m p : nat, m <= n -> n < m + p -> n - m < p. +Proof using . + intros * Hle Hlt. rewrite <- Nat.le_succ_l, minus_Sn_m, Nat.le_sub_le_add_l, + Nat.le_succ_l. all: assumption. +Qed. + +Lemma lt_sub_lt_add_r : ∀ n m p : nat, p <= n -> n < m + p -> n - p < m. +Proof using . + intros * Hle Hlt. apply lt_sub_lt_add_l. + 2: rewrite Nat.add_comm. all: assumption. +Qed. + +Lemma sub_le_mono : ∀ n m p q : nat, n <= m -> p <= q -> n - q <= m - p. +Proof using . + intros * H1 H2. destruct (Nat.eq_dec (n - q) 0) as [Hd | Hd]. + rewrite Hd. apply Nat.le_0_l. apply Nat.neq_0_lt_0, lt_O_minus_lt, + Nat.lt_le_incl in Hd. apply Nat.le_add_le_sub_r. rewrite <- Nat.add_sub_swap. + apply Nat.le_sub_le_add_r, Nat.add_le_mono. all: assumption. +Qed. + +Lemma lt_sub : ∀ n m : nat, 0 < n -> 0 < m -> n - m < n. +Proof using . + intros * Hn Hm. erewrite (S_pred n) at 2. rewrite Nat.lt_succ_r, + <- Nat.sub_1_r. apply sub_le_mono. reflexivity. apply lt_le_S, Hm. apply Hn. +Qed. + +Lemma sub_lt_mono : ∀ n m p q : nat, p <= q -> p < m -> n < m -> n - q < m - p. +Proof using . + intros * H1 H2 H3. destruct (le_dec q n) as [Hd | Hd]. + - rewrite <- Nat.le_succ_l, minus_Sn_m. apply sub_le_mono. + rewrite Nat.le_succ_l. all: assumption. + - rewrite (not_le_minus_0 n). apply lt_minus_O_lt. all: assumption. +Qed. + +(* Statement improvable? Weaker precondition than m <= p? *) +Lemma sub_lt_mono_l : ∀ m n p : nat, m <= p -> n < m -> p - m < p - n. +Proof using . + intros * H1 H2. rewrite <- Nat.le_succ_l, <- Nat.sub_succ_l, + <- Nat.add_1_r, <- sub_sub, Nat.sub_1_r. apply Nat.sub_le_mono_l, + Nat.lt_le_pred, H2. eapply lt_le_S, Nat.lt_lt_0, H2. apply H1. +Qed. + +Lemma add_sub_le_sub_add : ∀ n m p : nat, n + m - p <= n - p + m. +Proof using . + intros. destruct (le_dec p n) as [Hd | Hd]. + - rewrite Nat.add_sub_swap. reflexivity. apply Hd. + - rewrite (not_le_minus_0 n), Nat.le_sub_le_add_r, Nat.add_0_l, Nat.add_comm. + apply plus_le_compat_l. apply Nat.lt_le_incl, Nat.nle_gt. all: apply Hd. +Qed. + Lemma S_pred2 : ∀ n : nat, 1 < n -> n = S (S (Nat.pred (Nat.pred n))). Proof using . intros n H. erewrite <-2 S_pred. reflexivity. 2: apply lt_pred. all: exact H. @@ -255,11 +334,11 @@ Qed. Lemma S_pred3 : ∀ n : nat, 2 < n -> n = S (S (S (Nat.pred (Nat.pred (Nat.pred n))))). Proof using . - intros n H. erewrite <-3 S_pred. reflexivity. 2,3: apply lt_pred. - 3: apply lt_pred. all: exact H. + intros n H. erewrite <-3 S_pred. reflexivity. + 2,3: apply lt_pred. 3: apply lt_pred. all: exact H. Qed. -Lemma mod_locally_inj : ∀ p m n : nat, p ≠0 -> m <= n +Lemma mod_bounded_diffI : ∀ p m n : nat, p ≠0 -> m <= n -> n - m < p -> m mod p = n mod p -> m = n. Proof using . intros * Hne Hle Hsu Heq. rewrite (Nat.div_mod m p), (Nat.div_mod n p), Heq. @@ -271,25 +350,54 @@ Proof using . rewrite Heq at 1. rewrite <-2 Nat.div_mod by exact Hne. exact Hsu. Qed. -Definition bounded_diff (n p m : nat) : Prop := p <= m < p + n. +Lemma between_1 : ∀ m : nat, m <= m < m +1. +Proof using . + intros. split. reflexivity. rewrite Nat.add_1_r. apply Nat.lt_succ_diag_r. +Qed. -Lemma bounded_diff_neq_0 : ∀ n p m : nat, bounded_diff n p m -> n ≠0. +Lemma between_neq_0 : ∀ n p m : nat, p <= m < p + n -> n ≠0. Proof using . intros * [H1 H2] Heq. subst. eapply (Nat.lt_irrefl p), Nat.le_lt_trans. apply H1. rewrite <- (Nat.add_0_r p). apply H2. Qed. -Lemma locally_bounded_diff : ∀ P : nat -> nat -> Prop, Symmetric P +Lemma between_addn : ∀ n1 n2 p1 p2 m1 m2 : nat, p1 <= m1 < p1 + n1 + -> p2 <= m2 < p2 + n2 -> p1 + p2 <= m1 + m2 < p1 + p2 + Nat.pred (n1 + n2). +Proof using . + intros * H1 H2. split. apply plus_le_compat. apply H1. apply H2. + rewrite <- Nat.add_pred_r, Nat.add_assoc, (Nat.add_shuffle0 p1), + <- (Nat.add_assoc (_ + _)). apply plus_lt_le_compat. apply H1. + rewrite Nat.add_pred_r. apply Nat.lt_le_pred, H2. + all: eapply between_neq_0. all: eassumption. +Qed. + +Lemma between_subn : ∀ n1 n2 p1 p2 m1 m2 : nat, p1 <= m1 < p1 + n1 + -> p2 <= m2 < p2 + n2 -> p1 - Nat.pred (p2 + n2) + <= m1 - m2 < p1 - Nat.pred (p2 + n2) + Nat.pred (n1 + n2). +Proof using . + intros * H1 H2. split. apply sub_le_mono. apply H1. apply Nat.lt_le_pred, H2. + rewrite <-2 Nat.add_pred_r, Nat.sub_add_distr, (Nat.add_comm n1), + Nat.add_assoc. eapply Nat.lt_le_trans. 2: apply plus_le_compat_r, + Nat.sub_add_le. 2,3: eapply between_neq_0, H2. + destruct (le_dec p2 p1) as [Hd | Hd]. + - rewrite <- Nat.add_sub_swap. apply sub_lt_mono. apply H2. 2: apply H1. + rewrite <- (Nat.add_0_r p2). apply plus_le_lt_compat. 1,3: apply Hd. + eapply Nat.neq_0_lt_0, between_neq_0, H1. + - rewrite (not_le_minus_0 p1), Nat.add_0_l. eapply Nat.le_lt_trans. + eapply Nat.sub_le_mono_l. eapply Nat.le_trans. 2: apply H2. + apply Nat.lt_le_incl, Nat.nle_gt. 1,3: apply Hd. apply lt_sub_lt_add_l. + all: apply H1. +Qed. + +Lemma bounded_diff_between : ∀ P : nat -> nat -> Prop, Symmetric P -> ∀ n : nat, (∀ m1 m2 : nat, m1 <= m2 -> m2 - m1 < n -> P m1 m2) - <-> ∀ (p m1 m2: nat), bounded_diff n p m1 -> bounded_diff n p m2 - -> P m1 m2. + <-> ∀ (p m1 m2: nat), p <= m1 < p + n -> p <= m2 < p + n -> P m1 m2. Proof using . intros * HS *. split. - intros H * [H11 H12] [H21 H22]. destruct (le_ge_dec m1 m2) as [Hd | Hd]. 2: symmetry. all: apply H. 1,3: exact Hd. all: eapply (Nat.le_lt_trans _ (_ - p)). 1,3: apply Nat.sub_le_mono_l. - 3,4: rewrite <- Nat.le_succ_l, minus_Sn_m, Nat.le_sub_le_add_l, - Nat.le_succ_l. all: assumption. + 3,4: apply lt_sub_lt_add_l. all: assumption. - intros H m1 m2 Hle Hlt. apply (H m1). all: split. reflexivity. 1,3: apply Nat.lt_sub_lt_add_l. rewrite Nat.sub_diag. eapply Nat.le_lt_trans. apply Nat.le_0_l. all: eassumption. @@ -301,33 +409,127 @@ Proof using . intros * a1 a2 H1 H2. symmetry. apply H1. symmetry. exact H2. Qed. -Lemma mod_bounded_diff_inj : ∀ p n m1 m2 : nat, bounded_diff n p m1 - -> bounded_diff n p m2 -> m1 mod n = m2 mod n -> m1 = m2. +Lemma mod_betweenI : ∀ p n m1 m2 : nat, p <= m1 < p + n -> p <= m2 < p + n + -> m1 mod n = m2 mod n -> m1 = m2. Proof using . - intros * H. pose proof (bounded_diff_neq_0 _ _ _ H) as Hn. revert m1 m2 H. - eapply (locally_bounded_diff _ (inj_sym _ _ (eq_setoid nat) (eq_setoid nat) - (λ x, x mod n))). intros * H1 H2 H3. eapply mod_locally_inj. all: eassumption. + intros * H. pose proof (between_neq_0 _ _ _ H) as Hn. revert m1 m2 H. + eapply (bounded_diff_between _ (inj_sym _ _ (eq_setoid nat) (eq_setoid nat) + (λ x, x mod n))). intros * H1 H2 H3. eapply mod_bounded_diffI. all: eassumption. Qed. -Lemma sub_sub : ∀ m n p : nat, p <= n → m - (n - p) = m + p - n. +Lemma between_muln : ∀ p n m : nat, n * p <= m < n * p + n -> m / n = p. Proof using . - intros * Hle. rewrite <- (Nat.sub_add _ _ Hle) at 2. - rewrite 2 (Nat.add_comm _ p), <- minus_plus_simpl_l_reverse. reflexivity. + intros * [H1 H2]. symmetry. apply Nat.div_unique with (m - n * p). + 2: apply le_plus_minus. eapply Nat.add_lt_mono_r. rewrite Nat.sub_add, + Nat.add_comm. apply H2. all: apply H1. Qed. -Lemma sub_cancel_l : ∀ p m n : nat, m < p -> p - m = p - n <-> m = n. +Lemma mod_le_between_compat : ∀ p n m1 m2 : nat, n * p <= m1 < n * p + n + -> n * p <= m2 < n * p + n -> m1 <= m2 -> m1 mod n <= m2 mod n. Proof using . - intros * Hlt. split; intros H. 2:{ subst. reflexivity. } symmetry. - rewrite <- (Nat.add_sub m n), <- (Nat.add_sub n m), Nat.add_comm, - <- (Nat.add_cancel_l _ _ p), 2 Nat.add_sub_assoc, 2 (Nat.add_comm p), - <-2 Nat.add_sub_assoc, Nat.add_cancel_l at 1. exact H. - 3: rewrite Nat.add_comm. 3,4: apply Nat.le_add_r. 1,2: apply Nat.lt_le_incl. - apply lt_O_minus_lt. rewrite <- H. apply lt_minus_O_lt. all: exact Hlt. + intros * H1 H2 Hle. erewrite 2 Nat.mod_eq, 2 between_muln, + Nat.add_le_mono_r, 2 Nat.sub_add. 6,7: eapply between_neq_0. + 1,4-7: eassumption. apply H2. apply H1. +Qed. + +Lemma mod_lt_between_compat : ∀ p n m1 m2 : nat, n * p <= m1 < n * p + n + -> n * p <= m2 < n * p + n -> m1 < m2 -> m1 mod n < m2 mod n. +Proof using . + intros * H1 H2 Hlt. erewrite 2 Nat.mod_eq, 2 between_muln, + Nat.add_lt_mono_r, 2 Nat.sub_add. 6,7: eapply between_neq_0. + 1,4-7: eassumption. apply H2. apply H1. +Qed. + +Lemma divide_neq : ∀ a b : nat, b ≠0 -> Nat.divide a b -> a ≠0. +Proof using . intros * Hn Hd H. subst. apply Hn, Nat.divide_0_l, Hd. Qed. + +Lemma divide_div_neq : ∀ a b : nat, b ≠0 -> Nat.divide a b -> b / a ≠0. +Proof using . + intros * Hn Hd H. apply Hn. unshelve erewrite (Nat.div_mod b a), + (proj2 (Nat.mod_divide b a _)), Nat.add_0_r, H. 2: apply Nat.mul_0_r. + 1,3: eapply divide_neq. all: eassumption. +Qed. + +Lemma addn_muln_divn : ∀ n q1 q2 r1 r2 : nat, n ≠0 -> r1 < n -> r2 < n + -> n * q1 + r1 = n * q2 + r2 -> q1 = q2 ∧ r1 = r2. +Proof using . + intros * Hn H1 H2 H. apply Nat.div_unique in H as H'. rewrite Nat.mul_comm, + Nat.div_add_l, Nat.div_small, Nat.add_0_r in H'. subst. split. reflexivity. + eapply Nat.add_cancel_l. all: eassumption. +Qed. + +Lemma divide_mul : ∀ a b : nat, b ≠0 -> Nat.divide a b -> b = a * (b / a). +Proof using . + intros * Hn Hd. erewrite (Nat.div_mod b a) at 1. + unshelve erewrite (proj2 (Nat.mod_divide b a _)). 2: apply Nat.add_0_r. + 1,3: eapply divide_neq. all: eassumption. +Qed. + +Lemma divide_mod : + ∀ a b c : nat, b ≠0 -> Nat.divide a b -> (c mod b) mod a = c mod a. +Proof using . + intros * Hn Hd. rewrite (divide_mul a b), Nat.mod_mul_r, Nat.mul_comm, + Nat.mod_add. apply Nat.mod_mod. 4: apply divide_div_neq. + 1-3: eapply divide_neq. all: eassumption. +Qed. + +Lemma mul_add : ∀ a b : nat, + b ≠0 -> a * b = a + Nat.pred b * a. +Proof using . + intros * Hb. destruct (Nat.eq_dec a 0) as [?|Ha]. + - subst. rewrite Nat.mul_0_r. apply Nat.mul_0_l. + - rewrite <- (Nat.mul_1_l a) at 2. rewrite <- Nat.mul_add_distr_r, + Nat.add_1_l, Nat.succ_pred. apply Nat.mul_comm. apply Hb. Qed. -Lemma add_mul_lt : ∀ a b c d : nat, a < d -> b < c -> a + b * d < c * d. +Lemma pred_mul_add : ∀ a b : nat, + b ≠0 -> Nat.pred (a * b) = Nat.pred a + Nat.pred b * a. Proof using . - intros * a_lt_d b_lt_c. apply (Mult.mult_le_compat_r _ _ d) in b_lt_c. lia. + intros * Hb. destruct (Nat.eq_dec a 0) as [?|Ha]. + - subst. rewrite Nat.mul_0_r. apply Nat.mul_0_l. + - rewrite mul_add by apply Hb. symmetry. apply Nat.add_pred_l, Ha. +Qed. + +Lemma lt_pred_mul : ∀ a b c d : nat, + a < b -> c < d -> Nat.pred ((S a) * (S c)) < b * d. +Proof using . + intros * Hab Hcd. erewrite <- Nat.le_succ_l, Nat.lt_succ_pred. + - apply Nat.mul_le_mono. apply Hab. apply Hcd. + - apply Nat.mul_pos_pos. all: apply Nat.lt_0_succ. +Qed. + +Fact mul_le_lt_compat : ∀ (n m p q : nat), + 0 < n -> n <= m -> p < q -> n * p < m * q. +Proof using. + intros. + apply Nat.le_lt_trans with (m * p). + - apply Nat.mul_le_mono_r; assumption. + - apply Nat.mul_lt_mono_pos_l; try assumption. + apply Nat.lt_le_trans with n; assumption. +Qed. + +Corollary mul_lt_compat : ∀ (m p q : nat), 0 < m -> p < q -> p < m * q. +Proof using. + intros. + rewrite <- (Nat.mul_1_l p). + apply (@mul_le_lt_compat 1); try assumption. + apply Nat.lt_0_1. +Qed. + +Lemma neq_lt : ∀ a b : nat, a < b -> b ≠a. +Proof using . + intros * Hlt Heq. apply (lt_irrefl a). rewrite <- Heq at 2. exact Hlt. +Qed. + +Lemma lt_pred : ∀ a b : nat, a < b -> Nat.pred b < b. +Proof using . + intros * H. eapply lt_pred_n_n, Nat.le_lt_trans. apply Nat.le_0_l. apply H. +Qed. + +Lemma neq_pred : ∀ a b c : nat, a < b -> b < c -> Nat.pred c ≠a. +Proof using . + intros * Hab Hbc Habs. eapply (lt_not_le (Nat.pred _)). + rewrite Habs. exact Hab. apply Nat.lt_le_pred, Hbc. Qed. Open Scope Z. @@ -374,3 +576,115 @@ Lemma Zmin_bounds : forall n m, n < m -> Z.min n (m - n) <= m / 2. Proof. intros. apply Z.min_case_strong; intro; apply Zdiv.Zdiv_le_lower_bound; lia. Qed. Close Scope Z. + +Class ltc (l u : nat) := lt_l_u : l < u. +Infix "<c" := ltc (at level 70, no associativity). + +Section ltc_l_u. + +Context {l u : nat} {ltc_l_u : l <c u}. + +Lemma neq_u_l : u ≠l. +Proof using ltc_l_u. apply neq_lt, lt_l_u. Qed. + +Lemma le_l_pred_u : l <= Nat.pred u. +Proof using ltc_l_u. apply Nat.lt_le_pred, lt_l_u. Qed. + +Lemma lt_pred_u : Nat.pred u < u. +Proof using ltc_l_u. eapply lt_pred, lt_l_u. Qed. + +Lemma S_pred_u : S (Nat.pred u) = u. +Proof using ltc_l_u. eapply Nat.lt_succ_pred, lt_l_u. Qed. + +End ltc_l_u. + +Section ltc_s_u. + +Context {l u : nat} {ltc_l_u : l <c u}. + +Lemma lt_s_u : ∀ s : nat, s <= l -> s < u. +Proof using ltc_l_u. + intros * H. eapply Nat.le_lt_trans. exact H. exact lt_l_u. +Qed. + +Lemma lt_s_pred_u : ∀ s : nat, s < l -> s < Nat.pred u. +Proof using ltc_l_u. + intros * H. eapply Nat.lt_le_trans. exact H. exact le_l_pred_u. +Qed. + +Lemma le_s_pred_u : ∀ s : nat, s <= l -> s <= Nat.pred u. +Proof using ltc_l_u. + intros * H. eapply Nat.le_trans. exact H. exact le_l_pred_u. +Qed. + +Lemma neq_pred_u_s : ∀ s : nat, s < l -> Nat.pred u ≠s. +Proof using ltc_l_u. intros * H. eapply neq_pred. apply H. apply lt_l_u. Qed. + +Lemma lt_S_l_u : u ≠S l -> S l < u. +Proof using ltc_l_u. + intros H. apply le_neq_lt. apply ltc_l_u. apply not_eq_sym, H. +Qed. + +Lemma eq_S_l_lt_dec : {u = S l} + {S l < u}. +Proof using ltc_l_u. + destruct (Nat.eq_dec u (S l)) as [Hd | Hd]. + left. apply Hd. right. apply lt_S_l_u, Hd. +Qed. + +Global Instance lt_0_u : 0 <c u. +Proof using ltc_l_u. apply lt_s_u, Nat.le_0_l. Qed. + +Lemma neq_u_0 : u ≠0. +Proof using ltc_l_u. apply neq_u_l. Qed. + +Lemma le_0_pred_u : 0 <= Nat.pred u. +Proof using ltc_l_u. apply le_l_pred_u. Qed. + +Lemma lt_l_g : ∀ g : nat, u <= g -> l < g. +Proof using ltc_l_u. + intros * H. eapply Nat.lt_le_trans. exact lt_l_u. exact H. +Qed. + +Lemma lt_sub_u : ∀ s : nat, 0 < s -> u - s < u. +Proof using ltc_l_u. intros * H. apply lt_sub, H. apply lt_0_u. Qed. + +End ltc_s_u. + +Section lt_pred_mul_ul_ur. + +Context {ll lr ul ur : nat} {ltc_ll_ul : ll <c ul} {ltc_lr_ur : lr <c ur}. + +Lemma lt_pred_mul_ul_ur : Nat.pred ((S ll) * (S lr)) < ul * ur. +Proof using ltc_ll_ul ltc_lr_ur. + apply lt_pred_mul. all: apply lt_l_u. +Qed. + +Lemma lt_ll_mul_ul_ur : ll < ul * ur. +Proof using ltc_ll_ul ltc_lr_ur. + rewrite Nat.mul_comm. apply mul_lt_compat. apply lt_0_u. apply lt_l_u. +Qed. + +Lemma lt_lr_mul_ul_ur : lr < ul * ur. +Proof using ltc_ll_ul ltc_lr_ur. + apply mul_lt_compat. apply lt_0_u. apply lt_l_u. +Qed. + +End lt_pred_mul_ul_ur. + +Section lt_pred_mul_ul_ur. + +Context {ll lr ul ur : nat} {ltc_ll_ul : ll <c ul} {ltc_lr_ur : lr <c ur}. + +Lemma lt_pred_ul_mul_ul_ur : Nat.pred ul < ul * ur. +Proof using ltc_ll_ul ltc_lr_ur. + eapply @lt_s_u. eapply @lt_pred_mul_ul_ur. apply lt_pred_u. + apply lt_0_u. rewrite Nat.mul_1_r. reflexivity. +Qed. + +Lemma lt_pred_ur_mul_ul_ur : Nat.pred ur < ul * ur. +Proof using ltc_ll_ul ltc_lr_ur. + eapply @lt_s_u. eapply @lt_pred_mul_ul_ur. apply lt_0_u. + apply lt_pred_u. rewrite Nat.mul_1_l. reflexivity. +Qed. + +End lt_pred_mul_ul_ur. diff --git a/Util/Preliminary.v b/Util/Preliminary.v index cefff3adf7d16a83f8c013511f0b0fbbdf7bbb6a..bf9d436e3bffe3a917b2205f7eb5d6c6ada1ba6a 100644 --- a/Util/Preliminary.v +++ b/Util/Preliminary.v @@ -22,7 +22,6 @@ Require Import Utf8 Relations Morphisms SetoidClass. Set Implicit Arguments. - Ltac autoclass := eauto with typeclass_instances. Ltac inv H := inversion H; subst; clear H. Hint Extern 1 (equiv ?x ?x) => reflexivity : core. @@ -104,9 +103,82 @@ Class Inverse T `{Setoid T} := { Notation "bij â»Â¹" := (inverse bij) (at level 39). Arguments Inverse T {_}. +(* pick_spec can be useful when you want to test whether some 'oT : option T' is + 'Some t' or 'None'. Destructing 'pick_spec P oT' gives you directly 'P t' + in the first case and '∀ t : T, ¬ (P t)' in the other one *) Variant pick_spec (T : Type) (P : T -> Prop) : option T -> Type := Pick : ∀ x : T, P x → pick_spec P (Some x) | Nopick : (∀ x : T, ¬ (P x)) → pick_spec P None. Arguments pick_spec [T] _ _. Arguments Pick [T P x] _. Arguments Nopick [T P] _. + +Lemma injective_compat_iff : + ∀ {A B : Type} {eqA : relation A} {eqB : relation B} (f : A -> B) + {compat : Proper (eqA ==> eqB) f}, injective eqA eqB f + -> ∀ a1 a2 : A, eqB (f a1) (f a2) <-> eqA a1 a2. +Proof using . + intros * Hcompat Hinj *. split. apply Hinj. apply Hcompat. +Qed. + +Lemma injective_eq_iff : ∀ {A B : Type} (f : A -> B), + injective eq eq f -> ∀ a1 a2 : A, f a1 = f a2 <-> a1 = a2. +Proof using . intros *. apply injective_compat_iff. autoclass. Qed. + +Lemma eq_sym_iff : ∀ (A : Type) (x y : A), x = y <-> y = x. +Proof using . intros. split. all: apply eq_sym. Qed. + +Lemma decidable_not_and_iff : + ∀ P Q : Prop, Decidable.decidable P -> ¬ (P ∧ Q) <-> ¬ P ∨ ¬ Q. +Proof using . + intros * Hd. split. apply Decidable.not_and, Hd. intros Ho Ha. + destruct Ho as [Ho | Ho]. all: apply Ho, Ha. +Qed. + +Lemma decidable_not_not_iff : ∀ P : Prop, Decidable.decidable P -> ¬ ¬ P <-> P. +Proof using . + intros * Hd. split. apply Decidable.not_not, Hd. intros H Hn. apply Hn, H. +Qed. + +Lemma sumbool_iff_compat : + ∀ P Q R S: Prop, P <-> R -> Q <-> S -> {P} + {Q} -> {R} + {S}. +Proof using . + intros * H1 H2 [Hd | Hd]. left. apply H1, Hd. right. apply H2, Hd. +Qed. + +Lemma sumbool_not_iff_compat : + ∀ P Q : Prop, P <-> Q -> {P} + {¬ P} -> {Q} + {¬ Q}. +Proof using . + intros * H Hd. eapply sumbool_iff_compat. + 2: apply not_iff_compat. 1,2: apply H. apply Hd. +Qed. + +Lemma sumbool_and_compat : + ∀ P Q R S: Prop, {P} + {Q} -> {R} + {S} -> {P ∧ R} + {Q ∨ S}. +Proof using . + intros * [H1 | H1] [H2 | H2]. left. split. apply H1. apply H2. + all: right. right. apply H2. all: left. all: apply H1. +Qed. + +Lemma sumbool_decidable : ∀ P : Prop, {P} + {¬ P} -> Decidable.decidable P. +Proof using . intros P [H | H]. left. apply H. right. apply H. Qed. + +Lemma pair_eqE : + ∀ {A B : Type} (p1 p2 : A * B), fst p1 = fst p2 ∧ snd p1 = snd p2 <-> p1 = p2. +Proof using . + intros. rewrite (surjective_pairing p1), (surjective_pairing p2). + symmetry. apply pair_equal_spec. +Qed. + +Lemma pair_dec : ∀ {A B : Type}, + (∀ a1 a2 : A, {a1 = a2} + {a1 ≠a2}) -> (∀ b1 b2 : B, {b1 = b2} + {b1 ≠b2}) + -> ∀ p1 p2 : A * B, {p1 = p2} + {p1 ≠p2}. +Proof using . + intros * Ha Hb *. eapply sumbool_not_iff_compat. apply pair_eqE. + eapply sumbool_iff_compat. reflexivity. symmetry. apply decidable_not_and_iff. + 2: apply sumbool_and_compat. 2: apply Ha. 2: apply Hb. + apply sumbool_decidable, Ha. +Qed. + +Lemma and_cancel : ∀ P : Prop, P -> P ∧ P. +Proof using . intros * H. split. all: apply H. Qed. diff --git a/Util/Ratio.v b/Util/Ratio.v index 1e6166937587b51f2f838ba022fa09ff6592c834..07a280e14791f98efa3359f5a849b36606430b65 100644 --- a/Util/Ratio.v +++ b/Util/Ratio.v @@ -37,6 +37,9 @@ Instance ratio_EqDec : EqDec ratio_Setoid := @sig_EqDec _ _ _ _. Definition proj_ratio : ratio -> R := @proj1_sig _ _. +Lemma proj_ratio_inj : injective equiv equiv proj_ratio. +Proof using . apply proj1_sig_inj. Qed. + Instance proj_ratio_compat : Proper (equiv ==> equiv) proj_ratio. Proof. intros ? ? Heq. apply Heq. Qed. @@ -81,7 +84,7 @@ Proof. intros [] [] ? [] [] ?. unfold add_ratio. simpl in *. subst. destruct_mat (** A strict ratio is a [ratio] that is neither [0] nor [1]. *) Definition strict_ratio := {x : R | 0 < x < 1}%R. -Instance strict_ratio_Setoid : Setoid ratio := sig_Setoid _. +Instance strict_ratio_Setoid : Setoid strict_ratio := sig_Setoid _. Instance strict_ratio_EqDec : EqDec strict_ratio_Setoid := @sig_EqDec _ _ _ _. Definition proj_strict_ratio (x : strict_ratio) : ratio := @@ -91,6 +94,9 @@ Definition proj_strict_ratio (x : strict_ratio) : ratio := Instance proj_strict_ratio_compat : Proper (equiv ==> equiv) proj_strict_ratio. Proof. intros [] [] Heq. apply Heq. Qed. +Lemma proj_strict_ratio_inj : injective equiv equiv proj_strict_ratio. +Proof using . intros [][] H. cbn in H. subst. reflexivity. Qed. + Coercion proj_strict_ratio : strict_ratio >-> ratio. Lemma strict_ratio_bounds : forall r : strict_ratio, (0 < r < 1)%R. diff --git a/Util/SetoidDefs.v b/Util/SetoidDefs.v index 3963647279e6ea4d934ab31c1a8138ee10da02fd..e95477fd5dd04e2bae5018367283a38d9d38858c 100644 --- a/Util/SetoidDefs.v +++ b/Util/SetoidDefs.v @@ -19,7 +19,7 @@ (**************************************************************************) -Require Import Rbase RelationPairs SetoidDec. +Require Import Utf8 Rbase RelationPairs SetoidDec. Require Import Pactole.Util.Preliminary. Set Implicit Arguments. @@ -42,6 +42,7 @@ Instance bool_EqDec : EqDec bool_Setoid := Bool.bool_dec. Instance unit_EqDec : EqDec unit_Setoid := fun x y => match x, y with tt, tt => left eq_refl end. Notation "x == y" := (equiv x y). +Notation "x == y :> A" := (@equiv A _ x y) (at level 70, y at next level, no associativity, only parsing). Arguments complement A R x y /. Arguments Proper {A}%type R%signature m. @@ -68,6 +69,14 @@ Instance fun_Setoid_pointwise_compat A B `{Setoid B} : subrelation (@equiv _ (fun_Setoid A _)) (pointwise_relation _ equiv). Proof using . intros f g Hfg x. apply Hfg. Qed. +Instance eq_Setoid_eq_compat (T : Type) : + subrelation (@equiv T (eq_setoid T)) (@eq T). +Proof using . apply subrelation_refl. Qed. + +Instance eq_eq_Setoid_compat (T : Type) : + subrelation (@eq T) (@equiv T (eq_setoid T)). +Proof using . apply subrelation_refl. Qed. + Notation "x =?= y" := (equiv_dec x y) (at level 70, no associativity). (** Lifting an equivalence relation to an option type. *) @@ -78,13 +87,13 @@ Definition opt_eq {T} (eqT : T -> T -> Prop) (xo yo : option T) := | Some x, Some y => eqT x y end. -Instance opt_eq_refl : forall T (R : T -> T -> Prop), Reflexive R -> Reflexive (opt_eq R). +Instance opt_eq_refl : ∀ T (R : T -> T -> Prop), Reflexive R -> Reflexive (opt_eq R). Proof using . intros T R HR [x |]; simpl; auto. Qed. -Instance opt_eq_sym : forall T (R : T -> T -> Prop), Symmetric R -> Symmetric (opt_eq R). +Instance opt_eq_sym : ∀ T (R : T -> T -> Prop), Symmetric R -> Symmetric (opt_eq R). Proof using . intros T R HR [x |] [y |]; simpl; auto. Qed. -Instance opt_eq_trans : forall T (R : T -> T -> Prop), Transitive R -> Transitive (opt_eq R). +Instance opt_eq_trans : ∀ T (R : T -> T -> Prop), Transitive R -> Transitive (opt_eq R). Proof using . intros T R HR [x |] [y |] [z |]; simpl; intros; eauto; contradiction. Qed. Instance opt_equiv T eqT (HeqT : @Equivalence T eqT) : Equivalence (opt_eq eqT). @@ -95,9 +104,63 @@ Instance opt_Setoid T (S : Setoid T) : Setoid (option T) := {| equiv := opt_eq e Instance Some_compat `(Setoid) : Proper (equiv ==> @equiv _ (opt_Setoid _)) Some. Proof using . intros ? ? Heq. apply Heq. Qed. -Lemma Some_injective T (S : Setoid T) : forall (t1 t2 : T), - @equiv (option T) (opt_Setoid S) (Some t1) (Some t2) -> @equiv T S t1 t2. -Proof using . cbn. intros t1 t2 H. exact H. Qed. +Lemma Some_inj {T : Type} {ST : Setoid T} : + injective (@equiv _ ST) (@equiv _ (opt_Setoid ST)) Some. +Proof using . intros t1 t2 H. exact H. Qed. + +Lemma Some_eq_inj {T : Type} : @injective T _ eq eq Some. +Proof using . intros t1 t2 H1. inversion_clear H1. reflexivity. Qed. + +Lemma not_None {T : Type} {ST : Setoid T} : + ∀ ot : option T, ¬ (@equiv _ (opt_Setoid ST) ot None) + -> {t : T | @equiv _ (opt_Setoid ST) ot (Some t)}. +Proof using . + intros * H. destruct ot as [t|]. constructor 1 with t. + reflexivity. contradict H. reflexivity. +Qed. + +Lemma not_Some {T : Type} {ST : Setoid T} : ∀ ot : option T, + (∀ t : T, ¬ (@equiv _ (opt_Setoid ST) ot (Some t))) + -> @equiv _ (opt_Setoid ST) ot None. +Proof using . + intros * H. destruct ot as [t|]. exfalso. apply (H t). all: reflexivity. +Qed. + +Lemma not_None_iff {T : Type} {ST : Setoid T} : + ∀ ot : option T, ¬ (@equiv _ (opt_Setoid ST) ot None) + <-> ∃ t : T, @equiv _ (opt_Setoid ST) ot (Some t). +Proof using . + intros. etransitivity. 2: split. 2: apply inhabited_sig_to_exists. + 2: apply exists_to_inhabited_sig. split. + - intros H. constructor. apply not_None, H. + - intros [[t H]] Hn. rewrite H in Hn. inversion Hn. +Qed. + +Lemma option_decidable {T : Type} {ST : Setoid T} : ∀ ot : option T, + Decidable.decidable (@equiv _ (opt_Setoid ST) ot None). +Proof using . + intros. destruct ot as [t|]. right. intros H. inversion H. left. reflexivity. +Qed. + +Lemma pick_Some_None : + ∀ {T : Type} {ST : Setoid T} (P : T -> Prop) (oT : option T), + (∀ t : T, @equiv _ (opt_Setoid ST) oT (Some t) <-> P t) + -> (@equiv _ (opt_Setoid ST) oT None <-> (∀ t : T, ¬ (P t))) + -> pick_spec P oT. +Proof using . + intros T ST * Hs Hn. destruct oT as [t|]. apply Pick, Hs. + 2: apply Nopick, Hn. all: reflexivity. +Qed. + +Lemma opt_Setoid_eq : ∀ {T : Type} {ST : Setoid T} (o1 o2 : option T), + (∀ t1 t2 : T, @equiv _ ST t1 t2 <-> t1 = t2) + -> @equiv _ (opt_Setoid ST) o1 o2 <-> o1 = o2. +Proof using . + intros * H1. destruct o1 as [t1|], o2 as [t2|]. + 4: split; intros; reflexivity. 2,3: split; intros H2; inversion H2. + rewrite (injective_compat_iff Some_inj), (injective_compat_iff Some_eq_inj). + apply H1. +Qed. (* This Setoid could be written using RelProd in which case, revealing the "and" inside would be more troublesome than @@ -133,6 +196,14 @@ Instance pair_compat (A B : Type) (SA : Setoid A) (SB : Setoid B) : Proper (@equiv A SA ==> @equiv B SB ==> @equiv (A*B) (prod_Setoid SA SB)) pair. Proof using . intros ?? Hf ?? Hs. split; cbn. exact Hf. exact Hs. Qed. +Lemma prod_Setoid_eq : ∀ {A B : Type} {SA : Setoid A} {SB : Setoid B} + (p1 p2 : A * B), (∀ a1 a2 : A, @equiv _ SA a1 a2 <-> a1 = a2) + -> (∀ b1 b2 : B, @equiv _ SB b1 b2 <-> b1 = b2) + -> @equiv _ (prod_Setoid SA SB) p1 p2 <-> p1 = p2. +Proof using . + intros * Ha Hb. rewrite <- pair_eqE, <- Ha, <- Hb. reflexivity. +Qed. + (* Setoid over [sig] types *) Instance sig_Setoid {T} (S : Setoid T) {P : T -> Prop} : Setoid (sig P). simple refine {| equiv := fun x y => proj1_sig x == proj1_sig y |}; auto; []. @@ -147,9 +218,11 @@ Proof. intros ? ?. simpl. apply equiv_dec. Defined. Instance proj1_sig_compat {T} {S : Setoid T} (E : EqDec S) (P : T -> Prop) : Proper (@equiv _ (sig_Setoid S) ==> equiv) (@proj1_sig T P). -Proof using . - intros ?? H. apply H. -Qed. +Proof using . intros ?? H. apply H. Qed. + +Lemma proj1_sig_inj : ∀ {T : Type} {S : Setoid T} (P : T -> Prop), + injective (@equiv (sig P) (sig_Setoid S)) (@equiv T S) (@proj1_sig T P). +Proof using . intros * ?? H. cbn. apply H. Qed. Instance sigT_Setoid {T} (S : Setoid T) {P : T -> Type} : Setoid (sigT P). simple refine {| equiv := fun x y => projT1 x == projT1 y |}; auto; []. @@ -164,9 +237,7 @@ Proof. intros ? ?. simpl. apply equiv_dec. Defined. Instance projT1_compat {T} {S : Setoid T} (E : EqDec S) (P : T -> Type) : Proper (@equiv _ (sigT_Setoid S) ==> equiv) (@projT1 T P). -Proof using . - intros ?? H. apply H. -Qed. +Proof using . intros ?? H. apply H. Qed. (** The intersection of equivalence relations is still an equivalence relation. *) Lemma inter_equivalence T R1 R2 (E1 : Equivalence R1) (E2 : Equivalence R2) @@ -184,7 +255,7 @@ Definition inter_Setoid {T} (S1 : Setoid T) (S2 : Setoid T) : Setoid T := {| Definition inter_EqDec {T} {S1 S2 : Setoid T} (E1 : EqDec S1) (E2 : EqDec S2) : EqDec (inter_Setoid S1 S2). -Proof. +Proof using . intros x y. destruct (E1 x y), (E2 x y); (now left; split) || (right; intros []; contradiction). Defined. diff --git a/_CoqProject b/_CoqProject index d4783f9107623496f9de6be94c364fb8cb422321..1caec51725b2cc0b53e0d0bf204cc5c2f352bca8 100644 --- a/_CoqProject +++ b/_CoqProject @@ -73,6 +73,7 @@ Spaces/R.v Spaces/R2.v Spaces/Graph.v Spaces/Isomorphism.v +Spaces/ThresholdIsomorphism.v Spaces/Ring.v Spaces/Grid.v