├── .depend
├── .gitignore
├── Banach_Tarski.v
├── Countable.v
├── Equidecomp.v
├── GroupTransf.v
├── LICENSE
├── Makefile
├── Matrix.v
├── Misc.v
├── MiscReals.v
├── MiscTrigo.v
├── Normalize.v
├── NotEmptyPath.v
├── Orbit.v
├── OrbitRepr.v
├── Partition.v
├── Pset.v
├── README.md
├── Reverse.v
├── RnCountable.v
├── Words.v
└── interface.sh
/.depend:
--------------------------------------------------------------------------------
1 | Banach_Tarski.vo Banach_Tarski.glob Banach_Tarski.v.beautified Banach_Tarski.required_vo: Banach_Tarski.v Countable.vo Equidecomp.vo GroupTransf.vo Matrix.vo Misc.vo MiscReals.vo MiscTrigo.vo Normalize.vo NotEmptyPath.vo Orbit.vo OrbitRepr.vo Partition.vo Pset.vo Reverse.vo RnCountable.vo Words.vo
2 | Cohen.vo Cohen.glob Cohen.v.beautified Cohen.required_vo: Cohen.v
3 | Countable.vo Countable.glob Countable.v.beautified Countable.required_vo: Countable.v Misc.vo
4 | Equidecomp.vo Equidecomp.glob Equidecomp.v.beautified Equidecomp.required_vo: Equidecomp.v GroupTransf.vo Matrix.vo Misc.vo OrbitRepr.vo Partition.vo Pset.vo
5 | GroupTransf.vo GroupTransf.glob GroupTransf.v.beautified GroupTransf.required_vo: GroupTransf.v Matrix.vo Misc.vo MiscReals.vo Orbit.vo OrbitRepr.vo Partition.vo Pset.vo Words.vo
6 | Matrix.vo Matrix.glob Matrix.v.beautified Matrix.required_vo: Matrix.v Misc.vo MiscReals.vo MiscTrigo.vo Normalize.vo Reverse.vo Words.vo
7 | Misc.vo Misc.glob Misc.v.beautified Misc.required_vo: Misc.v
8 | MiscReals.vo MiscReals.glob MiscReals.v.beautified MiscReals.required_vo: MiscReals.v Misc.vo Normalize.vo Reverse.vo Words.vo
9 | MiscTrigo.vo MiscTrigo.glob MiscTrigo.v.beautified MiscTrigo.required_vo: MiscTrigo.v MiscReals.vo
10 | Normalize.vo Normalize.glob Normalize.v.beautified Normalize.required_vo: Normalize.v Misc.vo Words.vo
11 | NotEmptyPath.vo NotEmptyPath.glob NotEmptyPath.v.beautified NotEmptyPath.required_vo: NotEmptyPath.v Matrix.vo MiscReals.vo Normalize.vo Reverse.vo Words.vo
12 | Orbit.vo Orbit.glob Orbit.v.beautified Orbit.required_vo: Orbit.v Matrix.vo Misc.vo MiscReals.vo Normalize.vo Pset.vo Reverse.vo Words.vo
13 | OrbitRepr.vo OrbitRepr.glob OrbitRepr.v.beautified OrbitRepr.required_vo: OrbitRepr.v Matrix.vo Misc.vo Normalize.vo Orbit.vo Partition.vo Pset.vo Reverse.vo Words.vo
14 | Partition.vo Partition.glob Partition.v.beautified Partition.required_vo: Partition.v Misc.vo Pset.vo
15 | Pset.vo Pset.glob Pset.v.beautified Pset.required_vo: Pset.v Misc.vo
16 | Reverse.vo Reverse.glob Reverse.v.beautified Reverse.required_vo: Reverse.v Misc.vo Normalize.vo Words.vo
17 | RnCountable.vo RnCountable.glob RnCountable.v.beautified RnCountable.required_vo: RnCountable.v Countable.vo MiscReals.vo
18 | Words.vo Words.glob Words.v.beautified Words.required_vo: Words.v Misc.vo
19 |
--------------------------------------------------------------------------------
/.gitignore:
--------------------------------------------------------------------------------
1 | *.aux
2 | *.bak
3 | *.glob
4 | *.vo
5 | .*.cache
6 | *.vok
7 | *.vos
8 |
--------------------------------------------------------------------------------
/Countable.v:
--------------------------------------------------------------------------------
1 | (* Banach-Tarski paradox. *)
2 |
3 | From Stdlib Require Import Utf8 Bool Compare_dec ZArith Psatz.
4 | Require Import Misc.
5 |
6 | Definition prod_nat_of_nat n :=
7 | let s := Nat.sqrt n in
8 | (s - (n - s ^ 2), n - s ^ 2)%nat.
9 |
10 | Definition prod_nat_to_nat '(i, j) :=
11 | ((i + j) ^ 2 + j)%nat.
12 |
13 | Theorem nat_sqrt_add : ∀ n p, (p ≤ 2 * n)%nat → Nat.sqrt (n * n + p) = n.
14 | Proof.
15 | intros * Hp.
16 | apply Nat.sqrt_unique.
17 | split; [ apply Nat.le_add_r | simpl ].
18 | setoid_rewrite Nat.mul_comm; simpl.
19 | rewrite Nat.add_comm, Nat.add_assoc.
20 | eapply Nat.le_lt_trans; [ apply Nat.add_le_mono_r, Hp | ].
21 | simpl; rewrite Nat.add_0_r.
22 | apply Nat.lt_succ_diag_r.
23 | Qed.
24 |
25 | Theorem prod_nat_to_nat_id : ∀ i j,
26 | prod_nat_of_nat (prod_nat_to_nat (i, j)) = (i, j).
27 | Proof.
28 | intros; simpl.
29 | unfold prod_nat_of_nat.
30 | rewrite Nat.mul_1_r.
31 | remember ((i + j) * (i + j) + j)%nat as n eqn:Hn.
32 | remember (Nat.sqrt n) as s eqn:Hs.
33 | rewrite Hn in Hs.
34 | rewrite nat_sqrt_add in Hs. {
35 | rewrite Nat.add_comm in Hn.
36 | subst; rewrite Nat.pow_2_r.
37 | rewrite Nat.add_sub.
38 | now rewrite Nat.add_sub.
39 | }
40 | simpl.
41 | rewrite Nat.add_0_r, Nat.add_assoc, Nat.add_comm.
42 | apply Nat.le_add_r.
43 | Qed.
44 |
45 | Definition Z_of_nat_surj n :=
46 | if zerop (n mod 2) then Z.of_nat (n / 2)
47 | else (- Z.of_nat (S n / 2))%Z.
48 |
49 | Definition Z_to_nat_inj z :=
50 | if Z_lt_dec z 0 then Z.to_nat (- z * 2 - 1) else Z.to_nat (z * 2).
51 |
52 | Theorem Z2Nat_bij_id : ∀ k, Z_of_nat_surj (Z_to_nat_inj k) = k.
53 | Proof.
54 | intros.
55 | unfold Z_of_nat_surj, Z_to_nat_inj.
56 | destruct (Z_lt_dec k 0) as [Hk | Hk]. {
57 | rewrite Z2Nat.inj_sub; [ simpl | easy ].
58 | unfold Pos.to_nat; simpl.
59 | rewrite <- nat_mod_add_once.
60 | rewrite <- Nat.add_sub_swap. {
61 | rewrite <- Nat.add_sub_assoc; [ simpl | lia ].
62 | rewrite Z2Nat.inj_mul; [ simpl | lia | easy ].
63 | unfold Pos.to_nat; simpl.
64 | rewrite Nat.add_comm.
65 | rewrite Nat.Div0.mod_add.
66 | rewrite <- Nat.sub_succ_l. {
67 | rewrite Nat.sub_succ, Nat.sub_0_r.
68 | rewrite Nat.div_mul; [ | easy ].
69 | rewrite Z2Nat.id; [ | lia ].
70 | now rewrite Z.opp_involutive.
71 | }
72 | remember (Z.to_nat (- k)) as n eqn:Hn.
73 | destruct n; [ | lia ].
74 | apply (f_equal Z.of_nat) in Hn.
75 | rewrite Z2Nat.id in Hn; lia.
76 | }
77 | rewrite Z2Nat.inj_mul; [ simpl | lia | easy ].
78 | unfold Pos.to_nat; simpl.
79 | remember (- k)%Z as l eqn:Hl.
80 | apply (f_equal Z.opp) in Hl.
81 | rewrite Z.opp_involutive in Hl.
82 | subst k; rename l into k; f_equal.
83 | assert (H : (0 < k)%Z) by lia.
84 | clear Hk; rename H into Hk.
85 | remember (Z.to_nat k) as n eqn:Hn.
86 | destruct n; [ | lia ].
87 | apply (f_equal Z.of_nat) in Hn.
88 | rewrite Z2Nat.id in Hn; lia.
89 | }
90 | apply Z.le_ngt in Hk.
91 | rewrite Z2Nat.inj_mul; [ simpl | easy | easy ].
92 | unfold Pos.to_nat; simpl.
93 | rewrite Nat.Div0.mod_mul.
94 | rewrite Nat.div_mul; [ | easy ].
95 | now rewrite Z2Nat.id.
96 | Qed.
97 |
98 | (* Rémi Nollet's code, modified *)
99 |
100 | Theorem Cantor : ∀ E (F : E → (E → bool)), ∃ f : E → bool, ∀ x, f x ≠ F x x.
101 | Proof.
102 | intros E F; exists (fun e => negb (F e e)); intros x H.
103 | exact (no_fixpoint_negb _ H).
104 | Qed.
105 |
106 | Lemma Cantor_gen : ∀ E X Y (Yss : Y → Prop),
107 | ∀ (sX : E → X) (sY : Y → (E → bool)),
108 | ∀ (sX_surj : ∀ e, ∃ x, sX x = e),
109 | ∀ (sY_surj : ∀ f, ∃ y, Yss y ∧ ∀ x, sY y x = f x),
110 | ∀ f : X → Y, ∃ y, ∀ x, Yss y ∧ y ≠ f x.
111 | Proof.
112 | intros * sX_surj sY_surj F.
113 | destruct Cantor with E (fun e => sY (F (sX e))) as [f H].
114 | destruct sY_surj with f as [y Hy]; subst.
115 | destruct Hy as (Hy, Hyf).
116 | exists y; intros x; split; [ easy | ]; subst.
117 | destruct sX_surj with x as [e]; subst.
118 | specialize (H e).
119 | now intros H2; apply H; subst.
120 | Qed.
121 |
122 | (* End Rémi Nollet's code *)
123 |
--------------------------------------------------------------------------------
/Equidecomp.v:
--------------------------------------------------------------------------------
1 | (* Banach-Tarski paradox. *)
2 |
3 | (* Equidecomposability *)
4 |
5 | From Stdlib Require Import Utf8 Arith List Relations.
6 | From Stdlib Require Import Reals.Rdefinitions.
7 | Import ListNotations.
8 |
9 | Require Import Misc Matrix Pset.
10 | Require Import Partition OrbitRepr GroupTransf.
11 |
12 | Definition equidecomposable E₁ E₂ :=
13 | ∃ P₁ P₂, is_partition E₁ P₁ ∧ is_partition E₂ P₂ ∧
14 | List.Forall2 (λ S₁ S₂, ∃ g, (app_gr g S₁ = S₂)%S) P₁ P₂.
15 |
16 | Theorem equidec_refl : reflexive _ equidecomposable.
17 | Proof.
18 | intros E.
19 | exists (E :: []), (E :: []).
20 | split; [ apply is_partition_single | ].
21 | split; [ apply is_partition_single | ].
22 | constructor; [ | constructor ].
23 | exists gr_ident; apply app_gr_ident.
24 | Qed.
25 |
26 | Theorem equidec_sym : symmetric _ equidecomposable.
27 | Proof.
28 | intros E F (P₁ & P₂ & HP₁ & HP₂ & HEF).
29 | exists P₂, P₁.
30 | split; [ easy | ].
31 | split; [ easy | ].
32 | apply Forall2_sym; [ | easy ].
33 | clear -HEF.
34 | intros E F (g & Hg).
35 | exists (gr_inv g); rewrite <- Hg.
36 | apply app_gr_inv_l.
37 | Qed.
38 |
39 | Definition partition_combine {A B} (fl : list (set A → set B)) PE PF :=
40 | flat_map (λ '(fi, Ei), map (λ Fj, Ei ∩ fi Fj) PF) (combine fl PE).
41 |
42 | Definition partition_combine_swi {A B} (fl : list (set A → set B)) PE PF :=
43 | flat_map (λ Fj, map (λ '(fi, Ei), Ei ∩ fi Fj) (combine fl PE)) PF.
44 |
45 | Theorem partition_combine_nth :
46 | ∀ A fl (PE : list (set A)) PF i len,
47 | len = length PF
48 | → length fl = length PE
49 | → (∀ f, List.In f fl → (f ∅ = ∅)%S)
50 | → ((partition_combine fl PE PF).[i] =
51 | PE.[i / len] ∩ nth (i / len) fl id PF.[i mod len])%S.
52 | Proof.
53 | intros * Hlen HlfP Hf.
54 | subst len.
55 | unfold partition_combine; simpl.
56 | revert fl PF i HlfP Hf.
57 | induction PE as [| E₁ PE]; intros. {
58 | destruct fl as [| f₁ fl]; [ | easy ].
59 | intros x.
60 | split; intros Hx; [ now destruct i | ].
61 | destruct Hx as (Hx, _).
62 | now destruct (i / length PF)%nat.
63 | }
64 | destruct fl as [| f₁ fl]; [ easy | ].
65 | simpl in HlfP; apply Nat.succ_inj in HlfP; simpl.
66 | destruct (lt_dec i (length PF)) as [Hi| Hi]. {
67 | rewrite app_nth1; [| now rewrite length_map ].
68 | rewrite Nat.div_small; [ simpl | easy ].
69 | rewrite Nat.mod_small; [ simpl | easy ].
70 | intros x; clear - HlfP Hf.
71 | split; intros Hx. {
72 | revert i Hx.
73 | induction PF as [| F₁ PF]; intros; [ now destruct i | ].
74 | simpl in Hx; simpl.
75 | destruct i; [ easy | now apply IHPF ].
76 | }
77 | revert i Hx.
78 | induction PF as [| F₁ PF]; intros. {
79 | destruct Hx as (_, Hx).
80 | destruct i; simpl in Hx; simpl.
81 | rewrite Hf in Hx; [ easy | now left ].
82 | rewrite Hf in Hx; [ easy | now left ].
83 | }
84 | destruct i; simpl in Hx; simpl; [ easy | ].
85 | now apply IHPF.
86 | }
87 | apply Nat.nlt_ge in Hi.
88 | rewrite app_nth2; [| now rewrite length_map ].
89 | rewrite length_map.
90 | remember (i - length PF)%nat as j eqn:Hj.
91 | assert (H : (i = j + length PF)%nat). {
92 | rewrite Hj.
93 | now rewrite Nat.sub_add.
94 | }
95 | subst i; clear Hi Hj.
96 | destruct PF as [| F₁ PF]. {
97 | simpl.
98 | rewrite match_id.
99 | intros x.
100 | split; intros Hx. {
101 | destruct j; simpl in Hx. {
102 | induction (combine fl PE) as [| (y, z) l]; [ easy | ].
103 | apply IHl, Hx.
104 | }
105 | induction (combine fl PE) as [| (y, z) l]; [ easy | ].
106 | apply IHl, Hx.
107 | }
108 | rewrite Hf in Hx; [ | now left ].
109 | now rewrite set_inter_empty_r in Hx.
110 | }
111 | rewrite nat_mod_add_once.
112 | rewrite nat_div_add_once; [ | easy ].
113 | apply IHPE; [ easy | ].
114 | intros f Hfi.
115 | now apply Hf; right.
116 | Qed.
117 |
118 | Theorem partition_combine_swi_nth :
119 | ∀ A fl (PE : list (set A)) PF i len,
120 | len = length PE
121 | → length fl = length PE
122 | → (∀ f, List.In f fl → (f ∅ = ∅)%S)
123 | → ((partition_combine_swi fl PE PF).[i] =
124 | PE.[i mod len] ∩ nth (i mod len) fl id PF.[i / len])%S.
125 | Proof.
126 | intros * Hlen HlfP Hf.
127 | subst len.
128 | unfold partition_combine_swi; simpl.
129 | revert fl PE i HlfP Hf.
130 | induction PF as [| F₁ PF]; intros. {
131 | simpl; do 2 rewrite match_id.
132 | clear -Hf.
133 | remember (i mod length PE) as j eqn:Hj; clear Hj.
134 | assert (Hj : ∀ j, (nth j fl id ∅ = ∅)%S). {
135 | clear -Hf.
136 | induction fl as [| f₁ fl]; intros; [ simpl; now rewrite match_id | ].
137 | destruct j; [ now apply Hf; left | simpl ].
138 | apply IHfl.
139 | intros f₂ Hf₂; now apply Hf; right.
140 | }
141 | now rewrite Hj, set_inter_empty_r.
142 | }
143 | simpl.
144 | destruct (lt_dec i (length fl)) as [Hi| Hi]. {
145 | rewrite app_nth1. {
146 | rewrite HlfP in Hi.
147 | rewrite Nat.mod_small; [ | easy ].
148 | rewrite Nat.div_small; [ | easy ].
149 | intros x; clear - HlfP Hf.
150 | split; intros Hx. {
151 | revert i fl Hx HlfP Hf.
152 | induction PE as [| E₁ PE]; intros. {
153 | apply length_zero_iff_nil in HlfP; subst fl.
154 | simpl in Hx; now rewrite match_id in Hx.
155 | }
156 | destruct fl as [| f₁ fl]; [ easy | ].
157 | simpl in HlfP; apply Nat.succ_inj in HlfP.
158 | destruct i; [ easy | ].
159 | simpl in Hx; simpl.
160 | apply IHPE; [ easy | easy | ].
161 | intros f Hfi.
162 | now apply Hf; right.
163 | }
164 | destruct Hx as (Hx, Hxn).
165 | revert x i fl HlfP Hf Hx Hxn.
166 | induction PE as [| E₁ PE]; intros. {
167 | now simpl in Hx; rewrite match_id in Hx.
168 | }
169 | destruct fl as [| f₁ fl]; [ easy | ].
170 | simpl in HlfP; apply Nat.succ_inj in HlfP.
171 | simpl in Hx; simpl.
172 | destruct i; [ now split | ].
173 | apply IHPE; [ easy | | easy | easy ].
174 | intros f Hfi.
175 | now apply Hf; right.
176 | }
177 | rewrite length_map, length_combine, Nat.min_l; [ easy | ].
178 | now rewrite HlfP.
179 | }
180 | apply Nat.nlt_ge in Hi.
181 | rewrite app_nth2. {
182 | rewrite IHPF; [ | easy | easy ].
183 | rewrite length_map, length_combine, Nat.min_l; [ | now rewrite HlfP ].
184 | rewrite HlfP in Hi.
185 | remember (length PE) as len eqn:Hlen; symmetry in Hlen.
186 | destruct (eq_nat_dec len 0) as [Hzl| Hnzl]. {
187 | move Hzl at top; subst len.
188 | apply length_zero_iff_nil in Hlen.
189 | apply length_zero_iff_nil in HlfP.
190 | subst PE fl; simpl.
191 | do 4 rewrite match_id.
192 | now do 2 rewrite set_inter_empty_l.
193 | }
194 | subst len.
195 | generalize Hi; intros H.
196 | apply Nat.Div0.div_le_mono with (c := length PE) in H.
197 | rewrite Nat.div_same in H; [ | easy ].
198 | remember (i / length PE) as j eqn:Hj; symmetry in Hj.
199 | destruct j; [ now apply Nat.le_0_r in H | ].
200 | rewrite HlfP.
201 | remember (i - length PE) as k eqn:Hk.
202 | assert (i = k + length PE) by (now subst k; rewrite Nat.sub_add).
203 | subst i; clear Hk.
204 | rewrite nat_mod_add_once.
205 | rewrite nat_div_add_once in Hj; [ | easy ].
206 | apply Nat.succ_inj in Hj.
207 | now rewrite Hj.
208 | }
209 | rewrite length_map, length_combine, Nat.min_l; [ easy | ].
210 | now rewrite HlfP.
211 | Qed.
212 |
213 | Theorem partition_combine_is_partition :
214 | ∀ E F PE PF P'F gl,
215 | is_partition E PE
216 | → is_partition F PF
217 | → length PE = length PF
218 | → length gl = length PE
219 | → is_partition F P'F
220 | → (∀ i : nat, (app_gr (nth i gl gr_ident) PE.[i] = PF.[i])%S)
221 | → ∀ fl, fl = map app_gr_inv gl
222 | → is_partition E (partition_combine fl PE P'F).
223 | Proof.
224 | intros * HPE HPF Hlen1 Hlen3 HP'F Hgl * Hfl.
225 | split. {
226 | destruct HPE as (HPEU, _).
227 | destruct HPF as (HPFU, _).
228 | destruct HP'F as (HP'FU, _).
229 | assert (HUP'F : F ⊂ ⋃ P'F) by now rewrite HP'FU; intros x H.
230 | clear HP'FU.
231 | progress unfold partition_combine.
232 | subst fl.
233 | revert E F gl PF P'F HPEU HPFU HUP'F Hlen1 Hlen3 Hgl.
234 | induction PE as [| E₁ PE]; intros. {
235 | now apply length_zero_iff_nil in Hlen3; subst gl.
236 | }
237 | destruct gl as [| g₁ gl]; [ easy | ].
238 | rewrite HPEU; simpl.
239 | rewrite set_union_list_app.
240 | simpl in Hlen3; apply Nat.succ_inj in Hlen3.
241 | apply set_union_morph. {
242 | pose proof set_union_inter_self vector E₁ (map (app_gr_inv g₁) P'F).
243 | rewrite map_map in H.
244 | apply H.
245 | assert (HEF : E₁ ⊂ app_gr_inv g₁ F). {
246 | rewrite HPFU.
247 | apply included_group with g₁.
248 | rewrite app_gr_inv_r.
249 | intros p Hp.
250 | pose proof Hgl 0 p as Hgl₁; simpl in Hgl₁.
251 | apply Hgl₁ in Hp.
252 | destruct PF as [| P₁ PF]; [ easy | simpl in Hp ].
253 | now left.
254 | }
255 | apply included_group with (g := gr_inv g₁) in HUP'F.
256 | rewrite group_set_union_list_distr in HUP'F.
257 | rewrite fold_app_gr_inv in HUP'F.
258 | eapply included_trans; eassumption.
259 | }
260 | destruct PF as [| F₁ PF]; [ easy | ].
261 | simpl in Hlen1; apply Nat.succ_inj in Hlen1.
262 | eapply IHPE; [ | | | eassumption | easy | ]; try easy. {
263 | rewrite HPFU in HUP'F.
264 | intros p Hp; apply HUP'F.
265 | now right.
266 | }
267 | intros i.
268 | now pose proof (Hgl (S i)) as H; simpl in H.
269 | }
270 | intros i j Hij.
271 | erewrite partition_combine_nth; [ | easy | | ]. {
272 | erewrite partition_combine_nth; [ | easy | | ]. {
273 | remember (length P'F) as len eqn:Hlen.
274 | destruct len. {
275 | symmetry in Hlen.
276 | apply length_zero_iff_nil in Hlen; subst P'F; simpl.
277 | do 2 rewrite match_id.
278 | subst fl.
279 | destruct gl as [| g₁ gl]. {
280 | simpl; unfold id at 2; simpl.
281 | now do 2 rewrite set_inter_empty_r.
282 | }
283 | simpl; unfold app_gr_inv, Nat.div; rewrite app_gr_empty_set.
284 | now do 2 rewrite set_inter_empty_r.
285 | }
286 | destruct HPE as (HPEU, HPEI).
287 | destruct HP'F as (HP'FU, HP'FI).
288 | destruct (eq_nat_dec (i / S len) (j / S len)) as [Hd| Hd]. {
289 | destruct (eq_nat_dec (i mod S len) (j mod S len)) as [Hm| Hm]. {
290 | assert (Hnlen : (S len ≠ 0)%nat) by easy.
291 | pose proof Nat.div_mod i (S len) Hnlen as Hi.
292 | pose proof Nat.div_mod j (S len) Hnlen as Hj.
293 | now rewrite Hd, Hm, <- Hj in Hi.
294 | }
295 | subst fl; rewrite <- Hd; simpl.
296 | pose proof map_nth app_gr_inv gl gr_ident (i / S len) as Hi.
297 | destruct (lt_dec (i / S len) (length gl)) as [Hil| Hil]. {
298 | rewrite nth_indep with (d' := id) in Hi. {
299 | rewrite Hi, set_inter_shuffle0.
300 | rewrite set_inter_assoc, <- set_inter_assoc.
301 | unfold app_gr_inv; rewrite <- group_set_inter_distr.
302 | apply not_eq_sym in Hm.
303 | rewrite HP'FI; [ | easy ].
304 | rewrite app_gr_empty_set.
305 | apply set_inter_empty_r.
306 | }
307 | now rewrite length_map.
308 | }
309 | apply Nat.nlt_ge in Hil.
310 | rewrite Hlen3 in Hil.
311 | rewrite nth_overflow; [ | easy ].
312 | now do 2 rewrite set_inter_empty_l.
313 | }
314 | rewrite set_inter_shuffle0, set_inter_assoc.
315 | rewrite HPEI; [ | easy ].
316 | now do 2 rewrite set_inter_empty_l.
317 | } {
318 | now subst fl; rewrite length_map.
319 | }
320 | subst fl.
321 | intros f Hf.
322 | apply in_map_iff in Hf.
323 | destruct Hf as (g & Hg & Hix).
324 | subst f; apply app_gr_empty_set.
325 | } {
326 | now subst fl; rewrite length_map.
327 | }
328 | subst fl.
329 | intros f Hf.
330 | apply in_map_iff in Hf.
331 | destruct Hf as (g & Hg & Hix).
332 | subst f; apply app_gr_empty_set.
333 | Qed.
334 |
335 | From Stdlib Require Import Permutation.
336 |
337 | Theorem partition_combine_swi_is_permutation :
338 | ∀ A (fl : list (set A → set A)) PE P'F,
339 | Permutation (partition_combine_swi fl PE P'F) (partition_combine fl PE P'F).
340 | Proof.
341 | intros.
342 | unfold partition_combine, partition_combine_swi.
343 | rewrite Permutation_flat_map_map.
344 | induction (combine fl PE) as [| a la]; intros; [ easy | ].
345 | now simpl; rewrite IHla; destruct a.
346 | Qed.
347 |
348 | Theorem permuted_partition_is_partition :
349 | ∀ A (E : set A) PE P'E,
350 | Permutation PE P'E
351 | → is_partition E PE
352 | → is_partition E P'E.
353 | Proof.
354 | intros * Hpe Hpa.
355 | destruct Hpa as (Hpau, Hpai).
356 | split. {
357 | rewrite Hpau; clear -Hpe.
358 | induction Hpe; [ easy | | | ]. {
359 | now simpl; rewrite IHHpe.
360 | } {
361 | simpl; rewrite set_union_comm, <- set_union_assoc.
362 | apply set_union_morph; [ easy | apply set_union_comm ].
363 | }
364 | etransitivity; eassumption.
365 | }
366 | intros i j Hij x.
367 | split; [ intros Hx; simpl | easy ].
368 | apply Permutation_nth_error in Hpe.
369 | destruct Hpe as (Hlen & f & Hfi & Hn).
370 | unfold FinFun.Injective in Hfi.
371 | assert (Hfij : f i ≠ f j) by now intros H; apply Hfi in H.
372 | assert (HP'P : ∀ i, P'E.[i] = PE.[f i]). {
373 | intros k.
374 | pose proof Hn k as Hk.
375 | remember (nth_error P'E k) as p'k eqn:H'k.
376 | symmetry in Hk, H'k.
377 | destruct p'k as [v | ]. {
378 | apply nth_error_split in Hk.
379 | apply nth_error_split in H'k.
380 | destruct Hk as (l1 & l2 & HPE & Hlen1).
381 | destruct H'k as (l'1 & l'2 & HP'E & Hlen'1).
382 | rewrite HPE, HP'E, <- Hlen1, <- Hlen'1.
383 | rewrite app_nth2; [ | now unfold ge ].
384 | rewrite app_nth2; [ | now unfold ge ].
385 | now do 2 rewrite Nat.sub_diag.
386 | }
387 | apply nth_error_None in Hk.
388 | apply nth_error_None in H'k.
389 | rewrite nth_overflow; [ | easy ].
390 | now rewrite nth_overflow.
391 | }
392 | do 2 rewrite HP'P in Hx.
393 | now rewrite Hpai in Hx.
394 | Qed.
395 |
396 | Theorem partition_combine_partition_combine_swi :
397 | ∀ A E (fl : list (set A → set A)) PE P'F,
398 | is_partition E (partition_combine fl PE P'F)
399 | → is_partition E (partition_combine_swi fl PE P'F).
400 | Proof.
401 | intros * HP.
402 | eapply permuted_partition_is_partition; [ | eassumption ].
403 | symmetry.
404 | apply partition_combine_swi_is_permutation.
405 | Qed.
406 |
407 | Theorem partition_combine_swi_is_partition :
408 | ∀ E F PE PF P'F gl,
409 | is_partition E PE
410 | → is_partition F PF
411 | → length PE = length PF
412 | → length gl = length PE
413 | → is_partition F P'F
414 | → (∀ i : nat, (app_gr (nth i gl gr_ident) PE.[i] = PF.[i])%S)
415 | → ∀ fl, fl = map app_gr_inv gl
416 | → is_partition E (partition_combine_swi fl PE P'F).
417 | Proof.
418 | intros * HPE HPF Hlen1 Hlen3 HP'F Hgl * Hfl.
419 | apply partition_combine_partition_combine_swi.
420 | eapply partition_combine_is_partition with (F := F) (PF := PF); eassumption.
421 | Qed.
422 |
423 | Theorem partition_length_combine :
424 | ∀ A fl (PE PF : list (set A)),
425 | length fl = length PE
426 | → length (partition_combine fl PE PF) = (length PE * length PF)%nat.
427 | Proof.
428 | intros * Hlen.
429 | unfold partition_combine; simpl.
430 | revert fl PF Hlen.
431 | induction PE as [| E₁ PE]; intros; [ now destruct fl | simpl ].
432 | destruct fl as [| f₁ fl]; [ easy | ].
433 | destruct PF as [| F₁ FL]. {
434 | unfold partition_combine; simpl.
435 | rewrite Nat.mul_0_r.
436 | induction (combine fl PE) as [| (x, y) l]; [ easy | apply IHl ].
437 | }
438 | simpl in Hlen; simpl; f_equal.
439 | rewrite length_app, length_map.
440 | apply Nat.succ_inj in Hlen.
441 | apply IHPE with (PF := F₁ :: FL) in Hlen.
442 | now simpl in Hlen; rewrite Hlen.
443 | Qed.
444 |
445 | Theorem partition_combine_swi_length :
446 | ∀ A fl (PE PF : list (set A)),
447 | length fl = length PE
448 | → length (partition_combine_swi fl PE PF) = (length PE * length PF)%nat.
449 | Proof.
450 | intros * Hlen.
451 | pose proof partition_combine_swi_is_permutation _ fl PE PF as H.
452 | apply Permutation_length in H.
453 | now rewrite H; apply partition_length_combine.
454 | Qed.
455 |
456 | Require Import Setoid.
457 | Add Parametric Morphism :
458 | (λ n fl, @nth (set vector → set vector) n (map app_gr_inv fl) id)
459 | with signature eq ==> eq ==> set_eq ==> set_eq
460 | as nth_map_app_gr_inv_morph.
461 | Proof.
462 | intros n fl E F HEF x.
463 | split; intros Hx. {
464 | revert n Hx.
465 | induction fl as [| f₁ fl]; intros. {
466 | simpl in Hx; simpl; rewrite match_id in Hx |-*; now apply HEF.
467 | }
468 | simpl in Hx; simpl.
469 | now destruct n; [ rewrite <- HEF | apply IHfl ].
470 | }
471 | revert n Hx.
472 | induction fl as [| f₁ fl]; intros. {
473 | simpl in Hx; simpl; rewrite match_id in Hx |-*; now apply HEF.
474 | }
475 | simpl in Hx; simpl.
476 | now destruct n; [ rewrite HEF | apply IHfl ].
477 | Qed.
478 |
479 | Theorem equidec_trans : transitive _ equidecomposable.
480 | Proof.
481 | intros E F G HEF HFG.
482 | unfold equidecomposable.
483 | destruct HEF as (PE & P₁F & HPE & HP₁F & HEF).
484 | destruct HFG as (P₂F & PG & HP₂F & HPG & HFG).
485 | assert
486 | (Hgl : ∃ gl, length gl = length PE ∧
487 | ∀ i, (app_gr (nth i gl gr_ident) (nth i PE ∅) = nth i P₁F ∅)%S). {
488 | apply Forall2_Forall_combine in HEF.
489 | destruct HEF as (HEF, Hlen1).
490 | clear HPE HP₁F.
491 | revert P₁F Hlen1 HEF.
492 | induction PE as [| E₁ PE]; intros. {
493 | exists []; split; [ easy | ].
494 | symmetry in Hlen1; apply length_zero_iff_nil in Hlen1; subst P₁F.
495 | intros i; simpl.
496 | do 2 rewrite match_id; simpl.
497 | now intros (x, y, z); split.
498 | }
499 | destruct P₁F as [| F₁ P₁F]; [ easy | ].
500 | simpl in Hlen1; apply Nat.succ_inj in Hlen1.
501 | simpl in HEF; apply Forall_inv2 in HEF.
502 | destruct HEF as ((g₁, HgEF), HEF).
503 | apply IHPE in HEF; [ | easy ].
504 | destruct HEF as (gl & Hlen3 & HEF).
505 | exists (g₁ :: gl).
506 | split; [ now simpl; rewrite Hlen3 | ].
507 | intros i; simpl.
508 | destruct i; [ easy | apply HEF ].
509 | }
510 | assert
511 | (Hhl : ∃ hl, length hl = length PG ∧
512 | ∀ i, (app_gr (nth i hl gr_ident) (nth i PG ∅) = nth i P₂F ∅)%S). {
513 | apply Forall2_Forall_combine in HFG.
514 | destruct HFG as (HFG, Hlen2).
515 | clear HPG HP₂F.
516 | revert P₂F Hlen2 HFG.
517 | induction PG as [| G₁ PG]; intros. {
518 | exists []; split; [ easy | ].
519 | apply length_zero_iff_nil in Hlen2; subst P₂F.
520 | intros i; simpl.
521 | do 2 rewrite match_id; simpl.
522 | now intros (x, y, z); split.
523 | }
524 | destruct P₂F as [| F₁ P₂F]; [ easy | ].
525 | simpl in Hlen2; apply Nat.succ_inj in Hlen2.
526 | simpl in HFG; apply Forall_inv2 in HFG.
527 | destruct HFG as ((h₁, HhFG), HFG).
528 | apply IHPG in HFG; [ | easy ].
529 | destruct HFG as (hl & Hlen3 & HFG).
530 | exists (gr_inv h₁ :: hl).
531 | split; [ now simpl; rewrite Hlen3 | ].
532 | intros i; simpl.
533 | destruct i; [ | apply HFG ].
534 | rewrite <- HhFG, fold_app_gr_inv.
535 | now rewrite app_gr_inv_l.
536 | }
537 | destruct Hgl as (gl & Hlen3 & Hgl).
538 | destruct Hhl as (hl & Hlen4 & Hhl).
539 | apply Forall2_Forall_combine in HEF.
540 | destruct HEF as (HEF, Hlen1).
541 | apply Forall2_Forall_combine in HFG.
542 | destruct HFG as (HFG, Hlen2).
543 | remember (map app_gr_inv gl) as g'l eqn:Hg'l.
544 | assert (Hpcf : is_partition E (partition_combine g'l PE P₂F)). {
545 | eapply partition_combine_is_partition with (PF := P₁F); eassumption.
546 | }
547 | exists (partition_combine g'l PE P₂F).
548 | remember (map app_gr_inv hl) as h'l eqn:Hh'l.
549 | assert (Hpcg : is_partition G (partition_combine_swi h'l PG P₁F)). {
550 | symmetry in Hlen2.
551 | eapply partition_combine_swi_is_partition with (PF := P₂F); eassumption.
552 | }
553 | exists (partition_combine_swi h'l PG P₁F).
554 | split; [ easy | ].
555 | split; [ easy | ].
556 | apply Forall2_Forall_combine.
557 | split. {
558 | apply Forall_forall.
559 | intros (U, V) HUV.
560 | apply In_nth with (d := (∅, ∅)) in HUV.
561 | rewrite length_combine in HUV.
562 | rewrite partition_length_combine in HUV. {
563 | rewrite partition_combine_swi_length in HUV. {
564 | rewrite <- Hlen1, Hlen2, Nat.mul_comm in HUV.
565 | rewrite Nat.min_l in HUV; [ | easy ].
566 | destruct HUV as (i & Hi & HUV).
567 | rewrite combine_nth in HUV. {
568 | injection HUV; clear HUV; intros HV HU.
569 | apply eq_set_eq in HU.
570 | apply eq_set_eq in HV.
571 | remember (partition_combine g'l PE P₂F) as PE' eqn:HPE'.
572 | remember (partition_combine_swi h'l PG P₁F) as PG' eqn:HPG'.
573 | subst g'l h'l.
574 | destruct Hpcf as (HpcfU, HpcfI).
575 | destruct Hpcg as (HpcgU, HpcgI).
576 | remember (nth (i / length PG) gl gr_ident) as gi.
577 | remember (nth (i mod length PG) (map gr_inv hl) gr_ident) as hj.
578 | exists (Comb hj gi); subst gi hj; simpl.
579 | rewrite <- HU, <- HV; clear HU HV.
580 | rewrite HPE', HPG'.
581 | rewrite partition_combine_nth; [ | easy | | ]. {
582 | rewrite partition_combine_swi_nth; [ | easy | | ]. {
583 | do 2 rewrite group_set_inter_distr.
584 | rewrite Hlen2, Hgl.
585 | rewrite set_inter_comm.
586 | apply set_inter_morph. {
587 | rewrite app_gr_nth.
588 | replace Datatypes.id with (@id (set vector)) by easy.
589 | rewrite map_map.
590 | (* does not work, I don't know why
591 | rewrite <- Hhl.
592 | *)
593 | (* using transitivity instead *)
594 | etransitivity. {
595 | apply nth_map_app_gr_inv_morph_Proper; [ easy | easy | ].
596 | apply app_gr_morph_Proper; [ easy | ].
597 | apply nth_map_app_gr_inv_morph_Proper; [ easy | easy | ].
598 | symmetry; apply Hhl.
599 | }
600 | do 2 rewrite <- app_gr_nth_inv.
601 | assert (HPGnz : length PG ≠ 0). {
602 | intros H; rewrite H in Hi.
603 | now apply Nat.nlt_0_r in Hi.
604 | }
605 | setoid_rewrite nth_indep with (d' := gr_inv gr_ident). {
606 | do 2 rewrite map_nth.
607 | rewrite gr_inv_ident.
608 | remember (nth (i / length PG) gl gr_ident) as x.
609 | do 2 rewrite fold_app_gr_inv.
610 | rewrite app_gr_app_gr_inv.
611 | now rewrite app_gr_inv_app_gr.
612 | } {
613 | rewrite length_map, Hlen4.
614 | now apply Nat.mod_upper_bound.
615 | } {
616 | now rewrite Hlen3; apply Nat.Div0.div_lt_upper_bound.
617 | } {
618 | rewrite length_map, Hlen3.
619 | now apply Nat.Div0.div_lt_upper_bound.
620 | } {
621 | rewrite Hlen4.
622 | now apply Nat.mod_upper_bound.
623 | }
624 | }
625 | rewrite app_gr_nth.
626 | replace Datatypes.id with (@id (set vector)) by easy.
627 | now rewrite map_map.
628 | } {
629 | now rewrite length_map.
630 | }
631 | intros f Hif.
632 | clear -Hif.
633 | induction hl as [| h₁ hl]; [ easy | ].
634 | destruct Hif; [ subst f; apply app_gr_empty_set | now apply IHhl ].
635 | } {
636 | now rewrite length_map.
637 | }
638 | intros f Hif.
639 | clear -Hif.
640 | induction gl as [| g₁ gl]; [ easy | ].
641 | destruct Hif; [ subst f; apply app_gr_empty_set | now apply IHgl ].
642 | }
643 | rewrite partition_length_combine. {
644 | rewrite partition_combine_swi_length. {
645 | rewrite Hlen1, Hlen2.
646 | apply Nat.mul_comm.
647 | }
648 | now subst h'l; rewrite length_map.
649 | }
650 | now subst g'l; rewrite length_map.
651 | }
652 | now rewrite Hh'l, length_map.
653 | }
654 | now rewrite Hg'l, length_map.
655 | }
656 | rewrite partition_length_combine. {
657 | rewrite partition_combine_swi_length. {
658 | rewrite Hlen1, Hlen2.
659 | apply Nat.mul_comm.
660 | }
661 | now subst h'l; rewrite length_map.
662 | }
663 | now subst g'l; rewrite length_map.
664 | Qed.
665 |
666 | Add Parametric Relation : (set vector) equidecomposable
667 | reflexivity proved by equidec_refl
668 | symmetry proved by equidec_sym
669 | transitivity proved by equidec_trans
670 | as equidec_rel.
671 |
672 | Theorem equidec_set_union : ∀ E₁ E₂ F₁ F₂,
673 | (E₁ ∩ F₁ = ∅)%S
674 | → (E₂ ∩ F₂ = ∅)%S
675 | → equidecomposable E₁ E₂
676 | → equidecomposable F₁ F₂
677 | → equidecomposable (E₁ ∪ F₁) (E₂ ∪ F₂).
678 | Proof.
679 | intros * HEF₁ HEF₂ HE HF.
680 | destruct HE as (PE₁ & PE₂ & HE₁ & HE₂ & HE).
681 | destruct HF as (PF₁ & PF₂ & HF₁ & HF₂ & HF).
682 | unfold equidecomposable.
683 | exists (PE₁ ++ PF₁), (PE₂ ++ PF₂).
684 | split; [ now apply partition_union | ].
685 | split; [ now apply partition_union | ].
686 | now apply Forall2_app.
687 | Qed.
688 |
689 | Add Parametric Morphism : equidecomposable
690 | with signature set_eq ==> set_eq ==> iff
691 | as equidec_morph.
692 | Proof.
693 | intros E E' HE F F' HF.
694 | split; intros H. {
695 | destruct H as (EL & FL & HEL & HFL & HA).
696 | rewrite HE in HEL; rewrite HF in HFL.
697 | exists EL, FL.
698 | now split; [ | split ].
699 | }
700 | destruct H as (EL & FL & HEL & HFL & HA).
701 | rewrite <- HE in HEL; rewrite <- HF in HFL.
702 | exists EL, FL.
703 | now split; [ | split ].
704 | Qed.
705 |
--------------------------------------------------------------------------------
/GroupTransf.v:
--------------------------------------------------------------------------------
1 | (* Banach-Tarski paradox. *)
2 |
3 | From Stdlib Require Import Utf8.
4 | From Stdlib Require Import Reals Nsatz Psatz.
5 |
6 | Require Import Misc Words MiscReals Matrix Pset Orbit.
7 | Require Import Partition OrbitRepr.
8 |
9 | Definition transl d (E : set vector) := mkset (λ v, (v - d)%vec ∈ E).
10 | Arguments transl d%_vec E%_S.
11 |
12 | Inductive Gr :=
13 | | Rot : ∀ M, is_rotation_matrix M → Gr
14 | | Transl : vector → Gr
15 | | Comb : Gr → Gr → Gr.
16 |
17 | Arguments Transl _%_vec.
18 |
19 | Fixpoint app_gr f p :=
20 | match f with
21 | | Rot M Hrm => set_map (mat_vec_mul M) p
22 | | Transl d => transl d p
23 | | Comb g h => app_gr g (app_gr h p)
24 | end.
25 |
26 | Fixpoint app_gr_vec f p :=
27 | match f with
28 | | Rot M Hrm => (mat_transp M * p)%vec
29 | | Transl d => (p - d)%vec
30 | | Comb g h => app_gr_vec h (app_gr_vec g p)
31 | end.
32 |
33 | Fixpoint gr_inv f :=
34 | match f with
35 | | Rot M Hrm => Rot (mat_transp M) (rotation_transp_is_rotation _ Hrm)
36 | | Transl d => Transl (-d)
37 | | Comb g h => Comb (gr_inv h) (gr_inv g)
38 | end.
39 |
40 | Definition gr_ident := Transl 0.
41 | Definition app_gr_inv g := app_gr (gr_inv g).
42 |
43 | Theorem gr_subst : ∀ g E F,
44 | (E = F)%S → ∀ p, p ∈ app_gr g E → p ∈ app_gr g F.
45 | Proof.
46 | intros * HEF * HE.
47 | revert E F p HEF HE.
48 | induction g as [ M Hrm| dx | g IHg h IHh]; intros. {
49 | destruct HE as (u & HuE & HMu).
50 | now exists u; split; [ apply HEF | ].
51 | } {
52 | destruct p as (x, y, z).
53 | apply HEF, HE.
54 | } {
55 | simpl in HE; simpl.
56 | eapply IHg; [ | eassumption ].
57 | split; intros H; [ eapply IHh; eassumption | ].
58 | eapply IHh; [ symmetry; eassumption | eassumption ].
59 | }
60 | Qed.
61 |
62 | Theorem transl_0 : ∀ E, (transl 0 E = E)%S.
63 | Proof.
64 | intros; intros (x, y, z); simpl.
65 | now rewrite Ropp_0; do 3 rewrite Rplus_0_r.
66 | Qed.
67 |
68 | Theorem transl_transl : ∀ E d₁ d₂,
69 | (transl d₁ (transl d₂ E) = transl (d₁ + d₂) E)%S.
70 | Proof.
71 | intros E (dx₁, dy₁, dz₁) (dx₂, dy₂, dz₂) (x, y, z); simpl.
72 | do 9 rewrite fold_Rminus.
73 | now do 3 rewrite Rminus_plus_distr.
74 | Qed.
75 |
76 | Theorem app_gr_ident : ∀ E, (app_gr gr_ident E = E)%S.
77 | Proof.
78 | intros.
79 | unfold gr_ident; simpl.
80 | unfold transl; simpl.
81 | intros (x, y, z); simpl.
82 | now do 3 rewrite fold_Rminus, Rminus_0_r.
83 | Qed.
84 |
85 | Theorem app_gr_nth : ∀ E fl i,
86 | (app_gr (List.nth i fl gr_ident) E =
87 | List.nth i (map app_gr fl) Datatypes.id E)%S.
88 | Proof.
89 | intros.
90 | revert E i.
91 | induction fl as [| f₁ fl]; intros. {
92 | simpl; do 2 rewrite match_id.
93 | apply app_gr_ident.
94 | }
95 | destruct i; [ easy | apply IHfl ].
96 | Qed.
97 |
98 | Theorem app_gr_nth_inv : ∀ E fl i,
99 | (app_gr (List.nth i (map gr_inv fl) gr_ident) E =
100 | List.nth i (map app_gr_inv fl) Datatypes.id E)%S.
101 | Proof.
102 | intros.
103 | revert E i.
104 | induction fl as [| f₁ fl]; intros. {
105 | simpl; do 2 rewrite match_id.
106 | apply app_gr_ident.
107 | }
108 | destruct i; [ easy | apply IHfl ].
109 | Qed.
110 |
111 | Theorem gr_inv_ident : gr_inv gr_ident = gr_ident.
112 | Proof.
113 | unfold gr_ident; simpl.
114 | now rewrite Ropp_0.
115 | Qed.
116 |
117 | Add Parametric Morphism : app_gr
118 | with signature eq ==> set_eq ==> set_eq
119 | as app_gr_morph.
120 | Proof.
121 | intros g p q Hpq r.
122 | split; intros H; [ eapply gr_subst; eassumption | ].
123 | symmetry in Hpq; eapply gr_subst; eassumption.
124 | Qed.
125 |
126 | Add Parametric Morphism : app_gr_inv
127 | with signature eq ==> set_eq ==> set_eq
128 | as app_gr_inv_morph.
129 | Proof.
130 | intros g p q Hpq r.
131 | split; intros H; [ eapply gr_subst; eassumption | ].
132 | symmetry in Hpq; eapply gr_subst; eassumption.
133 | Qed.
134 |
135 | Add Parametric Morphism : transl
136 | with signature eq ==> set_eq ==> set_eq
137 | as transl_morph.
138 | Proof.
139 | intros dx E F HEF (x, y, z); simpl; now rewrite HEF.
140 | Qed.
141 |
142 | Theorem fold_app_gr_inv : ∀ g, app_gr (gr_inv g) = app_gr_inv g.
143 | Proof. easy. Qed.
144 |
145 | Theorem set_map_mul_transp : ∀ M E,
146 | is_rotation_matrix M
147 | → (set_map (mat_vec_mul (mat_transp M)) (set_map (mat_vec_mul M) E) = E)%S.
148 | Proof.
149 | intros * Hrm.
150 | intros p; simpl.
151 | split; intros H. {
152 | destruct H as (u & (v & Hv & Hvu) & Hu).
153 | rewrite <- Hvu, <- mat_vec_mul_assoc in Hu.
154 | rewrite rotation_mat_mul_transp_l in Hu; [ | easy ].
155 | rewrite mat_vec_mul_id in Hu.
156 | now rewrite <- Hu.
157 | } {
158 | exists (M * p)%vec.
159 | split; [ now exists p; split | ].
160 | rewrite <- mat_vec_mul_assoc.
161 | rewrite rotation_mat_mul_transp_l; [ | easy ].
162 | apply mat_vec_mul_id.
163 | }
164 | Qed.
165 |
166 | Theorem app_gr_inv_l : ∀ g E,
167 | (app_gr_inv g (app_gr g E) = E)%S.
168 | Proof.
169 | intros.
170 | unfold app_gr_inv.
171 | revert E.
172 | induction g as [ M Hrm | | ]; intros; simpl. {
173 | now apply set_map_mul_transp.
174 | } {
175 | intros (x, y, z); simpl.
176 | rewrite neg_vec_involutive.
177 | destruct v as (xv, yv, zv); simpl.
178 | now do 3 rewrite Rplus_assoc, Rplus_opp_r, Rplus_0_r.
179 | } {
180 | intros p.
181 | split; intros H.
182 | rewrite IHg1 in H; now apply IHg2.
183 | rewrite IHg1; apply IHg2, H.
184 | }
185 | Qed.
186 |
187 | Theorem app_gr_inv_r : ∀ g E,
188 | (app_gr g (app_gr_inv g E) = E)%S.
189 | Proof.
190 | intros.
191 | unfold app_gr_inv.
192 | revert E.
193 | induction g as [ M Hrm | | ]; intros; simpl. {
194 | apply set_map_mul_transp.
195 | now apply rotation_transp_is_rotation.
196 | } {
197 | intros (x, y, z); simpl.
198 | rewrite vec_sub_opp_r.
199 | destruct v as (xv, yv, zv); simpl.
200 | now do 3 rewrite Rplus_assoc, Rplus_opp_l, Rplus_0_r.
201 | } {
202 | intros p.
203 | split; intros H. {
204 | rewrite IHg2 in H; now apply IHg1.
205 | } {
206 | rewrite IHg2; apply IHg1, H.
207 | }
208 | }
209 | Qed.
210 |
211 | Theorem app_gr_app_gr_vec : ∀ g E p, p ∈ app_gr g E → app_gr_vec g p ∈ E.
212 | Proof.
213 | intros * Hp.
214 | revert E p Hp.
215 | induction g as [ M Hrm | | ]; intros. {
216 | simpl in Hp; simpl.
217 | destruct Hp as (u & Hu & Hmu).
218 | rewrite <- Hmu, <- mat_vec_mul_assoc.
219 | rewrite rotation_mat_mul_transp_l; [ | easy ].
220 | now rewrite mat_vec_mul_id.
221 | } {
222 | now destruct p.
223 | } {
224 | simpl in Hp; simpl.
225 | apply IHg1 in Hp.
226 | now apply IHg2 in Hp.
227 | }
228 | Qed.
229 |
230 | Theorem app_gr_empty_set : ∀ f, (app_gr f ∅ = ∅)%S.
231 | Proof.
232 | intros * p.
233 | split; intros H; [ | easy ].
234 | revert p H.
235 | induction f as [M Hrm | | ]; intros. {
236 | now destruct H.
237 | } {
238 | now destruct p.
239 | }
240 | simpl in H.
241 | eapply gr_subst in H; [ now apply IHf1 in H | ].
242 | split; [ apply IHf2 | easy ].
243 | Qed.
244 |
245 | Theorem app_gr_app_gr_inv : ∀ E g,
246 | (app_gr g (app_gr_inv g E) = E)%S.
247 | Proof.
248 | intros.
249 | unfold app_gr_inv.
250 | revert E.
251 | induction g as [M Hrm | | ]; intros. {
252 | now apply set_map_mul_transp, rotation_transp_is_rotation.
253 | } {
254 | intros (x, y, z); simpl.
255 | rewrite vec_sub_opp_r.
256 | destruct v as (xv, yv, zv); simpl.
257 | now do 3 rewrite Rplus_assoc, Rplus_opp_l, Rplus_0_r.
258 | }
259 | simpl.
260 | rewrite IHg2.
261 | apply IHg1.
262 | Qed.
263 |
264 | Theorem app_gr_inv_app_gr : ∀ E g,
265 | (app_gr_inv g (app_gr g E) = E)%S.
266 | Proof.
267 | intros.
268 | unfold app_gr_inv.
269 | revert E.
270 | induction g as [ M Hrm | | ]; intros. {
271 | now apply set_map_mul_transp.
272 | } {
273 | intros (x, y, z); simpl.
274 | rewrite neg_vec_involutive.
275 | destruct v as (xv, yv, zv); simpl.
276 | now do 3 rewrite Rplus_assoc, Rplus_opp_r, Rplus_0_r.
277 | }
278 | simpl.
279 | rewrite IHg1.
280 | apply IHg2.
281 | Qed.
282 |
283 | Theorem group_set_inter_distr : ∀ g E F,
284 | (app_gr g (E ∩ F) = app_gr g E ∩ app_gr g F)%S.
285 | Proof.
286 | intros.
287 | revert E F.
288 | induction g; intros. {
289 | apply set_map_inter_distr.
290 | now apply rot_mat_vec_mul_is_inj.
291 | } {
292 | easy.
293 | }
294 | intros p; simpl; now rewrite IHg2, IHg1.
295 | Qed.
296 |
297 | Theorem group_set_union_distr : ∀ g E F,
298 | (app_gr g (E ∪ F) = app_gr g E ∪ app_gr g F)%S.
299 | Proof.
300 | intros.
301 | revert E F.
302 | induction g; intros. {
303 | apply set_map_union_distr.
304 | } {
305 | easy.
306 | }
307 | now intros p; simpl; rewrite IHg2, IHg1.
308 | Qed.
309 |
310 | Theorem group_set_union_list_distr : ∀ f EL,
311 | (app_gr f (⋃ EL) = ⋃ map (app_gr f) EL)%S.
312 | Proof.
313 | intros.
314 | induction EL as [| E₁ EL]. {
315 | intros x; rewrite app_gr_empty_set; now split.
316 | }
317 | intros x; simpl.
318 | rewrite group_set_union_distr.
319 | split; intros Hx. {
320 | destruct Hx as [Hx| Hx]; [ now left | ].
321 | right; now apply IHEL.
322 | } {
323 | destruct Hx as [Hx| Hx]; [ now left | ].
324 | right; now apply IHEL.
325 | }
326 | Qed.
327 |
328 | Theorem included_group : ∀ E F g, E ⊂ F ↔ app_gr g E ⊂ app_gr g F.
329 | Proof.
330 | intros.
331 | split; intros HEF. {
332 | revert E F HEF.
333 | induction g as [M Hrm| dx| ]; intros. {
334 | intros u (v & Hv & Hm).
335 | now exists v; split; [ apply HEF | ].
336 | } {
337 | intros (x, y, z) Hp; now apply HEF.
338 | }
339 | now apply IHg1, IHg2.
340 | }
341 | intros p Hp.
342 | revert p E F HEF Hp.
343 | induction g as [M Hrm| d| ]; intros. {
344 | simpl in HEF.
345 | apply in_set_map with (f := mat_vec_mul M) in Hp.
346 | specialize (HEF (mat_vec_mul M p) Hp).
347 | apply set_map_in in HEF; [ easy | ].
348 | now apply rot_mat_vec_mul_is_inj.
349 | } {
350 | destruct p as (x, y, z); simpl in HEF.
351 | destruct d as (dx, dy, dz).
352 | pose proof HEF (V (x + dx) (y + dy) (z + dz)) as H; simpl in H.
353 | unfold Rminus in H.
354 | do 3 rewrite Rplus_assoc, Rplus_opp_r, Rplus_0_r in H.
355 | apply H, Hp.
356 | }
357 | simpl in HEF.
358 | eapply IHg2; [ | eassumption ].
359 | intros q Hq; eapply IHg1; eassumption.
360 | Qed.
361 |
362 | Theorem partition_group_map : ∀ (F : set vector) EL g,
363 | is_partition F EL → is_partition (app_gr g F) (map (app_gr g) EL).
364 | Proof.
365 | intros F EL * HP.
366 | unfold is_partition in HP |-*.
367 | destruct HP as (HF, HP).
368 | split. {
369 | induction g as [M Hrm| dx | g IHg h IHh ]; intros; simpl. {
370 | split. {
371 | intros Hr.
372 | revert F HF Hr.
373 | induction EL as [| E EL]; intros. {
374 | rewrite HF in Hr.
375 | now destruct Hr.
376 | }
377 | simpl in HF; simpl.
378 | generalize Hr; intros H.
379 | rewrite HF in H.
380 | rewrite set_map_union_distr in H.
381 | destruct H as [H| H]; [ now left | right ].
382 | eapply IHEL; [ | easy | eassumption ].
383 | intros i j Hij.
384 | unfold set_eq; simpl; intros y.
385 | assert (HSij : S i ≠ S j). {
386 | intros HSij; now apply Hij, Nat.succ_inj.
387 | }
388 | pose proof HP (S i) (S j) HSij y as HP; simpl in HP.
389 | destruct HP as (HQ, _).
390 | split; [ intros (HPi, HPj) | easy ].
391 | apply HQ; now split.
392 | }
393 | intros Hme.
394 | revert F HF.
395 | induction EL as [| E EL]; intros; [ easy | ].
396 | simpl in Hme; rewrite HF.
397 | destruct Hme as [(u, Hme)| Hme]. {
398 | now exists u; split; [ left | ].
399 | }
400 | rewrite set_map_union_list_distr; right.
401 | rewrite <- set_map_union_list_distr.
402 | apply IHEL; [ | easy | intros y; split; intros H; apply H ].
403 | intros i j Hij y.
404 | assert (HSij : S i ≠ S j). {
405 | intros HSij; now apply Hij, Nat.succ_inj.
406 | }
407 | pose proof HP (S i) (S j) HSij y as HP; simpl in HP.
408 | destruct HP as (HQ, _).
409 | split; [ intros (HPi, HPj) | easy ].
410 | apply HQ; now split.
411 | }
412 | intros (x, y, z).
413 | split. {
414 | intros Hp.
415 | revert F HF Hp.
416 | induction EL as [| E EL]; intros. {
417 | unfold set_eq in HF; simpl in HF.
418 | now apply HF in Hp.
419 | }
420 | simpl in HF; simpl.
421 | generalize Hp; intros H.
422 | apply HF in H; simpl in H.
423 | destruct H as [H| H]; [ now left | right ].
424 | eapply IHEL; [ | easy | eassumption ].
425 | intros i j Hij.
426 | unfold set_eq; simpl; intros q.
427 | assert (HSij : S i ≠ S j). {
428 | intros HSij; now apply Hij, Nat.succ_inj.
429 | }
430 | pose proof HP (S i) (S j) HSij q as HP; simpl in HP.
431 | destruct HP as (HQ, _).
432 | split; [ intros (HPi, HPj) | easy ].
433 | apply HQ; now split.
434 | } {
435 | intros Hme.
436 | revert F HF.
437 | induction EL as [| E EL]; intros; [ easy | ].
438 | simpl in HF, Hme; apply HF.
439 | destruct Hme as [Hme| Hme]; [ now left | ].
440 | right; simpl.
441 | apply IHEL; [ | easy | intros q; split; intros H; apply H ].
442 | intros i j Hij q.
443 | assert (HSij : S i ≠ S j). {
444 | intros HSij; now apply Hij, Nat.succ_inj.
445 | }
446 | pose proof HP (S i) (S j) HSij q as HP; simpl in HP.
447 | destruct HP as (HQ, _).
448 | split; [ intros (HPi, HPj) | easy ].
449 | apply HQ; now split.
450 | }
451 | intros p.
452 | split. {
453 | intros Hgh.
454 | revert F HF IHg IHh Hgh.
455 | induction EL as [| E EL]; intros. {
456 | rewrite IHh in Hgh; simpl in Hgh.
457 | eapply app_gr_empty_set, Hgh.
458 | }
459 | rewrite IHh in Hgh.
460 | simpl in Hgh.
461 | apply group_set_union_distr in Hgh.
462 | destruct Hgh as [Hgh| Hgh]; [ now left | right ].
463 | eapply IHEL; [ | easy | | | ]. {
464 | intros i j Hij.
465 | unfold set_eq; simpl; intros y.
466 | assert (HSij : S i ≠ S j). {
467 | intros HSij; now apply Hij, Nat.succ_inj.
468 | }
469 | pose proof HP (S i) (S j) HSij y as HP; simpl in HP.
470 | destruct HP as (HQ, _).
471 | split; [ intros (HPi, HPj) | easy ].
472 | apply HQ; now split.
473 | } {
474 | apply group_set_union_list_distr.
475 | } {
476 | apply group_set_union_list_distr.
477 | }
478 | pose proof group_set_union_list_distr h EL.
479 | now rewrite <- H in Hgh.
480 | }
481 | intros Hgh.
482 | revert F HF IHg IHh Hgh.
483 | induction EL as [| E ELl]; intros; [ easy | ].
484 | destruct Hgh as [Hgh| Hgh]. {
485 | rewrite IHh; simpl.
486 | rewrite set_eq_equiv; [ | now rewrite group_set_union_distr ].
487 | now left.
488 | }
489 | rewrite HF; simpl.
490 | rewrite set_eq_equiv; [ | now rewrite group_set_union_distr ].
491 | rewrite set_eq_equiv; [ | now rewrite group_set_union_distr ].
492 | right.
493 | rewrite group_set_union_list_distr.
494 | rewrite set_eq_equiv; [ | now rewrite group_set_union_list_distr ].
495 | now rewrite map_map.
496 | }
497 | intros i j Hij p.
498 | split; intros H; [ | easy ].
499 | rewrite <- app_gr_empty_set with (f := g) in H.
500 | do 2 rewrite map_nth in H.
501 | destruct H as (Hi, Hj).
502 | pose proof HP i j Hij (app_gr_vec g p) as Hp.
503 | destruct Hp as (Hpi, _).
504 | apply Hpi; clear Hpi.
505 | split. {
506 | clear - Hi.
507 | revert p EL Hi.
508 | induction i; intros. {
509 | destruct EL as [| E EL]; [ now apply app_gr_empty_set in Hi | ].
510 | simpl in Hi; simpl.
511 | now apply app_gr_app_gr_vec.
512 | }
513 | destruct EL as [| E EL]; [ now apply app_gr_empty_set in Hi | ].
514 | simpl in Hi; simpl.
515 | now apply IHi.
516 | }
517 | clear - Hj.
518 | revert p EL Hj.
519 | induction j; intros. {
520 | destruct EL as [| E EL]; [ now apply app_gr_empty_set in Hj | ].
521 | simpl in Hj; simpl.
522 | now apply app_gr_app_gr_vec.
523 | }
524 | destruct EL as [| E EL]; [ now apply app_gr_empty_set in Hj | ].
525 | simpl in Hj; simpl.
526 | now apply IHj.
527 | Qed.
528 |
--------------------------------------------------------------------------------
/LICENSE:
--------------------------------------------------------------------------------
1 | CeCILL-B FREE SOFTWARE LICENSE AGREEMENT
2 |
3 |
4 | Notice
5 |
6 | This Agreement is a Free Software license agreement that is the result
7 | of discussions between its authors in order to ensure compliance with
8 | the two main principles guiding its drafting:
9 |
10 | * firstly, compliance with the principles governing the distribution
11 | of Free Software: access to source code, broad rights granted to
12 | users,
13 | * secondly, the election of a governing law, French law, with which
14 | it is conformant, both as regards the law of torts and
15 | intellectual property law, and the protection that it offers to
16 | both authors and holders of the economic rights over software.
17 |
18 | The authors of the CeCILL-B (for Ce[a] C[nrs] I[nria] L[ogiciel] L[ibre])
19 | license are:
20 |
21 | Commissariat à l'Energie Atomique - CEA, a public scientific, technical
22 | and industrial research establishment, having its principal place of
23 | business at 25 rue Leblanc, immeuble Le Ponant D, 75015 Paris, France.
24 |
25 | Centre National de la Recherche Scientifique - CNRS, a public scientific
26 | and technological establishment, having its principal place of business
27 | at 3 rue Michel-Ange, 75794 Paris cedex 16, France.
28 |
29 | Institut National de Recherche en Informatique et en Automatique -
30 | INRIA, a public scientific and technological establishment, having its
31 | principal place of business at Domaine de Voluceau, Rocquencourt, BP
32 | 105, 78153 Le Chesnay cedex, France.
33 |
34 |
35 | Preamble
36 |
37 | This Agreement is an open source software license intended to give users
38 | significant freedom to modify and redistribute the software licensed
39 | hereunder.
40 |
41 | The exercising of this freedom is conditional upon a strong obligation
42 | of giving credits for everybody that distributes a software
43 | incorporating a software ruled by the current license so as all
44 | contributions to be properly identified and acknowledged.
45 |
46 | In consideration of access to the source code and the rights to copy,
47 | modify and redistribute granted by the license, users are provided only
48 | with a limited warranty and the software's author, the holder of the
49 | economic rights, and the successive licensors only have limited liability.
50 |
51 | In this respect, the risks associated with loading, using, modifying
52 | and/or developing or reproducing the software by the user are brought to
53 | the user's attention, given its Free Software status, which may make it
54 | complicated to use, with the result that its use is reserved for
55 | developers and experienced professionals having in-depth computer
56 | knowledge. Users are therefore encouraged to load and test the
57 | suitability of the software as regards their requirements in conditions
58 | enabling the security of their systems and/or data to be ensured and,
59 | more generally, to use and operate it in the same conditions of
60 | security. This Agreement may be freely reproduced and published,
61 | provided it is not altered, and that no provisions are either added or
62 | removed herefrom.
63 |
64 | This Agreement may apply to any or all software for which the holder of
65 | the economic rights decides to submit the use thereof to its provisions.
66 |
67 |
68 | Article 1 - DEFINITIONS
69 |
70 | For the purpose of this Agreement, when the following expressions
71 | commence with a capital letter, they shall have the following meaning:
72 |
73 | Agreement: means this license agreement, and its possible subsequent
74 | versions and annexes.
75 |
76 | Software: means the software in its Object Code and/or Source Code form
77 | and, where applicable, its documentation, "as is" when the Licensee
78 | accepts the Agreement.
79 |
80 | Initial Software: means the Software in its Source Code and possibly its
81 | Object Code form and, where applicable, its documentation, "as is" when
82 | it is first distributed under the terms and conditions of the Agreement.
83 |
84 | Modified Software: means the Software modified by at least one
85 | Contribution.
86 |
87 | Source Code: means all the Software's instructions and program lines to
88 | which access is required so as to modify the Software.
89 |
90 | Object Code: means the binary files originating from the compilation of
91 | the Source Code.
92 |
93 | Holder: means the holder(s) of the economic rights over the Initial
94 | Software.
95 |
96 | Licensee: means the Software user(s) having accepted the Agreement.
97 |
98 | Contributor: means a Licensee having made at least one Contribution.
99 |
100 | Licensor: means the Holder, or any other individual or legal entity, who
101 | distributes the Software under the Agreement.
102 |
103 | Contribution: means any or all modifications, corrections, translations,
104 | adaptations and/or new functions integrated into the Software by any or
105 | all Contributors, as well as any or all Internal Modules.
106 |
107 | Module: means a set of sources files including their documentation that
108 | enables supplementary functions or services in addition to those offered
109 | by the Software.
110 |
111 | External Module: means any or all Modules, not derived from the
112 | Software, so that this Module and the Software run in separate address
113 | spaces, with one calling the other when they are run.
114 |
115 | Internal Module: means any or all Module, connected to the Software so
116 | that they both execute in the same address space.
117 |
118 | Parties: mean both the Licensee and the Licensor.
119 |
120 | These expressions may be used both in singular and plural form.
121 |
122 |
123 | Article 2 - PURPOSE
124 |
125 | The purpose of the Agreement is the grant by the Licensor to the
126 | Licensee of a non-exclusive, transferable and worldwide license for the
127 | Software as set forth in Article 5 hereinafter for the whole term of the
128 | protection granted by the rights over said Software.
129 |
130 |
131 | Article 3 - ACCEPTANCE
132 |
133 | 3.1 The Licensee shall be deemed as having accepted the terms and
134 | conditions of this Agreement upon the occurrence of the first of the
135 | following events:
136 |
137 | * (i) loading the Software by any or all means, notably, by
138 | downloading from a remote server, or by loading from a physical
139 | medium;
140 | * (ii) the first time the Licensee exercises any of the rights
141 | granted hereunder.
142 |
143 | 3.2 One copy of the Agreement, containing a notice relating to the
144 | characteristics of the Software, to the limited warranty, and to the
145 | fact that its use is restricted to experienced users has been provided
146 | to the Licensee prior to its acceptance as set forth in Article 3.1
147 | hereinabove, and the Licensee hereby acknowledges that it has read and
148 | understood it.
149 |
150 |
151 | Article 4 - EFFECTIVE DATE AND TERM
152 |
153 |
154 | 4.1 EFFECTIVE DATE
155 |
156 | The Agreement shall become effective on the date when it is accepted by
157 | the Licensee as set forth in Article 3.1.
158 |
159 |
160 | 4.2 TERM
161 |
162 | The Agreement shall remain in force for the entire legal term of
163 | protection of the economic rights over the Software.
164 |
165 |
166 | Article 5 - SCOPE OF RIGHTS GRANTED
167 |
168 | The Licensor hereby grants to the Licensee, who accepts, the following
169 | rights over the Software for any or all use, and for the term of the
170 | Agreement, on the basis of the terms and conditions set forth hereinafter.
171 |
172 | Besides, if the Licensor owns or comes to own one or more patents
173 | protecting all or part of the functions of the Software or of its
174 | components, the Licensor undertakes not to enforce the rights granted by
175 | these patents against successive Licensees using, exploiting or
176 | modifying the Software. If these patents are transferred, the Licensor
177 | undertakes to have the transferees subscribe to the obligations set
178 | forth in this paragraph.
179 |
180 |
181 | 5.1 RIGHT OF USE
182 |
183 | The Licensee is authorized to use the Software, without any limitation
184 | as to its fields of application, with it being hereinafter specified
185 | that this comprises:
186 |
187 | 1. permanent or temporary reproduction of all or part of the Software
188 | by any or all means and in any or all form.
189 |
190 | 2. loading, displaying, running, or storing the Software on any or
191 | all medium.
192 |
193 | 3. entitlement to observe, study or test its operation so as to
194 | determine the ideas and principles behind any or all constituent
195 | elements of said Software. This shall apply when the Licensee
196 | carries out any or all loading, displaying, running, transmission
197 | or storage operation as regards the Software, that it is entitled
198 | to carry out hereunder.
199 |
200 |
201 | 5.2 ENTITLEMENT TO MAKE CONTRIBUTIONS
202 |
203 | The right to make Contributions includes the right to translate, adapt,
204 | arrange, or make any or all modifications to the Software, and the right
205 | to reproduce the resulting software.
206 |
207 | The Licensee is authorized to make any or all Contributions to the
208 | Software provided that it includes an explicit notice that it is the
209 | author of said Contribution and indicates the date of the creation thereof.
210 |
211 |
212 | 5.3 RIGHT OF DISTRIBUTION
213 |
214 | In particular, the right of distribution includes the right to publish,
215 | transmit and communicate the Software to the general public on any or
216 | all medium, and by any or all means, and the right to market, either in
217 | consideration of a fee, or free of charge, one or more copies of the
218 | Software by any means.
219 |
220 | The Licensee is further authorized to distribute copies of the modified
221 | or unmodified Software to third parties according to the terms and
222 | conditions set forth hereinafter.
223 |
224 |
225 | 5.3.1 DISTRIBUTION OF SOFTWARE WITHOUT MODIFICATION
226 |
227 | The Licensee is authorized to distribute true copies of the Software in
228 | Source Code or Object Code form, provided that said distribution
229 | complies with all the provisions of the Agreement and is accompanied by:
230 |
231 | 1. a copy of the Agreement,
232 |
233 | 2. a notice relating to the limitation of both the Licensor's
234 | warranty and liability as set forth in Articles 8 and 9,
235 |
236 | and that, in the event that only the Object Code of the Software is
237 | redistributed, the Licensee allows effective access to the full Source
238 | Code of the Software at a minimum during the entire period of its
239 | distribution of the Software, it being understood that the additional
240 | cost of acquiring the Source Code shall not exceed the cost of
241 | transferring the data.
242 |
243 |
244 | 5.3.2 DISTRIBUTION OF MODIFIED SOFTWARE
245 |
246 | If the Licensee makes any Contribution to the Software, the resulting
247 | Modified Software may be distributed under a license agreement other
248 | than this Agreement subject to compliance with the provisions of Article
249 | 5.3.4.
250 |
251 |
252 | 5.3.3 DISTRIBUTION OF EXTERNAL MODULES
253 |
254 | When the Licensee has developed an External Module, the terms and
255 | conditions of this Agreement do not apply to said External Module, that
256 | may be distributed under a separate license agreement.
257 |
258 |
259 | 5.3.4 CREDITS
260 |
261 | Any Licensee who may distribute a Modified Software hereby expressly
262 | agrees to:
263 |
264 | 1. indicate in the related documentation that it is based on the
265 | Software licensed hereunder, and reproduce the intellectual
266 | property notice for the Software,
267 |
268 | 2. ensure that written indications of the Software intended use,
269 | intellectual property notice and license hereunder are included in
270 | easily accessible format from the Modified Software interface,
271 |
272 | 3. mention, on a freely accessible website describing the Modified
273 | Software, at least throughout the distribution term thereof, that
274 | it is based on the Software licensed hereunder, and reproduce the
275 | Software intellectual property notice,
276 |
277 | 4. where it is distributed to a third party that may distribute a
278 | Modified Software without having to make its source code
279 | available, make its best efforts to ensure that said third party
280 | agrees to comply with the obligations set forth in this Article .
281 |
282 | If the Software, whether or not modified, is distributed with an
283 | External Module designed for use in connection with the Software, the
284 | Licensee shall submit said External Module to the foregoing obligations.
285 |
286 |
287 | 5.3.5 COMPATIBILITY WITH THE CeCILL AND CeCILL-C LICENSES
288 |
289 | Where a Modified Software contains a Contribution subject to the CeCILL
290 | license, the provisions set forth in Article 5.3.4 shall be optional.
291 |
292 | A Modified Software may be distributed under the CeCILL-C license. In
293 | such a case the provisions set forth in Article 5.3.4 shall be optional.
294 |
295 |
296 | Article 6 - INTELLECTUAL PROPERTY
297 |
298 |
299 | 6.1 OVER THE INITIAL SOFTWARE
300 |
301 | The Holder owns the economic rights over the Initial Software. Any or
302 | all use of the Initial Software is subject to compliance with the terms
303 | and conditions under which the Holder has elected to distribute its work
304 | and no one shall be entitled to modify the terms and conditions for the
305 | distribution of said Initial Software.
306 |
307 | The Holder undertakes that the Initial Software will remain ruled at
308 | least by this Agreement, for the duration set forth in Article 4.2.
309 |
310 |
311 | 6.2 OVER THE CONTRIBUTIONS
312 |
313 | The Licensee who develops a Contribution is the owner of the
314 | intellectual property rights over this Contribution as defined by
315 | applicable law.
316 |
317 |
318 | 6.3 OVER THE EXTERNAL MODULES
319 |
320 | The Licensee who develops an External Module is the owner of the
321 | intellectual property rights over this External Module as defined by
322 | applicable law and is free to choose the type of agreement that shall
323 | govern its distribution.
324 |
325 |
326 | 6.4 JOINT PROVISIONS
327 |
328 | The Licensee expressly undertakes:
329 |
330 | 1. not to remove, or modify, in any manner, the intellectual property
331 | notices attached to the Software;
332 |
333 | 2. to reproduce said notices, in an identical manner, in the copies
334 | of the Software modified or not.
335 |
336 | The Licensee undertakes not to directly or indirectly infringe the
337 | intellectual property rights of the Holder and/or Contributors on the
338 | Software and to take, where applicable, vis-à-vis its staff, any and all
339 | measures required to ensure respect of said intellectual property rights
340 | of the Holder and/or Contributors.
341 |
342 |
343 | Article 7 - RELATED SERVICES
344 |
345 | 7.1 Under no circumstances shall the Agreement oblige the Licensor to
346 | provide technical assistance or maintenance services for the Software.
347 |
348 | However, the Licensor is entitled to offer this type of services. The
349 | terms and conditions of such technical assistance, and/or such
350 | maintenance, shall be set forth in a separate instrument. Only the
351 | Licensor offering said maintenance and/or technical assistance services
352 | shall incur liability therefor.
353 |
354 | 7.2 Similarly, any Licensor is entitled to offer to its licensees, under
355 | its sole responsibility, a warranty, that shall only be binding upon
356 | itself, for the redistribution of the Software and/or the Modified
357 | Software, under terms and conditions that it is free to decide. Said
358 | warranty, and the financial terms and conditions of its application,
359 | shall be subject of a separate instrument executed between the Licensor
360 | and the Licensee.
361 |
362 |
363 | Article 8 - LIABILITY
364 |
365 | 8.1 Subject to the provisions of Article 8.2, the Licensee shall be
366 | entitled to claim compensation for any direct loss it may have suffered
367 | from the Software as a result of a fault on the part of the relevant
368 | Licensor, subject to providing evidence thereof.
369 |
370 | 8.2 The Licensor's liability is limited to the commitments made under
371 | this Agreement and shall not be incurred as a result of in particular:
372 | (i) loss due the Licensee's total or partial failure to fulfill its
373 | obligations, (ii) direct or consequential loss that is suffered by the
374 | Licensee due to the use or performance of the Software, and (iii) more
375 | generally, any consequential loss. In particular the Parties expressly
376 | agree that any or all pecuniary or business loss (i.e. loss of data,
377 | loss of profits, operating loss, loss of customers or orders,
378 | opportunity cost, any disturbance to business activities) or any or all
379 | legal proceedings instituted against the Licensee by a third party,
380 | shall constitute consequential loss and shall not provide entitlement to
381 | any or all compensation from the Licensor.
382 |
383 |
384 | Article 9 - WARRANTY
385 |
386 | 9.1 The Licensee acknowledges that the scientific and technical
387 | state-of-the-art when the Software was distributed did not enable all
388 | possible uses to be tested and verified, nor for the presence of
389 | possible defects to be detected. In this respect, the Licensee's
390 | attention has been drawn to the risks associated with loading, using,
391 | modifying and/or developing and reproducing the Software which are
392 | reserved for experienced users.
393 |
394 | The Licensee shall be responsible for verifying, by any or all means,
395 | the suitability of the product for its requirements, its good working
396 | order, and for ensuring that it shall not cause damage to either persons
397 | or properties.
398 |
399 | 9.2 The Licensor hereby represents, in good faith, that it is entitled
400 | to grant all the rights over the Software (including in particular the
401 | rights set forth in Article 5).
402 |
403 | 9.3 The Licensee acknowledges that the Software is supplied "as is" by
404 | the Licensor without any other express or tacit warranty, other than
405 | that provided for in Article 9.2 and, in particular, without any warranty
406 | as to its commercial value, its secured, safe, innovative or relevant
407 | nature.
408 |
409 | Specifically, the Licensor does not warrant that the Software is free
410 | from any error, that it will operate without interruption, that it will
411 | be compatible with the Licensee's own equipment and software
412 | configuration, nor that it will meet the Licensee's requirements.
413 |
414 | 9.4 The Licensor does not either expressly or tacitly warrant that the
415 | Software does not infringe any third party intellectual property right
416 | relating to a patent, software or any other property right. Therefore,
417 | the Licensor disclaims any and all liability towards the Licensee
418 | arising out of any or all proceedings for infringement that may be
419 | instituted in respect of the use, modification and redistribution of the
420 | Software. Nevertheless, should such proceedings be instituted against
421 | the Licensee, the Licensor shall provide it with technical and legal
422 | assistance for its defense. Such technical and legal assistance shall be
423 | decided on a case-by-case basis between the relevant Licensor and the
424 | Licensee pursuant to a memorandum of understanding. The Licensor
425 | disclaims any and all liability as regards the Licensee's use of the
426 | name of the Software. No warranty is given as regards the existence of
427 | prior rights over the name of the Software or as regards the existence
428 | of a trademark.
429 |
430 |
431 | Article 10 - TERMINATION
432 |
433 | 10.1 In the event of a breach by the Licensee of its obligations
434 | hereunder, the Licensor may automatically terminate this Agreement
435 | thirty (30) days after notice has been sent to the Licensee and has
436 | remained ineffective.
437 |
438 | 10.2 A Licensee whose Agreement is terminated shall no longer be
439 | authorized to use, modify or distribute the Software. However, any
440 | licenses that it may have granted prior to termination of the Agreement
441 | shall remain valid subject to their having been granted in compliance
442 | with the terms and conditions hereof.
443 |
444 |
445 | Article 11 - MISCELLANEOUS
446 |
447 |
448 | 11.1 EXCUSABLE EVENTS
449 |
450 | Neither Party shall be liable for any or all delay, or failure to
451 | perform the Agreement, that may be attributable to an event of force
452 | majeure, an act of God or an outside cause, such as defective
453 | functioning or interruptions of the electricity or telecommunications
454 | networks, network paralysis following a virus attack, intervention by
455 | government authorities, natural disasters, water damage, earthquakes,
456 | fire, explosions, strikes and labor unrest, war, etc.
457 |
458 | 11.2 Any failure by either Party, on one or more occasions, to invoke
459 | one or more of the provisions hereof, shall under no circumstances be
460 | interpreted as being a waiver by the interested Party of its right to
461 | invoke said provision(s) subsequently.
462 |
463 | 11.3 The Agreement cancels and replaces any or all previous agreements,
464 | whether written or oral, between the Parties and having the same
465 | purpose, and constitutes the entirety of the agreement between said
466 | Parties concerning said purpose. No supplement or modification to the
467 | terms and conditions hereof shall be effective as between the Parties
468 | unless it is made in writing and signed by their duly authorized
469 | representatives.
470 |
471 | 11.4 In the event that one or more of the provisions hereof were to
472 | conflict with a current or future applicable act or legislative text,
473 | said act or legislative text shall prevail, and the Parties shall make
474 | the necessary amendments so as to comply with said act or legislative
475 | text. All other provisions shall remain effective. Similarly, invalidity
476 | of a provision of the Agreement, for any reason whatsoever, shall not
477 | cause the Agreement as a whole to be invalid.
478 |
479 |
480 | 11.5 LANGUAGE
481 |
482 | The Agreement is drafted in both French and English and both versions
483 | are deemed authentic.
484 |
485 |
486 | Article 12 - NEW VERSIONS OF THE AGREEMENT
487 |
488 | 12.1 Any person is authorized to duplicate and distribute copies of this
489 | Agreement.
490 |
491 | 12.2 So as to ensure coherence, the wording of this Agreement is
492 | protected and may only be modified by the authors of the License, who
493 | reserve the right to periodically publish updates or new versions of the
494 | Agreement, each with a separate number. These subsequent versions may
495 | address new issues encountered by Free Software.
496 |
497 | 12.3 Any Software distributed under a given version of the Agreement may
498 | only be subsequently distributed under the same version of the Agreement
499 | or a subsequent version.
500 |
501 |
502 | Article 13 - GOVERNING LAW AND JURISDICTION
503 |
504 | 13.1 The Agreement is governed by French law. The Parties agree to
505 | endeavor to seek an amicable solution to any disagreements or disputes
506 | that may arise during the performance of the Agreement.
507 |
508 | 13.2 Failing an amicable solution within two (2) months as from their
509 | occurrence, and unless emergency proceedings are necessary, the
510 | disagreements or disputes shall be referred to the Paris Courts having
511 | jurisdiction, by the more diligent Party.
512 |
513 |
514 | Version 1.0 dated 2006-09-05.
515 |
--------------------------------------------------------------------------------
/Makefile:
--------------------------------------------------------------------------------
1 | TARGET=Banach_Tarski.vo
2 | COQ_VERSION=`coqc -v | grep version | sed -e 's/^.*version //;s/[ ~].*$$//;s/\./_/g;s/^/COQ_/;s/+.*$$//'`
3 |
4 | all: $(TARGET)
5 |
6 | FILESFORDEP=`LC_ALL=C ls *.v`
7 |
8 | clean:
9 | rm -f *.glob *.vo .*.aux
10 | rm -f *.vok *.vos
11 | rm -f *.cm[iox] *.o *.cmxs *.native
12 | rm -f .*.cache
13 |
14 | depend:
15 | mv .depend .depend.bak
16 | rocq dep -Q . . $(FILESFORDEP) | sed -e " s|$$HOME[^ ]*||" | \
17 | LC_ALL=C sort | sed -e 's/ *$$//' > .depend
18 |
19 | show_coq_version:
20 | @echo $(COQ_VERSION)
21 |
22 | .SUFFIXES: .v .vo .vp
23 |
24 | %.vo: %.v
25 | rocq compile $<
26 |
27 | %.v: %.vp
28 | @echo /lib/cpp -D$(COQ_VERSION) $< '>' $@
29 | @/bin/rm -f $@
30 | @sed -e 's|//|slsl|g' $< | \
31 | /lib/cpp -D$(COQ_VERSION) 2>/dev/null | \
32 | sed -e 's|slsl|//|g' | \
33 | grep -v '^#' > $@
34 |
35 | .PHONY: all clean depend
36 |
37 | include .depend
38 |
--------------------------------------------------------------------------------
/Misc.v:
--------------------------------------------------------------------------------
1 | (* Banach-Tarski paradox. *)
2 |
3 | From Stdlib Require Import Utf8 Arith List Relations.
4 | Import ListNotations.
5 |
6 | Arguments Nat.div : simpl never.
7 | Arguments Nat.modulo : simpl never.
8 |
9 | Theorem match_id : ∀ A a (b : A), match a with O => b | S _ => b end = b.
10 | Proof. intros A a b; now destruct a. Qed.
11 |
12 | Theorem nat_add_diag_mul_2 : ∀ n, (n + n = 2 * n)%nat.
13 | Proof.
14 | intros.
15 | destruct n; [ easy | now simpl; rewrite Nat.add_0_r ].
16 | Qed.
17 |
18 | Theorem nat_div_add_once : ∀ a b, b ≠ 0 → (a + b) / b = S (a / b).
19 | Proof.
20 | intros a b Hb.
21 | replace b with (1 * b) at 1 by apply Nat.mul_1_l.
22 | rewrite Nat.div_add; [ apply Nat.add_1_r | easy ].
23 | Qed.
24 |
25 | Theorem nat_mod_add_once : ∀ a b, (a + b) mod b = a mod b.
26 | Proof.
27 | intros a b.
28 | replace b with (1 * b) at 1 by apply Nat.mul_1_l.
29 | apply Nat.Div0.mod_add.
30 | Qed.
31 |
32 | Theorem neq_negb : ∀ x y, x ≠ y → x = negb y.
33 | Proof.
34 | intros.
35 | destruct x, y; try easy; exfalso; now apply H.
36 | Qed.
37 |
38 | Theorem negb_neq : ∀ b₁ b₂, negb b₁ ≠ b₂ → b₁ = b₂.
39 | Proof.
40 | intros b₁ b₂ H.
41 | destruct b₁, b₂; try easy; exfalso; now apply H.
42 | Qed.
43 |
44 | Theorem negb_eq_eq : ∀ b₁ b₂, negb b₁ = negb b₂ → b₁ = b₂.
45 | Proof.
46 | intros b₁ b₂ Hn.
47 | now destruct b₁, b₂.
48 | Qed.
49 |
50 | Theorem cons_comm_app : ∀ A (x : A) l l', l ++ x :: l' = l ++ [x] ++ l'.
51 | Proof. easy. Qed.
52 |
53 | Theorem app_of_cons : ∀ A (e : A) el, e :: el = [e] ++ el.
54 | Proof. easy. Qed.
55 |
56 | Theorem fold_right_single : ∀ A B (f : A → B → B) x y,
57 | fold_right f x [y] = f y x.
58 | Proof. easy. Qed.
59 |
60 | Theorem combine_map : ∀ A B C D (f : A → B) (g : C → D) l l',
61 | combine (map f l) (map g l') = map (λ '(x, y), (f x, g y)) (combine l l').
62 | Proof.
63 | intros.
64 | revert l'.
65 | induction l as [| x l]; intros; [ easy | simpl ].
66 | destruct l' as [| x' l']; [ easy | simpl; f_equal; apply IHl ].
67 | Qed.
68 |
69 | Theorem Forall_inv2 : ∀ A (P : A → Prop) a l,
70 | List.Forall P (a :: l) → P a ∧ List.Forall P l.
71 | Proof.
72 | intros A P a l H.
73 | inversion H; now split.
74 | Qed.
75 |
76 | Theorem Forall2_cons_nil : ∀ A B (R : A → B → Prop) x l,
77 | ¬Forall2 R (x :: l) [].
78 | Proof.
79 | intros A B * H.
80 | inversion H.
81 | Qed.
82 |
83 | Theorem Forall2_cons_cons : ∀ A B (R : A → B → Prop) x y l l',
84 | Forall2 R (x :: l) (y :: l')
85 | → R x y ∧ Forall2 R l l'.
86 | Proof.
87 | intros A B * H.
88 | inversion H; subst.
89 | now split.
90 | Qed.
91 |
92 | Theorem Forall2_Forall_combine : ∀ A B f (l1 : list A) (l2 : list B),
93 | Forall2 f l1 l2
94 | ↔ Forall (λ '(x, y), f x y) (combine l1 l2) ∧ length l1 = length l2.
95 | Proof.
96 | intros; revert l2.
97 | induction l1, l2.
98 | all: split; [ intros HF | intros (HF, H) ]; simpl.
99 | 1-6: easy.
100 | {
101 | apply Forall2_cons_cons in HF.
102 | split; [ | now f_equal; apply IHl1 ].
103 | constructor; [ easy | now apply IHl1 ].
104 | } {
105 | simpl in HF; apply Forall_inv2 in HF.
106 | constructor; [ easy | ].
107 | simpl in H; apply Nat.succ_inj in H.
108 | now apply IHl1; split.
109 | }
110 | Qed.
111 |
112 | Theorem flat_map_nil_fun : ∀ A B (f : A → list B) l,
113 | Forall (λ x, f x = []) l
114 | → flat_map f l = [].
115 | Proof.
116 | intros * HF.
117 | induction l as [| x l]; [ easy | simpl ].
118 | apply Forall_inv2 in HF.
119 | destruct HF as (Hx, HF).
120 | rewrite IHl; [ now rewrite Hx | easy ].
121 | Qed.
122 |
123 | Theorem app_repeat_diag : ∀ A (e : A) n,
124 | repeat e n ++ [e] = e :: repeat e n.
125 | Proof.
126 | intros.
127 | induction n; [ easy | ].
128 | simpl; now rewrite IHn.
129 | Qed.
130 |
131 | Theorem list_nil_app_dec {A} : ∀ (l : list A),
132 | {l = []} + {∃ x l', l = l' ++ [x]}.
133 | Proof.
134 | intros l.
135 | destruct l as [| x]; [ now left | right ].
136 | revert x.
137 | induction l as [| y] using rev_ind; intros; [ now exists x, [] | ].
138 | now exists y, (x :: l).
139 | Qed.
140 |
141 | Theorem split_app_eq : ∀ A (el₁ el₂ el₃ el₄ : list A),
142 | el₁ ++ el₂ = el₃ ++ el₄
143 | → { ∃ el, el₃ = el₁ ++ el ∧ el₂ = el ++ el₄ } +
144 | { ∃ el, el₁ = el₃ ++ el ∧ el₄ = el ++ el₂ }.
145 | Proof.
146 | intros A el₁ el₂ el₃ el₄ Hel.
147 | revert el₂ el₃ el₄ Hel.
148 | induction el₁ as [| e₁]; intros. {
149 | left; exists el₃.
150 | now split.
151 | }
152 | destruct el₃ as [| e₃]. {
153 | right; exists (e₁ :: el₁).
154 | now split.
155 | }
156 | simpl in Hel.
157 | injection Hel; clear Hel; intros; subst e₃.
158 | apply IHel₁ in H.
159 | destruct H as [H| H]. {
160 | left; destruct H as (el, (H₁, H₂)); subst el₂ el₃.
161 | exists el; now split.
162 | } {
163 | right; destruct H as (el, (H₁, H₂)); subst el₁ el₄.
164 | exists el; now split.
165 | }
166 | Qed.
167 |
168 | Definition false_neq_negb_false : false ≠ negb false :=
169 | λ p, False_ind False
170 | (eq_ind false (λ e : bool, if e then False else True) I true p).
171 |
172 | Definition true_neq_negb_true : true ≠ negb true :=
173 | λ p, False_ind False
174 | (eq_ind true (λ e : bool, if e then True else False) I false p).
175 |
176 | Definition negb_true_neq_true : negb true ≠ true := false_neq_negb_false.
177 | Definition negb_false_neq_false : negb false ≠ false := true_neq_negb_true.
178 |
179 | Theorem bool_dec_negb_l : ∀ b,
180 | Bool.bool_dec (negb b) b =
181 | right (if b return _ then negb_true_neq_true else negb_false_neq_false).
182 | Proof. intros b; now destruct b. Qed.
183 |
184 | Theorem bool_dec_negb_r : ∀ b,
185 | Bool.bool_dec b (negb b) =
186 | right (if b return _ then true_neq_negb_true else false_neq_negb_false).
187 | Proof. intros b; now destruct b. Qed.
188 |
189 | Theorem Forall2_sym : ∀ A (R : A → A → Prop) l1 l2,
190 | symmetric _ R → Forall2 R l1 l2 → Forall2 R l2 l1.
191 | Proof.
192 | intros * Hs HF; revert l2 HF.
193 | induction l1; intros; [ now destruct l2 | ].
194 | destruct l2; [ now apply Forall2_cons_nil in HF | ].
195 | apply Forall2_cons_cons in HF.
196 | now constructor; [ apply Hs | apply IHl1 ].
197 | Qed.
198 |
199 | (* Type-theoretical Choice Axiom *)
200 | Axiom TTCA : ∀ (A : Type) (R : A → A → Prop), equiv A R →
201 | ∃ f : A → A, (∀ x : A, R x (f x)) ∧ (∀ x y, R x y → f x = f y).
202 |
203 | (* Excluded Middle is the consequence of the Axiom of Choice
204 | (Diaconescu) *)
205 | Theorem EM : ∀ P, P ∨ ¬P.
206 | Proof.
207 | intros P.
208 | set (R (x y : bool) := x = y ∨ P).
209 | assert (He : equiv _ R). {
210 | split; [ intros b; now left | ].
211 | split. {
212 | now intros b c d Hbc [Hcd| Hcd]; [ subst c | right ].
213 | } {
214 | now intros b c [Hbc| Hbc]; [ left; symmetry | right ].
215 | }
216 | }
217 | destruct (TTCA bool R He) as (f & Hx & Hxy).
218 | subst R; simpl in Hx, Hxy.
219 | destruct (Bool.bool_dec (f false) (f true)) as [H| H]. {
220 | destruct (Hx true) as [Ht| Ht]; [ | now left ].
221 | destruct (Hx false) as [Hf| Hf]; [ | now left ].
222 | now rewrite <- Ht, <- Hf in H.
223 | } {
224 | right; intros H₁; apply H.
225 | now apply Hxy; right.
226 | }
227 | Qed.
228 |
229 | Record choice_function {A} (R : A → A → Prop) f := mkcf
230 | { cf_repr_uniqueness : ∀ x y, R x y → f x = f y;
231 | cf_repr_membership : ∀ x, R x (f x) }.
232 |
233 | From Stdlib Require Import Permutation.
234 |
235 | Theorem Permutation_flat_map_map : ∀ A B C (f : A → B → C) la lb,
236 | Permutation
237 | (flat_map (λ a, map (λ b, f a b) lb) la)
238 | (flat_map (λ b, map (λ a, f a b) la) lb).
239 | Proof.
240 | intros.
241 | revert lb.
242 | induction la as [| a la]; intros. {
243 | simpl; rewrite flat_map_nil_fun; [ easy | ].
244 | induction lb; now constructor.
245 | }
246 | simpl.
247 | rewrite IHla; clear IHla.
248 | revert a la.
249 | induction lb as [| b lb]; intros; [ easy | ].
250 | simpl; constructor; rewrite <- IHlb.
251 | do 2 rewrite app_assoc.
252 | apply Permutation_app_tail.
253 | apply Permutation_app_comm.
254 | Qed.
255 |
256 | Fixpoint map2 {A B C} (f : A → B → C) l1 l2 :=
257 | match l1 with
258 | | [] => []
259 | | a :: t =>
260 | match l2 with
261 | | [] => []
262 | | b :: u => f a b :: map2 f t u
263 | end
264 | end.
265 |
--------------------------------------------------------------------------------
/MiscReals.v:
--------------------------------------------------------------------------------
1 | (* Banach-Tarski paradox. *)
2 |
3 | From Stdlib Require Import Utf8 List Relations Wf_nat.
4 | From Stdlib Require Import Reals Psatz.
5 | From Stdlib Require Import ZArith.
6 | Import ListNotations.
7 |
8 | Require Import Misc Words Normalize Reverse.
9 |
10 | Notation "'ℝ'" := R.
11 | Notation "'ℤ'" := Z.
12 | Notation "'ℕ'" := nat.
13 |
14 | Notation "x '≤' y" := (le x y) : nat_scope.
15 |
16 | Notation "'√'" := sqrt.
17 | Notation "x '≤' y" := (Rle x y) : R_scope.
18 | Notation "x '≤' y '<' z" := (Rle x y ∧ Rlt y z)
19 | (at level 70, y at next level) : R_scope.
20 | Notation "x '≤' y '≤' z" := (Rle x y ∧ Rle y z)
21 | (at level 70, y at next level) : R_scope.
22 | Notation "x '<' y '≤' z" := (Rlt x y ∧ Rle y z)
23 | (at level 70, y at next level) : R_scope.
24 |
25 | Open Scope R_scope.
26 |
27 | Theorem fold_Rminus : ∀ x y, x + - y = x - y.
28 | Proof. intros. now fold (Rminus x y). Qed.
29 |
30 | Theorem fold_Rdiv : ∀ x y, x * / y = x / y.
31 | Proof. intros; now fold (Rdiv x y). Qed.
32 |
33 | Theorem fold_Rsqr : ∀ x, x * x = x².
34 | Proof. intros; now fold (Rsqr x). Qed.
35 |
36 | Theorem Rmult_div : ∀ x y z, x * y / z = x / z * y.
37 | Proof. intros; lra. Qed.
38 |
39 | Theorem Rdiv_mult : ∀ x y z, x * (y / z) = x * y / z.
40 | Proof. intros; lra. Qed.
41 |
42 | Theorem Rminus_plus_distr : ∀ x y z, x - (y + z) = x - y - z.
43 | Proof. intros; lra. Qed.
44 |
45 | Theorem Rminus_opp : ∀ x y, x - - y = x + y.
46 | Proof. intros; lra. Qed.
47 |
48 | Theorem Ropp_div_r : ∀ x y, x / - y = - (x / y).
49 | Proof.
50 | intros.
51 | unfold Rdiv.
52 | rewrite Rinv_opp.
53 | symmetry; apply Ropp_mult_distr_r.
54 | Qed.
55 |
56 | Theorem Rmult_div_same : ∀ x y, y ≠ 0 → x / y * y = x.
57 | Proof.
58 | intros * Hy.
59 | unfold Rdiv.
60 | rewrite Rmult_assoc.
61 | rewrite Rinv_l; [ lra | easy ].
62 | Qed.
63 |
64 | Theorem Rplus_shuffle0 : ∀ n m p : ℝ, n + m + p = n + p + m.
65 | Proof.
66 | intros.
67 | rewrite Rplus_comm, <- Rplus_assoc.
68 | f_equal; apply Rplus_comm.
69 | Qed.
70 |
71 | Theorem Rmult_shuffle0 : ∀ n m p : ℝ, n * m * p = n * p * m.
72 | Proof.
73 | intros.
74 | rewrite Rmult_comm, <- Rmult_assoc.
75 | f_equal; apply Rmult_comm.
76 | Qed.
77 |
78 | Theorem Rdiv_mult_simpl_l : ∀ x y z,
79 | x ≠ 0
80 | → (x * y) / (x * z) = y / z.
81 | Proof.
82 | intros * Hx.
83 | unfold Rdiv.
84 | rewrite Rinv_mult.
85 | rewrite <- Rmult_assoc.
86 | f_equal; rewrite Rmult_shuffle0.
87 | rewrite Rinv_r; [ | easy ].
88 | apply Rmult_1_l.
89 | Qed.
90 |
91 | Theorem Rdiv_mult_simpl_r : ∀ x y z,
92 | z ≠ 0
93 | → (x * z) / (y * z) = x / y.
94 | Proof.
95 | intros * Hz.
96 | specialize (Rdiv_mult_simpl_l z x y Hz) as H.
97 | rewrite <- H.
98 | f_equal; lra.
99 | Qed.
100 |
101 | Theorem Req_dec : ∀ x y : ℝ, { x = y } + { x ≠ y }.
102 | Proof.
103 | intros x y.
104 | destruct (Rle_dec x y) as [H₁| H₁]. {
105 | destruct (Rle_dec y x) as [H₂| H₂]. {
106 | left; now apply Rle_antisym.
107 | } {
108 | right; intros H; subst y; apply H₂, Rle_refl.
109 | }
110 | } {
111 | right; intros H; subst y.
112 | apply H₁, Rle_refl.
113 | }
114 | Qed.
115 |
116 | Theorem Rmult5_sqrt2_sqrt5 : ∀ a b c d, 0 <= b →
117 | a * √ b * c * d * √ b = a * b * c * d.
118 | Proof.
119 | intros a b c d Hb.
120 | rewrite Rmult_comm, <- Rmult_assoc; f_equal.
121 | rewrite <- Rmult_assoc; f_equal.
122 | rewrite Rmult_comm, Rmult_assoc; f_equal.
123 | now apply sqrt_sqrt.
124 | Qed.
125 |
126 | Theorem Rdiv_0_l : ∀ x, 0 / x = 0.
127 | Proof.
128 | intros x; unfold Rdiv; apply Rmult_0_l.
129 | Qed.
130 |
131 | Theorem Rdiv_1_r : ∀ x, x / 1 = x.
132 | Proof. intros x; lra. Qed.
133 |
134 | Theorem Rdiv_same : ∀ x, x ≠ 0 → x / x = 1.
135 | Proof.
136 | intros.
137 | unfold Rdiv.
138 | now rewrite Rinv_r.
139 | Qed.
140 |
141 | Theorem Int_part_close_to_1 : ∀ r n,
142 | INR n / INR (n + 1) <= r < 1
143 | → Int_part (r * (INR (n + 1))) = Z.of_nat n.
144 | Proof.
145 | intros * (Hnr, Hr1).
146 | apply Rmult_le_compat_r with (r := INR (n + 1)) in Hnr; [ | apply pos_INR ].
147 | rewrite <- Rmult_div in Hnr.
148 | unfold Rdiv in Hnr.
149 | rewrite Rmult_assoc in Hnr.
150 | rewrite Rinv_r in Hnr; [ | now apply not_0_INR; rewrite Nat.add_comm ].
151 | rewrite Rmult_1_r in Hnr.
152 | apply Rmult_lt_compat_r with (r := INR (n + 1)) in Hr1. 2: {
153 | rewrite plus_INR; simpl.
154 | apply Rplus_le_lt_0_compat; [ apply pos_INR | lra ].
155 | }
156 | rewrite Rmult_1_l in Hr1.
157 | remember (r * INR (n + 1)) as x eqn:Hx.
158 | rewrite plus_INR in Hr1; simpl in Hr1.
159 | rewrite INR_IZR_INZ in Hnr.
160 | rewrite INR_IZR_INZ in Hr1.
161 | unfold Int_part.
162 | apply Z.add_cancel_r with (p := 1%Z).
163 | rewrite Z.sub_add; symmetry.
164 | apply tech_up; [ now rewrite plus_IZR | ].
165 | rewrite plus_IZR; simpl.
166 | now apply Rplus_le_compat_r.
167 | Qed.
168 |
169 | Theorem Int_part_small : ∀ x, 0 <= x < 1 → Int_part x = 0%Z.
170 | Proof.
171 | intros * Hx.
172 | assert (INR 0 / INR (0 + 1) <= x < 1) by (now simpl; lra).
173 | pose proof Int_part_close_to_1 x 0 H as H1.
174 | now simpl in H1; rewrite Rmult_1_r in H1.
175 | Qed.
176 |
177 | Theorem frac_part_small : ∀ x, 0 <= x < 1 → frac_part x = x.
178 | Proof.
179 | intros * Hx.
180 | unfold frac_part.
181 | rewrite Int_part_small; [ lra | easy ].
182 | Qed.
183 |
184 | Theorem pow_INR : ∀ n k, INR (n ^ k) = INR n ^ k.
185 | Proof.
186 | intros.
187 | induction k; [ easy | ].
188 | simpl; rewrite mult_INR.
189 | now rewrite IHk.
190 | Qed.
191 |
192 | Theorem frac_part_INR : ∀ n, frac_part (INR n) = 0.
193 | Proof.
194 | intros.
195 | unfold frac_part.
196 | rewrite Int_part_INR.
197 | now rewrite <- INR_IZR_INZ, Rminus_diag_eq.
198 | Qed.
199 |
200 | Theorem Int_part_IZR : ∀ z, Int_part (IZR z) = z.
201 | Proof.
202 | intros.
203 | destruct (Z_le_dec 0 z) as [Hz| Hz]. {
204 | apply Z2Nat.id in Hz.
205 | rewrite <- Hz at 1.
206 | rewrite <- INR_IZR_INZ.
207 | now rewrite Int_part_INR.
208 | }
209 | apply Z.nle_gt in Hz.
210 | destruct z as [| p| p]; [ easy | easy | ].
211 | unfold IZR.
212 | replace (- IPR p) with (0 - IPR p) by lra.
213 | rewrite Rminus_Int_part1. {
214 | rewrite Int_part_small; [ | lra ].
215 | rewrite <- INR_IPR.
216 | rewrite Int_part_INR.
217 | now rewrite positive_nat_Z.
218 | }
219 | rewrite <- INR_IPR, frac_part_INR; apply base_fp.
220 | Qed.
221 |
222 | Theorem frac_part_IZR : ∀ z, frac_part (IZR z) = 0.
223 | Proof.
224 | intros.
225 | unfold frac_part.
226 | rewrite Int_part_IZR; lra.
227 | Qed.
228 |
229 | Theorem Rpow_div_sub : ∀ x i j,
230 | x ≠ 0
231 | → (j ≤ i)%nat
232 | → x ^ i / x ^ j = x ^ (i - j).
233 | Proof.
234 | intros * Hx Hij.
235 | unfold Rdiv.
236 | replace i with ((i - j) + j)%nat at 1 by now rewrite Nat.sub_add.
237 | now symmetry; apply pow_RN_plus.
238 | Qed.
239 |
240 | Theorem frac_part_interv : ∀ x, 0 ≤ frac_part x < 1.
241 | Proof.
242 | intros.
243 | unfold frac_part.
244 | specialize (base_Int_part x); intros Hx; lra.
245 | Qed.
246 |
247 | Theorem Int_part_interv : ∀ z x, IZR z ≤ x < IZR (z + 1) → Int_part x = z.
248 | Proof.
249 | intros * (Hzx, Hxz).
250 | rewrite plus_IZR in Hxz; simpl in Hxz.
251 | assert (H : 0 ≤ x - IZR z < 1) by lra.
252 | apply Int_part_small in H.
253 | rewrite Rminus_Int_part1 in H. {
254 | rewrite Int_part_IZR in H.
255 | now apply -> Z.sub_move_0_r in H.
256 | } {
257 | rewrite frac_part_IZR.
258 | apply Rle_ge, frac_part_interv.
259 | }
260 | Qed.
261 |
262 | Theorem Rabs_or : ∀ x y, Rabs x = y → x = y ∨ x = - y.
263 | Proof.
264 | intros * Hxy; subst y.
265 | unfold Rabs.
266 | destruct (Rcase_abs x) as [H₁| H₁]; [ right | now left ].
267 | symmetry; apply Ropp_involutive.
268 | Qed.
269 |
270 | Theorem Rabs_eq_0 : ∀ x, Rabs x = 0 → x = 0.
271 | Proof.
272 | intros * Hx.
273 | unfold Rabs in Hx.
274 | destruct (Rcase_abs x); [ | easy ].
275 | apply Ropp_eq_0_compat in Hx.
276 | now rewrite Ropp_involutive in Hx.
277 | Qed.
278 |
279 | Theorem Rabs_lt : ∀ x y, Rabs x < y ↔ - y < x < y.
280 | Proof.
281 | intros; split. {
282 | intros Hxy.
283 | unfold Rabs in Hxy.
284 | destruct (Rcase_abs x); lra.
285 | } {
286 | intros (Hyx, Hxy).
287 | unfold Rabs.
288 | destruct (Rcase_abs x); [ lra | easy ].
289 | }
290 | Qed.
291 |
292 | Theorem Rabs_le : ∀ x y, Rabs x ≤ y ↔ - y ≤ x ≤ y.
293 | Proof.
294 | intros; split. {
295 | intros Hxy.
296 | unfold Rabs in Hxy.
297 | destruct (Rcase_abs x); lra.
298 | } {
299 | intros (Hyx, Hxy).
300 | unfold Rabs.
301 | destruct (Rcase_abs x); [ lra | easy ].
302 | }
303 | Qed.
304 |
305 | Theorem Rabs_sqr : ∀ x, Rabs (x²) = x².
306 | Proof.
307 | intros.
308 | unfold Rabs.
309 | destruct (Rcase_abs x²) as [Hx| Hx]; [ | easy ].
310 | exfalso; apply Rlt_not_le in Hx; apply Hx, Rle_0_sqr.
311 | Qed.
312 |
313 | Theorem Rabs_sqrt : ∀ x, Rabs (√ x) = √ x.
314 | Proof.
315 | intros.
316 | unfold Rabs.
317 | destruct (Rcase_abs (√ x)) as [Hx| Hx]; [ exfalso | easy ].
318 | apply Rlt_not_le in Hx; apply Hx, sqrt_pos.
319 | Qed.
320 |
321 | Theorem Rmult_minus_distr_r : ∀ r1 r2 r3,
322 | (r1 - r2) * r3 = r1 * r3 - r2 * r3.
323 | Proof.
324 | intros.
325 | unfold Rminus.
326 | rewrite Rmult_plus_distr_r; lra.
327 | Qed.
328 |
329 | Theorem Rminus_plus : ∀ x y, x - y + y = x.
330 | Proof. intros; lra. Qed.
331 |
332 | Theorem Rdiv_div : ∀ x y z, x / y / z = x / (y * z).
333 | Proof.
334 | intros.
335 | unfold Rdiv.
336 | rewrite Rinv_mult; lra.
337 | Qed.
338 |
339 | Theorem Rmult_div_r : ∀ x y, y ≠ 0 → y * (x / y) = x.
340 | Proof.
341 | intros * Hy.
342 | unfold Rdiv; rewrite Rmult_comm, Rmult_assoc.
343 | rewrite Rinv_l; [ lra | easy ].
344 | Qed.
345 |
346 | Theorem Rinv_div : ∀ x, / x = 1 / x.
347 | Proof. intros; lra. Qed.
348 |
349 | Theorem nonneg_plus_sqr : ∀ x y, 0 ≤ x² + y².
350 | Proof.
351 | intros.
352 | apply Rplus_le_le_0_compat; apply Rle_0_sqr.
353 | Qed.
354 |
355 | Definition Rsignp x := if Rle_dec 0 x then 1 else -1.
356 | Definition Rsign x := if Req_dec x 0 then 0 else Rsignp x.
357 |
358 | Theorem Rsignp_of_pos : ∀ x, 0 ≤ x → Rsignp x = 1.
359 | Proof.
360 | intros * Hx.
361 | unfold Rsignp.
362 | destruct (Rle_dec 0 x); [ easy | lra ].
363 | Qed.
364 |
365 | Theorem Rsignp_of_neg : ∀ x, x < 0 → Rsignp x = -1.
366 | Proof.
367 | intros * Hx.
368 | unfold Rsignp.
369 | destruct (Rle_dec 0 x); [ lra | easy ].
370 | Qed.
371 |
372 | Theorem Rsign_of_pos : ∀ x, 0 < x → Rsign x = 1.
373 | Proof.
374 | intros * Hx.
375 | unfold Rsign, Rsignp.
376 | destruct (Req_dec x 0); [ lra | ].
377 | destruct (Rle_dec 0 x); [ easy | lra ].
378 | Qed.
379 |
380 | Theorem Rsign_of_neg : ∀ x, x < 0 → Rsign x = -1.
381 | Proof.
382 | intros * Hx.
383 | unfold Rsign, Rsignp.
384 | destruct (Req_dec x 0); [ lra | ].
385 | destruct (Rle_dec 0 x); [ lra | easy ].
386 | Qed.
387 |
388 | Theorem Rsign_mul_distr : ∀ x y, Rsign (x * y) = Rsign x * Rsign y.
389 | Proof.
390 | intros.
391 | unfold Rsign, Rsignp.
392 | destruct (Req_dec (x * y) 0) as [Hxyz| Hxyz]. {
393 | destruct (Req_dec x 0) as [Hx| Hx]; [ lra | ].
394 | destruct (Req_dec y 0) as [Hy| Hy]; [ lra | ].
395 | apply Rmult_integral in Hxyz; lra.
396 | }
397 | destruct (Req_dec x 0) as [Hxz| Hxz]; [ rewrite Hxz in Hxyz; lra | ].
398 | destruct (Req_dec y 0) as [Hyz| Hyz]; [ rewrite Hyz in Hxyz; lra | ].
399 | destruct (Rle_dec 0 (x * y)) as [Hxy| Hxy]. {
400 | destruct (Rle_dec 0 x) as [Hx| Hx]. {
401 | destruct (Rle_dec 0 y) as [Hy| Hy]; [ lra | exfalso ].
402 | apply Hy; clear Hy.
403 | apply Rmult_le_reg_l with (r := x); [ lra | ].
404 | now rewrite Rmult_0_r.
405 | }
406 | destruct (Rle_dec 0 y) as [Hy| Hy]; [ exfalso | lra ].
407 | apply Hx; clear Hx.
408 | apply Rmult_le_reg_r with (r := y); [ lra | ].
409 | now rewrite Rmult_0_l.
410 | }
411 | destruct (Rle_dec 0 x) as [Hx| Hx]. {
412 | destruct (Rle_dec 0 y) as [Hy| Hy]; [ exfalso | lra ].
413 | apply Hxy; clear Hxy.
414 | now apply Rmult_le_pos.
415 | } {
416 | destruct (Rle_dec 0 y) as [Hy| Hy]; [ lra | exfalso ].
417 | apply Hxy; clear Hxy.
418 | rewrite <- Rmult_opp_opp.
419 | apply Rmult_le_pos; lra.
420 | }
421 | Qed.
422 |
423 | Theorem Rneq_le_lt : ∀ x y, x ≠ y → x ≤ y → x < y.
424 | Proof.
425 | intros * Hnxy Hxy.
426 | apply Rnot_le_lt; intros H.
427 | now apply Rle_antisym in H.
428 | Qed.
429 |
430 | Theorem sqrt_diff_sqr_eq_0 : ∀ x y,
431 | 0 ≤ x ≤ y
432 | → √ (y² - x²) = 0
433 | → x = y.
434 | Proof.
435 | intros * Hxy Hyx.
436 | apply sqrt_eq_0 in Hyx; [ apply Rsqr_inj; lra | ].
437 | enough (x² ≤ y²) by lra.
438 | apply Rsqr_incr_1; lra.
439 | Qed.
440 |
441 | Definition Rediv_mod x y :=
442 | let k :=
443 | match Rcase_abs y with
444 | | left _ => (- Int_part (x / - y))%Z
445 | | right _ => Int_part (x / y)
446 | end
447 | in
448 | (k, x - IZR k * y).
449 |
450 | Definition Rediv x y := fst (Rediv_mod x y).
451 | Definition Rmod x y := snd (Rediv_mod x y).
452 |
453 | Notation "x '//' y" := (Rediv x y) (at level 40).
454 | Notation "x 'rmod' y" := (Rmod x y) (at level 40).
455 |
456 | Theorem Rmod_interv : ∀ x y, 0 < y → 0 ≤ x rmod y < y.
457 | Proof.
458 | intros * Hy.
459 | unfold Rmod, Rediv_mod, snd.
460 | destruct (Rcase_abs y) as [Hya| Hya]; [ lra | ].
461 | split. {
462 | apply Rmult_le_reg_r with (r := / y); [ now apply Rinv_0_lt_compat | ].
463 | rewrite Rmult_0_l, fold_Rdiv, Rdiv_minus_distr, Rmult_div.
464 | rewrite Rmult_div_same; [ | lra ].
465 | specialize (base_Int_part (x / y)); lra.
466 | } {
467 | apply Rmult_lt_reg_r with (r := / y); [ now apply Rinv_0_lt_compat | ].
468 | rewrite fold_Rdiv, fold_Rdiv, Rdiv_minus_distr, Rmult_div.
469 | rewrite Rmult_div_same; [ | lra ].
470 | rewrite Rdiv_same; [ | lra ].
471 | specialize (base_Int_part (x / y)); lra.
472 | }
473 | Qed.
474 |
475 | Theorem Rmod_from_ediv : ∀ x y, x rmod y = x - IZR (x // y) * y.
476 | Proof.
477 | intros.
478 | unfold Rmod, Rediv, fst, snd.
479 | remember (Rediv_mod x y) as rdm eqn:Hrdm.
480 | symmetry in Hrdm.
481 | destruct rdm as (q, r).
482 | unfold Rediv_mod in Hrdm.
483 | destruct (Rcase_abs y) as [Hy| Hy]. {
484 | remember Z.sub as f.
485 | injection Hrdm; clear Hrdm; intros Hr Hq; subst f.
486 | now rewrite Hq in Hr.
487 | } {
488 | remember Z.sub as f.
489 | injection Hrdm; clear Hrdm; intros Hr Hq; subst f.
490 | now rewrite Hq in Hr.
491 | }
492 | Qed.
493 |
494 | Theorem Int_part_neg : ∀ x,
495 | Int_part (- x) =
496 | (- Int_part x - if Req_dec x (IZR (Int_part x)) then 0 else 1)%Z.
497 | Proof.
498 | intros.
499 | destruct (Req_dec x (IZR (Int_part x))) as [Hx| Hx]. {
500 | rewrite Hx at 1.
501 | now rewrite <- opp_IZR, Int_part_IZR, Z.sub_0_r.
502 | }
503 | apply Int_part_interv.
504 | rewrite Z.sub_simpl_r, opp_IZR.
505 | unfold Z.sub; rewrite <- Z.opp_add_distr.
506 | rewrite opp_IZR, plus_IZR; simpl (IZR 1).
507 | specialize (base_Int_part x) as H; lra.
508 | Qed.
509 |
510 | Theorem Rediv_add_1 : ∀ x y, y ≠ 0 → (x + y) // y = (x // y + 1)%Z.
511 | Proof.
512 | intros * Hyz.
513 | unfold Rediv, Rediv_mod, fst.
514 | destruct (Rcase_abs y) as [Hy| Hy]. {
515 | unfold Rdiv.
516 | rewrite Rinv_opp.
517 | rewrite <- Ropp_mult_distr_r.
518 | rewrite Rmult_plus_distr_r.
519 | rewrite Rinv_r; [ | lra ].
520 | rewrite Ropp_plus_distr.
521 | rewrite fold_Rminus.
522 | rewrite <- Ropp_mult_distr_r.
523 | ring_simplify.
524 | rewrite Rminus_Int_part1. {
525 | rewrite Z.opp_sub_distr.
526 | replace 1 with (IZR 1) by lra.
527 | now rewrite Int_part_IZR.
528 | } {
529 | replace 1 with (IZR 1) by lra.
530 | rewrite frac_part_IZR.
531 | specialize (frac_part_interv (- (x * / y))) as (Hn, Hp); lra.
532 | }
533 | }
534 | rewrite Rdiv_plus_distr.
535 | rewrite Rdiv_same; [ | easy ].
536 | rewrite plus_Int_part2. {
537 | replace 1 with (IZR 1) by lra.
538 | now rewrite Int_part_IZR.
539 | } {
540 | replace 1 with (IZR 1) at 1 by lra.
541 | rewrite frac_part_IZR, Rplus_0_r.
542 | apply frac_part_interv.
543 | }
544 | Qed.
545 |
546 | Theorem Rediv_opp_r : ∀ x y, y ≠ 0 → x // - y = (- (x // y))%Z.
547 | Proof.
548 | intros * Hyz.
549 | unfold "//", fst, Rediv_mod.
550 | destruct (Rcase_abs (- y)) as [Hy| Hy]. {
551 | destruct (Rcase_abs y); [ lra | now rewrite Ropp_involutive ].
552 | } {
553 | destruct (Rcase_abs y); [ now rewrite Z.opp_involutive | lra ].
554 | }
555 | Qed.
556 |
557 | Theorem Rediv_add_nat : ∀ x y n,
558 | y ≠ 0
559 | → (x + INR n * y) // y = (x // y + Z.of_nat n)%Z.
560 | Proof.
561 | intros * Hyz.
562 | induction n; [ now simpl; rewrite Rmult_0_l, Rplus_0_r, Z.add_0_r | ].
563 | rewrite S_INR, Rmult_plus_distr_r, Rmult_1_l, <- Rplus_assoc.
564 | rewrite Rediv_add_1; [ | easy ].
565 | rewrite IHn; lia.
566 | Qed.
567 |
568 | Theorem Rediv_add_Z : ∀ x y a,
569 | y ≠ 0
570 | → (x + IZR a * y) // y = (x // y + a)%Z.
571 | Proof.
572 | intros * Hyz.
573 | destruct (Z_le_dec 0 a) as [Ha| Ha]. {
574 | apply IZN in Ha.
575 | destruct Ha as (n, Hn); subst a.
576 | rewrite <- INR_IZR_INZ.
577 | now apply Rediv_add_nat.
578 | }
579 | remember (- a)%Z as b eqn:Hb.
580 | assert (a = (- b)%Z) by lia.
581 | subst a; clear Hb.
582 | rename b into a.
583 | assert (Hb : (0 < a)%Z) by lia.
584 | clear Ha; rename Hb into Ha.
585 | apply Z.lt_le_incl in Ha.
586 | apply IZN in Ha.
587 | destruct Ha as (n, Hn); subst a.
588 | rewrite opp_IZR.
589 | rewrite <- INR_IZR_INZ.
590 | rewrite <- Ropp_mult_distr_l, Ropp_mult_distr_r.
591 | symmetry; rewrite <- Z.opp_involutive; symmetry.
592 | rewrite <- Rediv_opp_r; [ | easy ].
593 | rewrite Rediv_add_nat; [ | lra ].
594 | rewrite Z.opp_add_distr.
595 | rewrite Rediv_opp_r; [ | easy ].
596 | now rewrite Z.opp_involutive.
597 | Qed.
598 |
599 | Theorem Rmod_add_Z : ∀ x y a,
600 | y ≠ 0
601 | → (x + IZR a * y) rmod y = x rmod y.
602 | Proof.
603 | intros * Hy.
604 | rewrite Rmod_from_ediv.
605 | rewrite Rediv_add_Z; [ | easy ].
606 | rewrite plus_IZR.
607 | rewrite Rmult_plus_distr_r.
608 | remember (IZR a * y) as u.
609 | remember (IZR (x // y) * y) as v.
610 | now replace (x + u - (v + u)) with (x - v) by lra; subst u v.
611 | Qed.
612 |
613 | Theorem Int_part_0 : Int_part 0 = 0%Z.
614 | Proof. rewrite Int_part_small; [ easy | lra ]. Qed.
615 |
616 | Theorem Rmod_0_l : ∀ x, 0 rmod x = 0.
617 | Proof.
618 | intros x.
619 | unfold Rmod, snd, Rediv_mod.
620 | do 2 rewrite Rdiv_0_l.
621 | rewrite Int_part_0, Z.opp_0.
622 | destruct (Rcase_abs x); lra.
623 | Qed.
624 |
625 | Theorem Rmod_mul_same : ∀ x a, (IZR a * x) rmod x = 0.
626 | Proof.
627 | intros.
628 | destruct (Req_dec x 0) as [Hx| Hx]. {
629 | rewrite Hx, Rmult_0_r; apply Rmod_0_l.
630 | }
631 | specialize (Rmod_add_Z 0 x a Hx) as H.
632 | rewrite Rplus_0_l in H; rewrite H.
633 | apply Rmod_0_l.
634 | Qed.
635 |
636 | Theorem Rmod_small : ∀ x y, 0 ≤ x < y → x rmod y = x.
637 | Proof.
638 | intros * (Hx, Hxy).
639 | unfold Rmod, snd, Rediv_mod.
640 | destruct (Rcase_abs y) as [Hyn| Hyp]; [ lra | ].
641 | assert (H : 0 ≤ x / y < 1). {
642 | split. {
643 | apply Rmult_le_reg_r with (r := y); [ lra | ].
644 | rewrite Rmult_0_l, Rmult_div_same; [ easy | lra ].
645 | } {
646 | apply Rmult_lt_reg_r with (r := y); [ lra | ].
647 | rewrite Rmult_1_l, Rmult_div_same; [ easy | lra ].
648 | }
649 | }
650 | apply Int_part_small in H.
651 | rewrite H; simpl.
652 | now rewrite Rmult_0_l, Rminus_0_r.
653 | Qed.
654 |
655 | Theorem Rediv_mul_r : ∀ x y z,
656 | x // (y * z) =
657 | (Int_part (x / (y * z)) +
658 | if Rcase_abs (y * z) then
659 | if Req_dec (x / (y * z)) (IZR (Int_part (x / (y * z)))) then 0
660 | else 1
661 | else 0)%Z.
662 | Proof.
663 | intros.
664 | unfold "//", fst, Rediv_mod.
665 | destruct (Rcase_abs (y * z)) as [Hyz| Hyz]; [ | now rewrite Z.add_0_r ].
666 | rewrite Ropp_div_r.
667 | rewrite Int_part_neg.
668 | rewrite Z.opp_sub_distr.
669 | rewrite Z.opp_involutive.
670 | destruct (Req_dec (x / (y * z)) (IZR (Int_part (x / (y * z))))); [ | easy ].
671 | now rewrite Z.add_0_r.
672 | Qed.
673 |
674 | Theorem frac_part_double : ∀ x,
675 | frac_part (2 * x) =
676 | 2 * frac_part x - if Rlt_dec (frac_part x) (1 / 2) then 0 else 1.
677 | Proof.
678 | intros.
679 | do 2 rewrite <- Rplus_diag.
680 | destruct (Rlt_dec (frac_part x) (1 / 2)) as [Hx| Hx]. {
681 | rewrite Rminus_0_r; apply plus_frac_part2; lra.
682 | }
683 | apply plus_frac_part1; lra.
684 | Qed.
685 |
686 | Theorem Int_part_double : ∀ x,
687 | Int_part (2 * x) =
688 | (2 * Int_part x + if Rlt_dec (frac_part x) (1 / 2) then 0 else 1)%Z.
689 | Proof.
690 | intros.
691 | rewrite <- Rplus_diag.
692 | destruct (Rlt_dec (frac_part x) (1 / 2)) as [Hx| Hx]. {
693 | rewrite plus_Int_part2; [ lia | lra ].
694 | } {
695 | rewrite plus_Int_part1; [ lia | lra ].
696 | }
697 | Qed.
698 |
699 | Theorem pow_1_abs_nat_odd : ∀ n, (-1) ^ Z.abs_nat (2 * n + 1) = -1.
700 | Proof.
701 | intros n.
702 | destruct n as [| n| n]. {
703 | rewrite Z.mul_0_r, Z.add_0_l.
704 | simpl (Z.abs_nat _); unfold Pos.to_nat; simpl (Pos.iter_op _ _ _).
705 | now rewrite pow_1.
706 | } {
707 | rewrite Zabs2Nat.inj_add; [ | lia | lia ].
708 | rewrite Zabs2Nat.inj_mul.
709 | simpl (Z.abs_nat _); unfold Pos.to_nat; simpl (Pos.iter_op _ _ _).
710 | now rewrite Nat.add_1_r, pow_1_odd.
711 | } {
712 | replace (Z.neg n) with (- Z.pos n)%Z by apply Pos2Z.opp_pos.
713 | rewrite <- Zopp_mult_distr_r, <- Z.opp_sub_distr.
714 | rewrite <- Zabs_N_nat, Zabs2N.inj_opp, Zabs_N_nat.
715 | rewrite Zabs2Nat.inj_sub; [ | lia ].
716 | simpl (Z.abs_nat 1); unfold Pos.to_nat; simpl (Pos.iter_op _ _ _).
717 | rewrite <- Rpow_div_sub; [ | lra | lia ].
718 | rewrite pow_1, Zabs2Nat.inj_mul.
719 | simpl (Z.abs_nat 2); unfold Pos.to_nat; simpl (Pos.iter_op _ _ _).
720 | rewrite pow_1_even; lra.
721 | }
722 | Qed.
723 |
724 | Theorem Rdiv_mod : ∀ x y, y ≠ 0 → x = y * IZR (x // y) + x rmod y.
725 | Proof.
726 | intros x y Hy.
727 | rewrite Rmod_from_ediv.
728 | lra.
729 | Qed.
730 |
--------------------------------------------------------------------------------
/Normalize.v:
--------------------------------------------------------------------------------
1 | (* Banach-Tarski paradox. *)
2 |
3 | From Stdlib Require Import Utf8 List Relations Wf_nat.
4 | Import ListNotations.
5 | Require Import Misc Words.
6 |
7 | Fixpoint norm_list el :=
8 | match el with
9 | | nil => nil
10 | | e₁ :: el₁ =>
11 | match norm_list el₁ with
12 | | nil => e₁ :: nil
13 | | e₂ :: el₂ => if letter_opp_dec e₁ e₂ then el₂ else e₁ :: e₂ :: el₂
14 | end
15 | end.
16 |
17 | Theorem norm_list_no_consec : ∀ e el el₁ el₂,
18 | norm_list el ≠ el₁ ++ e :: negf e :: el₂.
19 | Proof.
20 | intros e el el₁ el₂.
21 | revert el₁.
22 | induction el as [| e₁]; intros; [ intros H; now destruct el₁ | ].
23 | simpl; remember (norm_list el) as nl eqn:Hnl; symmetry in Hnl.
24 | destruct nl as [| e₂]. {
25 | clear; intros H.
26 | destruct el₁ as [| e₂]; intros; [ easy | simpl in H ].
27 | injection H; clear H; intros; subst e₂.
28 | now destruct el₁.
29 | }
30 | destruct (letter_opp_dec e₁ e₂) as [H₁| H₁]. {
31 | intros H; subst nl.
32 | pose proof IHel (e₂ :: el₁) as H₂; simpl in H₂.
33 | now apply H₂.
34 | }
35 | unfold letter_opp in H₁.
36 | destruct e₁ as (x₁, d₁).
37 | destruct e₂ as (x₂, d₂).
38 | destruct (letter_dec x₁ x₂) as [H₂| H₂]. {
39 | subst x₂.
40 | destruct (Bool.bool_dec d₁ d₂) as [H₂| H₂]. {
41 | clear H₁; subst d₂.
42 | destruct el₁ as [| e₁]. {
43 | simpl; intros H.
44 | injection H; clear H; intros H₁ H₂ H₃; subst e.
45 | simpl in H₂.
46 | injection H₂; clear H₂; intros H.
47 | symmetry in H; revert H; apply Bool.no_fixpoint_negb.
48 | }
49 | simpl; intros H.
50 | injection H; clear H; intros H₁ H₂; subst e₁.
51 | eapply IHel, H₁.
52 | }
53 | exfalso; apply H₁; constructor.
54 | }
55 | clear H₁.
56 | destruct el₁ as [| e₁]. {
57 | simpl; intros H.
58 | injection H; clear H; intros H₁ H₃ H₄; subst e.
59 | simpl in H₃.
60 | injection H₃; clear H₃; intros; subst x₂ d₂.
61 | now apply H₂.
62 | }
63 | simpl; intros H.
64 | injection H; clear H; intros H₁ H₃.
65 | eapply IHel, H₁.
66 | Qed.
67 |
68 | Theorem norm_list_no_consec2 : ∀ e el el₁ el₂,
69 | norm_list el ≠ el₁ ++ negf e :: e :: el₂.
70 | Proof.
71 | intros e el el₁ el₂.
72 | pose proof norm_list_no_consec (negf e) el el₁ el₂ as H.
73 | now rewrite negf_involutive in H.
74 | Qed.
75 |
76 | Theorem norm_list_no_start : ∀ e el el₂,
77 | norm_list el ≠ e :: negf e :: el₂.
78 | Proof.
79 | intros e el el₂.
80 | apply norm_list_no_consec with (el₁ := []).
81 | Qed.
82 |
83 | Theorem norm_list_no_start2 : ∀ e el el₂,
84 | norm_list el ≠ negf e :: e :: el₂.
85 | Proof.
86 | intros e el el₂.
87 | apply norm_list_no_consec2 with (el₁ := []).
88 | Qed.
89 |
90 | Theorem norm_list_cancel : ∀ el e,
91 | norm_list (e :: negf e :: el) = norm_list el.
92 | Proof.
93 | intros el (t, d).
94 | revert t d.
95 | induction el as [| (t₁, d₁)]; intros. {
96 | simpl; now rewrite letter_dec_diag, bool_dec_negb_r.
97 | }
98 | remember (FE t₁ d₁ :: el) as el₁ eqn:Hel₁.
99 | symmetry in Hel₁; simpl.
100 | remember (norm_list el₁) as el₂ eqn:Hel₂.
101 | symmetry in Hel₂; simpl.
102 | destruct el₂ as [| (t₂, d₂)]. {
103 | now rewrite letter_dec_diag, bool_dec_negb_r.
104 | }
105 | subst el₁.
106 | destruct (letter_dec t t₂) as [H₁| H₁]. {
107 | subst t₂.
108 | destruct (Bool.bool_dec (negb d) d₂) as [H₁| H₁]. {
109 | subst d₂.
110 | now rewrite letter_dec_diag, bool_dec_negb_r.
111 | }
112 | apply negb_neq in H₁; subst d₂.
113 | destruct el₂ as [| (t₂, d₂)]; [ easy | ].
114 | destruct (letter_dec t t₂) as [H₁| H₁]; [ | easy ].
115 | subst t₂.
116 | destruct (Bool.bool_dec d d₂) as [H₁| H₁]; [ easy | ].
117 | apply not_eq_sym, neq_negb in H₁; subst d₂.
118 | exfalso; revert Hel₂; apply norm_list_no_start.
119 | }
120 | now rewrite letter_dec_diag, bool_dec_negb_r.
121 | Qed.
122 |
123 | Theorem norm_list_cancel_in : ∀ el₁ el₂ e,
124 | norm_list (el₁ ++ e :: negf e :: el₂) =
125 | norm_list (el₁ ++ el₂).
126 | Proof.
127 | intros.
128 | revert el₂ e.
129 | induction el₁ as [| e₁]; intros. {
130 | do 2 rewrite app_nil_l.
131 | apply norm_list_cancel.
132 | }
133 | simpl; now rewrite IHel₁.
134 | Qed.
135 |
136 | Theorem norm_list_cancel_in2 : ∀ el₁ el₂ e,
137 | norm_list (el₁ ++ negf e :: e :: el₂) =
138 | norm_list (el₁ ++ el₂).
139 | Proof.
140 | intros.
141 | pose proof norm_list_cancel_in el₁ el₂ (negf e) as H.
142 | now rewrite negf_involutive in H.
143 | Qed.
144 |
145 | Theorem is_normal : ∀ el₁ el₂ el₃,
146 | norm_list (el₁ ++ norm_list el₂ ++ el₃) =
147 | norm_list (el₁ ++ el₂ ++ el₃).
148 | Proof.
149 | intros.
150 | revert el₁ el₃.
151 | induction el₂ as [| e₂]; intros; [ easy | simpl ].
152 | remember (norm_list el₂) as el eqn:Hel₂; symmetry in Hel₂.
153 | destruct el as [| e]. {
154 | simpl in IHel₂; simpl.
155 | rewrite cons_comm_app, app_assoc.
156 | now rewrite IHel₂, <- app_assoc.
157 | }
158 | destruct (letter_opp_dec e₂ e) as [H₁| H₁]. {
159 | apply letter_opp_negf in H₁; subst e₂.
160 | rewrite cons_comm_app.
161 | do 2 rewrite app_assoc.
162 | rewrite <- IHel₂.
163 | do 2 rewrite <- app_assoc; simpl.
164 | rewrite norm_list_cancel_in2.
165 | easy.
166 | }
167 | rewrite cons_comm_app.
168 | do 2 rewrite app_assoc.
169 | rewrite <- IHel₂.
170 | do 2 rewrite <- app_assoc.
171 | easy.
172 | Qed.
173 |
174 | Theorem norm_list_normal_l : ∀ el₁ el₂,
175 | norm_list (el₁ ++ el₂) = norm_list (norm_list el₁ ++ el₂).
176 | Proof.
177 | intros.
178 | replace el₁ with ([] ++ el₁) by easy.
179 | rewrite <- app_assoc.
180 | now rewrite <- is_normal.
181 | Qed.
182 |
183 | Theorem norm_list_normal_r : ∀ el₁ el₂,
184 | norm_list (el₁ ++ el₂) = norm_list (el₁ ++ norm_list el₂).
185 | Proof.
186 | intros.
187 | replace el₂ with (el₂ ++ []) by apply app_nil_r.
188 | rewrite <- is_normal.
189 | now do 2 rewrite app_nil_r.
190 | Qed.
191 |
192 | Theorem norm_list_idemp : ∀ el, norm_list (norm_list el) = norm_list el.
193 | Proof.
194 | intros el.
195 | pose proof is_normal [] el [] as H.
196 | simpl in H; do 2 rewrite app_nil_r in H.
197 | easy.
198 | Qed.
199 |
200 | Theorem norm_list_cons : ∀ el e,
201 | norm_list (e :: el) = e :: el
202 | → norm_list el = el.
203 | Proof.
204 | intros el e Hn.
205 | revert e Hn.
206 | induction el as [| e₁]; intros; [ easy | ].
207 | remember (e₁ :: el) as el₁ eqn:Hel.
208 | simpl in Hn.
209 | remember (norm_list el₁) as el₂ eqn:Hel₁; symmetry in Hel₁.
210 | destruct el₂ as [| e₂]. {
211 | injection Hn; clear Hn; intros; now subst.
212 | }
213 | destruct (letter_opp_dec e e₂) as [H₁| H₁]. {
214 | apply letter_opp_negf in H₁; subst e el₂.
215 | exfalso; revert Hel₁; apply norm_list_no_start.
216 | }
217 | injection Hn; clear Hn; intros; subst el₁.
218 | easy.
219 | Qed.
220 |
221 | Theorem norm_list_app_diag : ∀ el₁ el₂,
222 | norm_list (el₁ ++ el₂) = el₁ ++ el₂ → norm_list el₁ = el₁.
223 | Proof.
224 | intros el₁ el₂ Hn.
225 | revert el₂ Hn.
226 | induction el₁ as [| e₁]; intros; [ easy | simpl ].
227 | generalize Hn; intros Hn₁.
228 | rewrite <- app_comm_cons in Hn₁.
229 | apply norm_list_cons, IHel₁ in Hn₁.
230 | rewrite Hn₁.
231 | destruct el₁ as [| e₂]; [ easy | ].
232 | destruct (letter_opp_dec e₁ e₂) as [H₁| H₁]; [ exfalso | easy ].
233 | apply letter_opp_negf in H₁; subst e₁.
234 | revert Hn; apply norm_list_no_start2.
235 | Qed.
236 |
237 | Theorem norm_list_dec : ∀ el,
238 | { norm_list el = el } +
239 | { ∃ el₁ t d el₂, el = el₁ ++ FE t d :: FE t (negb d) :: el₂ }.
240 | Proof.
241 | intros el.
242 | induction el as [| e]; [ now left | ].
243 | destruct IHel as [IHel| IHel]. {
244 | simpl.
245 | rewrite IHel.
246 | destruct el as [| e₁]; [ now left | ].
247 | destruct (letter_opp_dec e e₁) as [H₁| H₁]; [ right | now left ].
248 | apply letter_opp_sym, letter_opp_negf in H₁; subst e₁.
249 | destruct e as (t, d).
250 | exists [], t, d, el.
251 | easy.
252 | }
253 | right.
254 | destruct IHel as (el₁, (t, (d, (el₂, IHel)))).
255 | exists (e :: el₁), t, d, el₂; subst el.
256 | easy.
257 | Qed.
258 |
259 | Theorem norm_list_repeat : ∀ e n, norm_list (repeat e n) = repeat e n.
260 | Proof.
261 | intros e n.
262 | induction n; [ easy | simpl ].
263 | rewrite IHn.
264 | remember (repeat e n) as el eqn:Hel.
265 | symmetry in Hel.
266 | destruct el as [| e₁]; [ easy | ].
267 | destruct (letter_opp_dec e e₁) as [H| H]; [ | easy ].
268 | apply letter_opp_negf in H; subst e.
269 | exfalso.
270 | destruct n; [ easy | ].
271 | injection Hel; clear Hel; intros Hel H.
272 | revert H; apply no_fixpoint_negf.
273 | Qed.
274 |
275 | Theorem norm_list_is_cons : ∀ el e el₁,
276 | norm_list el = e :: el₁ → norm_list el₁ = el₁.
277 | Proof.
278 | intros * H.
279 | destruct (norm_list_dec el₁) as [H₁| H₁]; [ easy | ].
280 | destruct H₁ as (el₂ & t & d & el₃ & H₁).
281 | subst el₁.
282 | exfalso; revert H.
283 | replace (FE t (negb d)) with (negf (FE t d)) by easy.
284 | rewrite app_comm_cons.
285 | apply norm_list_no_consec.
286 | Qed.
287 |
288 | Theorem norm_list_app_split : ∀ el₁ el₂ el₃ el₄ e,
289 | norm_list el₁ ++ norm_list el₂ = el₃ ++ e :: negf e :: el₄
290 | → norm_list el₁ = el₃ ++ [e] ∧ norm_list el₂ = negf e :: el₄.
291 | Proof.
292 | intros el₁ el₂ el₃ el₄ e Hn.
293 | apply split_app_eq in Hn.
294 | destruct Hn as [(el, (H₁, H₂))| (el, (H₁, H₂))]. {
295 | exfalso; revert H₂; apply norm_list_no_consec.
296 | }
297 | rewrite app_of_cons in H₂.
298 | apply split_app_eq in H₂.
299 | destruct H₂ as [(el', (H₂, H₃))| (el', (H₂, H₃))]. {
300 | subst el.
301 | destruct el' as [| e']. {
302 | rewrite app_nil_r in H₁.
303 | rewrite app_nil_l in H₃; symmetry in H₃.
304 | now split.
305 | }
306 | simpl in H₃.
307 | injection H₃; clear H₃; intros H₂ H₃; subst e'.
308 | exfalso; revert H₁; apply norm_list_no_consec.
309 | }
310 | destruct el as [| e₁]. {
311 | simpl in H₂; subst el'.
312 | exfalso; revert H₃; apply norm_list_no_start.
313 | }
314 | simpl in H₂.
315 | injection H₂; clear H₂; intros H₂ H₄; subst e₁.
316 | symmetry in H₂.
317 | apply app_eq_nil in H₂.
318 | destruct H₂; subst el el'.
319 | now split.
320 | Qed.
321 |
--------------------------------------------------------------------------------
/NotEmptyPath.v:
--------------------------------------------------------------------------------
1 | (* Banach-Tarski paradox. *)
2 |
3 | From Stdlib Require Import Utf8 Arith List.
4 | From Stdlib Require Import Reals Psatz.
5 | Import ListNotations.
6 |
7 | Require Import MiscReals Words Normalize Reverse Matrix.
8 |
9 | Definition rotate_param e '(a, b, c, N) :=
10 | match e with
11 | | ạ => ((3 * a)%Z, (b - 2 * c)%Z, (4 * b + c)%Z, S N)
12 | | ạ⁻¹ => ((3 * a)%Z, (b + 2 * c)%Z, (- 4 * b + c)%Z, S N)
13 | | ḅ => ((a - 4 * b)%Z, (b + 2 * a)%Z, (3 * c)%Z, S N)
14 | | ḅ⁻¹ => ((a + 4 * b)%Z, (b - 2 * a)%Z, (3 * c)%Z, S N)
15 | end.
16 |
17 | Arguments Z.mul _ _ : simpl nomatch.
18 |
19 | Theorem rotate_param_rotate : ∀ el x y z n a b c N,
20 | fold_right rotate_param (x, y, z, n) el = (a, b, c, N)
21 | ↔ fold_right rotate (V (IZR x / 3^n) (IZR y * √2 / 3^n) (IZR z / 3^n)) el =
22 | V (IZR a / 3^N) (IZR b*√2 / 3^N) (IZR c / 3^N)
23 | ∧ N = (n + length el)%nat.
24 | Proof.
25 | intros el x y z n a₁ b₁ c₁ N₁.
26 | split. {
27 | intros Hr.
28 | simpl in Hr; simpl.
29 | revert a₁ b₁ c₁ N₁ Hr.
30 | induction el as [| (t, d)]; intros. {
31 | simpl; simpl in Hr; rewrite Nat.add_0_r.
32 | injection Hr; intros; subst; simpl.
33 | split; reflexivity.
34 | }
35 | simpl in Hr; simpl.
36 | remember (fold_right rotate_param (x, y, z, n) el) as rp eqn:Hrp.
37 | symmetry in Hrp.
38 | destruct rp as (((a, b), c), N).
39 | pose proof IHel _ _ _ _ (eq_refl _) as H.
40 | destruct H as (H, HN).
41 | erewrite H.
42 | simpl in Hr; simpl; unfold Rdiv.
43 | destruct t, d; injection Hr; clear Hr; intros; subst a₁ b₁ c₁ N₁ N; simpl. {
44 | split; [ | rewrite Nat.add_succ_r; reflexivity ].
45 | rewrite plus_IZR, plus_IZR.
46 | progress repeat rewrite mult_IZR.
47 | rewrite Rinv_mult.
48 | progress repeat rewrite Rmult_1_l.
49 | progress repeat rewrite Rmult_0_l.
50 | progress repeat rewrite Rplus_0_l.
51 | progress repeat rewrite Rplus_0_r.
52 | progress repeat rewrite <- Rmult_assoc.
53 | unfold Rdiv.
54 | rewrite Rmult5_sqrt2_sqrt5; [ f_equal; lra | lra ].
55 | } {
56 | split; [ | rewrite Nat.add_succ_r; reflexivity ].
57 | rewrite plus_IZR, minus_IZR.
58 | progress repeat rewrite mult_IZR.
59 | rewrite Rinv_mult.
60 | progress repeat rewrite Rmult_1_l.
61 | progress repeat rewrite Rmult_0_l.
62 | progress repeat rewrite Rplus_0_l.
63 | progress repeat rewrite Rplus_0_r.
64 | progress repeat rewrite <- Rmult_assoc.
65 | unfold Rdiv.
66 | rewrite Rmult5_sqrt2_sqrt5; [ f_equal; lra | lra ].
67 | } {
68 | split; [ | rewrite Nat.add_succ_r; reflexivity ].
69 | rewrite plus_IZR, minus_IZR.
70 | progress repeat rewrite mult_IZR.
71 | rewrite Rinv_mult.
72 | progress repeat rewrite Rmult_1_l.
73 | progress repeat rewrite Rmult_0_l.
74 | progress repeat rewrite Rplus_0_l.
75 | progress repeat rewrite Rplus_0_r.
76 | progress repeat rewrite <- Rmult_assoc.
77 | unfold Rdiv.
78 | rewrite Rmult5_sqrt2_sqrt5; [ f_equal; lra | lra ].
79 | } {
80 | split; [ | rewrite Nat.add_succ_r; reflexivity ].
81 | rewrite plus_IZR, minus_IZR.
82 | progress repeat rewrite mult_IZR.
83 | rewrite Rinv_mult.
84 | progress repeat rewrite Rmult_1_l.
85 | progress repeat rewrite Rmult_0_l.
86 | progress repeat rewrite Rplus_0_l.
87 | progress repeat rewrite Rplus_0_r.
88 | progress repeat rewrite <- Rmult_assoc.
89 | unfold Rdiv.
90 | rewrite Rmult5_sqrt2_sqrt5; [ f_equal; lra | lra ].
91 | }
92 | }
93 | intros Hr.
94 | revert x y z n a₁ b₁ c₁ N₁ Hr.
95 | induction el as [| e] using rev_ind; intros. {
96 | simpl in Hr; simpl; rewrite Nat.add_0_r in Hr.
97 | destruct Hr as (Hr, Hn); subst N₁.
98 | unfold Rdiv in Hr.
99 | injection Hr; intros Hz Hy Hx.
100 | f_equal; f_equal; f_equal. {
101 | apply Rmult_eq_reg_r, eq_IZR in Hx; [ assumption | ].
102 | apply Rinv_neq_0_compat, pow_nonzero; lra.
103 | } {
104 | apply Rmult_eq_reg_r in Hy. {
105 | apply Rmult_eq_reg_r in Hy; [ | apply sqrt2_neq_0 ].
106 | apply eq_IZR; assumption.
107 | }
108 | apply Rinv_neq_0_compat, pow_nonzero; lra.
109 | } {
110 | apply Rmult_eq_reg_r, eq_IZR in Hz; [ assumption | ].
111 | apply Rinv_neq_0_compat, pow_nonzero; lra.
112 | }
113 | }
114 | simpl in Hr; destruct Hr as (Hr, HN).
115 | rewrite length_app, Nat.add_1_r in HN.
116 | rewrite <- Nat.add_succ_comm in HN.
117 | simpl; destruct e as (t, d).
118 | rewrite fold_right_app; simpl.
119 | rewrite fold_right_app in Hr; simpl in Hr.
120 | destruct t, d. {
121 | apply IHel; split; [ | assumption ].
122 | rewrite <- Hr; simpl.
123 | unfold Rdiv.
124 | progress repeat rewrite Rmult_1_l.
125 | progress repeat rewrite Rmult_0_l.
126 | progress repeat rewrite Rplus_0_l.
127 | progress repeat rewrite Rplus_0_r.
128 | progress repeat rewrite <- Rmult_assoc.
129 | rewrite Rmult5_sqrt2_sqrt5; [ | lra ].
130 | rewrite plus_IZR, plus_IZR.
131 | progress repeat rewrite mult_IZR.
132 | f_equal; f_equal.
133 | all : rewrite Rinv_mult; lra.
134 | } {
135 | apply IHel; split; [ | assumption ].
136 | rewrite <- Hr; simpl.
137 | unfold Rdiv.
138 | progress repeat rewrite Rmult_1_l.
139 | progress repeat rewrite Rmult_0_l.
140 | progress repeat rewrite Rplus_0_l.
141 | progress repeat rewrite Rplus_0_r.
142 | progress repeat rewrite <- Rmult_assoc.
143 | progress repeat rewrite mult_IZR.
144 | rewrite Rmult5_sqrt2_sqrt5; [ | lra ].
145 | rewrite minus_IZR, plus_IZR.
146 | progress repeat rewrite mult_IZR.
147 | f_equal; f_equal.
148 | all : rewrite Rinv_mult; lra.
149 | } {
150 | apply IHel; split; [ | assumption ].
151 | rewrite <- Hr; simpl.
152 | unfold Rdiv.
153 | progress repeat rewrite Rmult_1_l.
154 | progress repeat rewrite Rmult_0_l.
155 | progress repeat rewrite Rplus_0_l.
156 | progress repeat rewrite Rplus_0_r.
157 | progress repeat rewrite <- Rmult_assoc.
158 | progress repeat rewrite mult_IZR.
159 | rewrite Rmult5_sqrt2_sqrt5; [ | lra ].
160 | rewrite plus_IZR, minus_IZR.
161 | progress repeat rewrite mult_IZR.
162 | f_equal; f_equal.
163 | all : rewrite Rinv_mult; lra.
164 | } {
165 | apply IHel; split; [ | assumption ].
166 | rewrite <- Hr; simpl.
167 | unfold Rdiv.
168 | progress repeat rewrite Rmult_1_l.
169 | progress repeat rewrite Rmult_0_l.
170 | progress repeat rewrite Rplus_0_l.
171 | progress repeat rewrite Rplus_0_r.
172 | progress repeat rewrite <- Rmult_assoc.
173 | progress repeat rewrite mult_IZR.
174 | rewrite Rmult5_sqrt2_sqrt5; [ | lra ].
175 | rewrite minus_IZR, plus_IZR.
176 | progress repeat rewrite mult_IZR.
177 | f_equal; f_equal.
178 | all : rewrite Rinv_mult; lra.
179 | }
180 | Qed.
181 |
182 | Theorem rotate_prop : ∀ p t d el el₁ el₂ e a b c,
183 | t = lb ∧ p = (1, 0, 0, O)%Z ∨
184 | t = la ∧ p = (0, 0, 1, O)%Z
185 | → el₁ = el₂ ++ [FE t d]
186 | → el = e :: el₁
187 | → norm_list el = el
188 | → fold_right rotate_param p el₁ = (a, b, c, length el₁)
189 | → (b mod 3)%Z ≠ 0%Z
190 | → match e with
191 | | ạ => ((b - 2 * c) mod 3)%Z ≠ 0%Z
192 | | ạ⁻¹ => ((b + 2 * c) mod 3)%Z ≠ 0%Z
193 | | ḅ => ((b + 2 * a) mod 3)%Z ≠ 0%Z
194 | | ḅ⁻¹ => ((b - 2 * a) mod 3)%Z ≠ 0%Z
195 | end.
196 | Proof.
197 | intros p t d el el₁ el₂ e a b c.
198 | intros Htp Hel₁ Hel Hn Hp Hb'.
199 | rewrite Hel₁ in Hp at 2; simpl in Hp.
200 | remember (length el₂) as len eqn:Hlen.
201 | destruct el₂ as [| e₁]. {
202 | simpl in Hlen; subst len; simpl in Hel.
203 | subst el₁; simpl in Hp.
204 | destruct Htp as [(Ht, Hq)| (Ht, Hq)]; subst t p. {
205 | destruct d; injection Hp; intros; subst. {
206 | destruct e as (t₁, d₁); destruct t₁, d₁; intros H; try discriminate H.
207 | revert Hn; apply norm_list_no_start.
208 | }
209 | destruct e as (t₁, d₁); destruct t₁, d₁; intros H; try discriminate H.
210 | revert Hn; apply norm_list_no_start.
211 | }
212 | destruct d; injection Hp; intros; subst. {
213 | destruct e as (t₁, d₁); destruct t₁, d₁; intros H; try discriminate H.
214 | revert Hn; apply norm_list_no_start.
215 | }
216 | destruct e as (t₁, d₁); destruct t₁, d₁; intros H; try discriminate H.
217 | revert Hn; apply norm_list_no_start.
218 | }
219 | rewrite Hel₁ in Hel; simpl in Hel.
220 | generalize Hn; intros H₂.
221 | do 2 rewrite app_comm_cons in Hel.
222 | rewrite Hel in H₂.
223 | apply norm_list_app_diag in H₂.
224 | destruct len; [ discriminate Hlen | ].
225 | simpl in Hlen; apply eq_add_S in Hlen.
226 | rewrite Hel₁, fold_right_app in Hp.
227 | simpl in Hp.
228 | remember (rotate_param (FE t d) p) as p₁ eqn:Hp₁.
229 | remember (fold_right rotate_param p₁ el₂) as p' eqn:Hp'.
230 | symmetry in Hp'.
231 | destruct p' as (((a', b'), c'), N').
232 | simpl in Hp.
233 | destruct e₁ as (t₁, d₁); destruct t₁, d₁. {
234 | injection Hp; clear Hp; intros HN Hc Hb Ha; subst a b c N'.
235 | destruct e as (t₂, d₂); destruct t₂, d₂. {
236 | rewrite <- Z.mod_add with (b := (3 * b')%Z); [ | intros; discriminate ].
237 | remember (b' + 2 * c' + 2 * (-4 * b' + c') + 3 * b' * 3)%Z as x eqn:Hx.
238 | ring_simplify in Hx; subst x.
239 | replace 4%Z with (2 * 2)%Z by reflexivity.
240 | rewrite <- Z.mul_assoc, <- Z.mul_add_distr_l.
241 | intros H; apply Hb'.
242 | apply Znumtheory.Zmod_divide in H; [ | intros; discriminate ].
243 | apply Z.gauss in H; [ | reflexivity ].
244 | destruct H as (k, H); rewrite H.
245 | apply Z.mod_mul; intros; discriminate.
246 | } {
247 | exfalso; revert Hn; rewrite Hel.
248 | apply norm_list_no_start.
249 | } {
250 | rewrite Z.mul_assoc, Z.mul_shuffle0.
251 | unfold Z.sub; rewrite Zopp_mult_distr_l.
252 | rewrite Z.mod_add; [ assumption | intros H; discriminate H ].
253 | }
254 | rewrite Z.mul_assoc, Z.mul_shuffle0.
255 | rewrite Z.mod_add; [ assumption | intros H; discriminate H ].
256 | } {
257 | injection Hp; clear Hp; intros HN Hc Hb Ha; subst a b c N'.
258 | destruct e as (t₂, d₂); destruct t₂, d₂. {
259 | exfalso; revert Hn; rewrite Hel.
260 | apply norm_list_no_start.
261 | } {
262 | rewrite <- Z.mod_add with (b := (3 * b')%Z); [ | intros; discriminate ].
263 | remember (b' - 2 * c' - 2 * (4 * b' + c') + 3 * b' * 3)%Z as x eqn:Hx.
264 | ring_simplify in Hx; subst x.
265 | replace 4%Z with (2 * 2)%Z by reflexivity.
266 | rewrite <- Z.mul_assoc, <- Z.mul_sub_distr_l.
267 | intros H; apply Hb'.
268 | apply Znumtheory.Zmod_divide in H; [ | intros; discriminate ].
269 | apply Z.gauss in H; [ | reflexivity ].
270 | destruct H as (k, H); rewrite H.
271 | apply Z.mod_mul; intros; discriminate.
272 | } {
273 | rewrite Z.mul_assoc, Z.mul_shuffle0.
274 | unfold Z.sub at 1; rewrite Zopp_mult_distr_l.
275 | rewrite Z.mod_add; [ assumption | intros H; discriminate H ].
276 | } {
277 | rewrite Z.mul_assoc, Z.mul_shuffle0.
278 | rewrite Z.mod_add; [ assumption | intros H; discriminate H ].
279 | }
280 | } {
281 | injection Hp; clear Hp; intros HN Hc Hb Ha; subst a b c N'.
282 | destruct e as (t₂, d₂); destruct t₂, d₂. {
283 | rewrite Z.mul_assoc, Z.mul_shuffle0.
284 | rewrite Z.mod_add; [ assumption | intros H; discriminate H ].
285 | } {
286 | rewrite Z.mul_assoc, Z.mul_shuffle0.
287 | unfold Z.sub at 1; rewrite Zopp_mult_distr_l.
288 | rewrite Z.mod_add; [ assumption | intros H; discriminate H ].
289 | } {
290 | rewrite <- Z.mod_add with (b := (3 * b')%Z); [ | intros; discriminate ].
291 | remember (b' - 2 * a' - 2 * (a' + 4 * b') + 3 * b' * 3)%Z as x eqn:Hx.
292 | ring_simplify in Hx; subst x.
293 | replace 4%Z with (2 * 2)%Z by reflexivity.
294 | rewrite <- Z.mul_assoc, <- Z.mul_sub_distr_l.
295 | intros H; apply Hb'.
296 | apply Znumtheory.Zmod_divide in H; [ | intros; discriminate ].
297 | apply Z.gauss in H; [ | reflexivity ].
298 | destruct H as (k, H); rewrite H.
299 | apply Z.mod_mul; intros; discriminate.
300 | }
301 | exfalso; revert Hn; rewrite Hel.
302 | apply norm_list_no_start.
303 | }
304 | injection Hp; clear Hp; intros HN Hc Hb Ha; subst a b c N'.
305 | destruct e as (t₂, d₂); destruct t₂, d₂. {
306 | rewrite Z.mul_assoc, Z.mul_shuffle0.
307 | rewrite Z.mod_add; [ assumption | intros H; discriminate H ].
308 | } {
309 | rewrite Z.mul_assoc, Z.mul_shuffle0.
310 | unfold Z.sub; rewrite Zopp_mult_distr_l.
311 | rewrite Z.mod_add; [ assumption | intros H; discriminate H ].
312 | } {
313 | exfalso; revert Hn; rewrite Hel.
314 | apply norm_list_no_start.
315 | }
316 | rewrite <- Z.mod_add with (b := (3 * b')%Z); [ | intros; discriminate ].
317 | remember (b' + 2 * a' + 2 * (a' - 4 * b') + 3 * b' * 3)%Z as x eqn:Hx.
318 | ring_simplify in Hx; subst x.
319 | replace 4%Z with (2 * 2)%Z by reflexivity.
320 | rewrite <- Z.mul_assoc, <- Z.mul_add_distr_l.
321 | intros H; apply Hb'.
322 | apply Znumtheory.Zmod_divide in H; [ | intros; discriminate ].
323 | apply Z.gauss in H; [ | reflexivity ].
324 | destruct H as (k, H); rewrite H.
325 | apply Z.mod_mul; intros; discriminate.
326 | Qed.
327 |
328 | Theorem rotate_param_b_nonzero : ∀ p t d el el₁ a b c,
329 | t = lb ∧ p = (1, 0, 0, O)%Z ∨
330 | t = la ∧ p = (0, 0, 1, O)%Z
331 | → el = el₁ ++ [FE t d]
332 | → norm_list el = el
333 | → fold_right rotate_param p el = (a, b, c, length el)
334 | → (b mod 3 ≠ 0)%Z.
335 | Proof.
336 | intros p t d el el₁ a b c Htp Hel Hn Hu.
337 | remember (length el₁) as len eqn:Hlen; symmetry in Hlen.
338 | revert el el₁ d a b c Hel Hn Hu Hlen.
339 | induction len as (len, IHlen) using lt_wf_rec; intros.
340 | destruct len. {
341 | apply length_zero_iff_nil in Hlen; subst el₁.
342 | subst el; simpl in Hu.
343 | destruct Htp as [(Ht, Hp)| (Ht, Hp)]; subst t p.
344 | 1,2 : destruct d; injection Hu; intros; subst; intros H; discriminate H.
345 | }
346 | destruct el₁ as [| e₁]; [ discriminate Hlen | simpl in Hlen ].
347 | apply eq_add_S in Hlen.
348 | rewrite <- app_comm_cons in Hel.
349 | remember (el₁ ++ [FE t d]) as el₂ eqn:Hel₁.
350 | generalize Hn; intros H₁; rewrite Hel in H₁.
351 | apply norm_list_cons in H₁.
352 | rewrite Hel in Hu; simpl in Hu.
353 | remember (fold_right rotate_param p el₂) as v eqn:Hp.
354 | symmetry in Hp.
355 | destruct v as (((a', b'), c'), N').
356 | assert (Hss : (len < S len)%nat) by apply Nat.lt_succ_diag_r.
357 | assert (N' = S len); [ | subst N' ]. {
358 | destruct e₁ as (t₁, d₁).
359 | rewrite Hel₁, length_app, Nat.add_1_r in Hu.
360 | destruct t₁, d₁; simpl in Hu; injection Hu; intros; subst; reflexivity.
361 | }
362 | rewrite <- Hlen in Hp.
363 | replace (S (length el₁)) with (length el₂) in Hp. {
364 | pose proof IHlen _ Hss _ _ _ _ _ _ Hel₁ H₁ Hp Hlen as Hb'; subst len.
365 | pose proof rotate_prop p t d el el₂ el₁ e₁ a' b' c' Htp Hel₁ Hel Hn Hp Hb'.
366 | destruct e₁ as (t₁, d₁).
367 | destruct t₁, d₁; injection Hu; intros; subst; assumption.
368 | }
369 | subst; rewrite length_app, Nat.add_1_r; reflexivity.
370 | Qed.
371 |
372 | Theorem rotate_0_0_1_b_nonzero : ∀ el el₁ d,
373 | el = el₁ ++ [FE la d]
374 | → norm_list el = el
375 | → ∃ a b c k,
376 | (mat_of_path el * V 0 0 1)%vec =
377 | V (IZR a/3^k) (IZR b*√2/3^k) (IZR c/3^k) ∧
378 | (b mod 3 ≠ 0)%Z.
379 | Proof.
380 | intros el el₁ d Hel Hn.
381 | remember (fold_right rotate_param (0, 0, 1, O)%Z el) as u eqn:Hu.
382 | symmetry in Hu; destruct u as (((a, b), c), len).
383 | generalize Hu; intros Hv.
384 | apply rotate_param_rotate in Hv; simpl in Hv.
385 | rewrite rotate_vec_mul in Hv.
386 | rewrite Rmult_0_l, Rdiv_0_l, Rdiv_1_r in Hv.
387 | destruct Hv as (Hv, Hlen).
388 | rewrite Hv.
389 | exists a, b, c, len.
390 | split; [ reflexivity | clear Hv ].
391 | symmetry in Hlen.
392 | rewrite Hel in Hlen; simpl in Hlen.
393 | rewrite length_app, Nat.add_1_r in Hlen.
394 | destruct len; [ discriminate Hlen | ].
395 | apply eq_add_S in Hlen; subst len.
396 | replace (S (length el₁)) with (length el) in Hu. {
397 | eapply rotate_param_b_nonzero; try eassumption.
398 | right; split; reflexivity.
399 | }
400 | subst; rewrite length_app, Nat.add_1_r; reflexivity.
401 | Qed.
402 |
403 | Theorem rotate_1_0_0_b_nonzero : ∀ el el₁ d,
404 | el = el₁ ++ [FE lb d]
405 | → norm_list el = el
406 | → ∃ a b c k,
407 | (mat_of_path el * V 1 0 0)%vec =
408 | V (IZR a/3^k) (IZR b*√2/3^k) (IZR c/3^k) ∧
409 | (b mod 3 ≠ 0)%Z.
410 | Proof.
411 | intros el el₁ d Hel Hn.
412 | remember (fold_right rotate_param (1, 0, 0, O)%Z el) as u eqn:Hu.
413 | symmetry in Hu; destruct u as (((a, b), c), len).
414 | generalize Hu; intros Hv.
415 | apply rotate_param_rotate in Hv; simpl in Hv.
416 | rewrite rotate_vec_mul in Hv.
417 | rewrite Rmult_0_l, Rdiv_0_l, Rdiv_1_r in Hv.
418 | destruct Hv as (Hv, Hlen).
419 | rewrite Hv.
420 | exists a, b, c, len.
421 | split; [ reflexivity | clear Hv ].
422 | symmetry in Hlen.
423 | rewrite Hel in Hlen; simpl in Hlen.
424 | rewrite length_app, Nat.add_1_r in Hlen.
425 | destruct len; [ discriminate Hlen | ].
426 | apply eq_add_S in Hlen; subst len.
427 | replace (S (length el₁)) with (length el) in Hu. {
428 | eapply rotate_param_b_nonzero; try eassumption.
429 | left; split; reflexivity.
430 | }
431 | subst; rewrite length_app, Nat.add_1_r; reflexivity.
432 | Qed.
433 |
434 | Theorem rotate_1_0_0_is_diff : ∀ el el₁ d,
435 | el = el₁ ++ [FE lb d]
436 | → norm_list el = el
437 | → (mat_of_path el * V 1 0 0)%vec ≠ V 1 0 0.
438 | Proof.
439 | intros el el₁ d Hel Hn.
440 | pose proof rotate_1_0_0_b_nonzero el el₁ d Hel Hn as H.
441 | destruct H as (a, (b, (c, (k, (Hp, Hm))))).
442 | rewrite Hp; intros H.
443 | injection H; intros Hc Hb Ha.
444 | apply Rmult_integral in Hb.
445 | destruct Hb as [Hb| Hb]. {
446 | apply Rmult_integral in Hb.
447 | destruct Hb as [Hb| Hb]. {
448 | apply eq_IZR_R0 in Hb; subst b; apply Hm; reflexivity.
449 | }
450 | revert Hb; apply sqrt2_neq_0.
451 | }
452 | apply Rmult_eq_compat_r with (r := (3 ^ k)%R) in Hb.
453 | rewrite Rinv_l in Hb; [ lra | apply pow_nonzero; lra ].
454 | Qed.
455 |
456 | Theorem rotate_0_0_1_is_diff : ∀ el el₁ d,
457 | el = el₁ ++ [FE la d]
458 | → norm_list el = el
459 | → (mat_of_path el * V 0 0 1)%vec ≠ V 0 0 1.
460 | Proof.
461 | intros el el₁ d Hel Hn.
462 | specialize (rotate_0_0_1_b_nonzero el el₁ d Hel Hn) as H.
463 | destruct H as (a, (b, (c, (k, (Hp, Hm))))).
464 | rewrite Hp; intros H.
465 | injection H; intros Hc Hb Ha.
466 | apply Rmult_integral in Hb.
467 | destruct Hb as [Hb| Hb]. {
468 | apply Rmult_integral in Hb.
469 | destruct Hb as [Hb| Hb]. {
470 | apply eq_IZR_R0 in Hb; subst b; apply Hm; reflexivity.
471 | }
472 | revert Hb; apply sqrt2_neq_0.
473 | }
474 | apply Rmult_eq_compat_r with (r := (3 ^ k)%R) in Hb.
475 | rewrite Rinv_l in Hb; [ lra | apply pow_nonzero; lra ].
476 | Qed.
477 |
478 | Theorem matrix_of_non_empty_path_is_not_identity : ∀ el,
479 | norm_list el ≠ []
480 | → mat_of_path el ≠ mat_id.
481 | Proof.
482 | intros * Hn.
483 | remember (rev_path (norm_list el)) as el₁ eqn:Hel₁.
484 | symmetry in Hel₁.
485 | destruct el₁ as [| e₁ el₁]; [ now apply rev_path_is_nil in Hel₁ | ].
486 | apply (f_equal rev_path) in Hel₁.
487 | rewrite rev_path_involutive in Hel₁.
488 | rewrite rev_path_cons, rev_path_single in Hel₁.
489 | destruct e₁ as (t, d).
490 | intros H.
491 | destruct t. {
492 | apply rotate_0_0_1_is_diff in Hel₁; [ | apply norm_list_idemp ].
493 | now rewrite mat_of_path_norm, H, mat_vec_mul_id in Hel₁.
494 | } {
495 | apply rotate_1_0_0_is_diff in Hel₁; [ | apply norm_list_idemp ].
496 | now rewrite mat_of_path_norm, H, mat_vec_mul_id in Hel₁.
497 | }
498 | Qed.
499 |
500 | Definition is_a_rotation_π M := M = mat_transp M ∧ M ≠ mat_id.
501 |
502 | Theorem mat_of_path_is_not_rotation_π : ∀ el,
503 | ¬ is_a_rotation_π (mat_of_path el).
504 | Proof.
505 | intros el H.
506 | unfold is_a_rotation_π in H.
507 | destruct H as (Hmt, Hid).
508 | remember (mat_of_path el) as M eqn:HM.
509 | apply Hid; clear Hid.
510 | assert (Hr : is_rotation_matrix M). {
511 | subst M.
512 | apply mat_of_path_is_rotation_matrix.
513 | }
514 | assert (HMI : (M * M = mat_id)%mat). {
515 | rewrite Hmt at 2.
516 | now destruct Hr.
517 | }
518 | rewrite <- mat_of_path_norm in HM.
519 | remember (norm_list el) as nel eqn:Hnel.
520 | symmetry in Hnel.
521 | destruct nel as [| e nel]; [ easy | ].
522 | rewrite HM in HMI.
523 | rewrite <- mat_of_path_app in HMI.
524 | exfalso; revert HMI.
525 | apply matrix_of_non_empty_path_is_not_identity.
526 | rewrite <- Hnel.
527 | intros H.
528 | apply norm_list_app_is_nil in H. {
529 | symmetry in H; apply rev_path_eq_path in H.
530 | now rewrite Hnel in H.
531 | } {
532 | now rewrite norm_list_idemp.
533 | } {
534 | now rewrite norm_list_idemp.
535 | }
536 | Qed.
537 |
538 | Theorem mat_of_path_elem_pow : ∀ e n,
539 | (mat_of_path [e] ^ n)%mat = mat_of_path (repeat e n).
540 | Proof.
541 | intros.
542 | induction n; [ easy | simpl ].
543 | rewrite IHn; simpl.
544 | unfold mat_of_path; simpl.
545 | now rewrite mat_mul_id_r.
546 | Qed.
547 |
--------------------------------------------------------------------------------
/Orbit.v:
--------------------------------------------------------------------------------
1 | (* Banach-Tarski paradox. *)
2 |
3 | From Stdlib Require Import Utf8 List Relations.
4 | From Stdlib Require Import Reals Nsatz.
5 | Import ListNotations.
6 |
7 | Require Import Misc MiscReals Words Normalize Reverse Matrix Pset.
8 |
9 | Definition same_orbit x y := ∃ el, (mat_of_path el * x)%vec = y.
10 |
11 | Theorem same_orbit_refl : reflexive _ same_orbit.
12 | Proof. now exists []; rewrite mat_vec_mul_id. Qed.
13 |
14 | Theorem same_orbit_sym : symmetric _ same_orbit.
15 | Proof.
16 | intros p₁ p₂ (el, H); simpl in H.
17 | exists (rev_path el).
18 | revert p₁ p₂ H.
19 | induction el as [| e]; intros; [ now rewrite mat_vec_mul_id in H |-* | ].
20 | rewrite rev_path_cons, mat_of_path_app, mat_vec_mul_assoc.
21 | apply IHel; rewrite <- H, <- mat_vec_mul_assoc.
22 | rewrite <- mat_of_path_app, rev_path_single; simpl.
23 | now rewrite mat_of_path_negf.
24 | Qed.
25 |
26 | Theorem same_orbit_trans : transitive _ same_orbit.
27 | Proof.
28 | intros p₁ p₂ p₃ (el₁, H₁) (el₂, H₂); simpl in H₁, H₂.
29 | exists (el₂ ++ el₁).
30 | now rewrite mat_of_path_app, mat_vec_mul_assoc, H₁, H₂.
31 | Qed.
32 |
33 | Add Parametric Relation : _ same_orbit
34 | reflexivity proved by same_orbit_refl
35 | symmetry proved by same_orbit_sym
36 | transitivity proved by same_orbit_trans
37 | as same_orbit_rel.
38 |
39 | Definition equiv_same_orbit : equiv vector same_orbit :=
40 | conj same_orbit_refl (conj same_orbit_trans same_orbit_sym).
41 |
42 | Definition not_in_fixpoints :=
43 | mkset (λ p, ∀ el, norm_list el ≠ [] → fold_right rotate p el ≠ p).
44 |
45 | Theorem not_in_fixpoints_one_path : ∀ f p e₁ e₂ el el₂ el₁ el₃,
46 | p ∈ not_in_fixpoints
47 | → (mat_of_path el * p)%vec = f p
48 | → (mat_of_path el₁ * (f p))%vec = p
49 | → norm_list el = el₂ ++ [e₁]
50 | → norm_list el₁ = e₂ :: el₃
51 | → e₂ ≠ negf e₁
52 | → False.
53 | Proof.
54 | intros f p e₁ e₂ el el₂ el₁ el₃ Hnf Hel H₆ H₂ H₄ Hd.
55 | rewrite rotate_rotate_norm in Hel, H₆.
56 | rewrite <- Hel in H₆.
57 | rewrite <- mat_vec_mul_assoc in H₆.
58 | rewrite <- mat_of_path_app in H₆.
59 | revert H₆.
60 | rewrite <- rotate_vec_mul.
61 | apply Hnf.
62 | intros H.
63 | apply norm_list_app_is_nil in H. {
64 | rewrite H₄, H₂ in H.
65 | apply rev_path_eq_eq in H.
66 | rewrite rev_path_involutive, rev_path_app in H.
67 | apply not_eq_sym in Hd.
68 | now injection H.
69 | } {
70 | now rewrite norm_list_idemp.
71 | } {
72 | now rewrite norm_list_idemp.
73 | }
74 | Qed.
75 |
76 | Definition orbit_selector := choice_function same_orbit.
77 |
78 | Definition sphere r := mkset (λ '(V x y z), (x² + y² + z² = r²)%R).
79 | Definition ball := mkset (λ '(V x y z), (x² + y² + z² <= 1)%R).
80 |
81 | Theorem on_sphere_norm : ∀ p r, (0 ≤ r)%R → p ∈ sphere r ↔ ‖p‖ = r.
82 | Proof.
83 | intros (x, y, z) r Hr; simpl.
84 | split; intros Hp. {
85 | now rewrite Hp; apply sqrt_Rsqr.
86 | }
87 | apply (f_equal Rsqr) in Hp.
88 | rewrite Rsqr_sqrt in Hp; [ easy | ].
89 | apply nonneg_sqr_vec_norm.
90 | Qed.
91 |
92 | Theorem in_its_sphere : ∀ v, v ∈ sphere ‖v‖.
93 | Proof.
94 | intros (x, y, z); simpl.
95 | rewrite Rsqr_sqrt; [ easy | ].
96 | apply nonneg_sqr_vec_norm.
97 | Qed.
98 |
99 | Theorem on_sphere_after_rotation : ∀ p m r,
100 | p ∈ sphere r
101 | → is_rotation_matrix m
102 | → (m * p)%vec ∈ sphere r.
103 | Proof.
104 | intros * His Hm.
105 | destruct p as (x, y, z).
106 | unfold sphere in His; simpl in His.
107 | unfold sphere; simpl.
108 | unfold is_rotation_matrix in Hm.
109 | destruct Hm as (Hm, Hd).
110 | unfold mat_det in Hd.
111 | unfold mat_mul, mat_id in Hm; simpl in Hm.
112 | injection Hm; clear Hm; intros H₁ H₂ H₃ H₄ H₅ H₆ H₇ H₈ H₉.
113 | progress unfold Rsqr in His.
114 | progress unfold Rsqr.
115 | nsatz.
116 | Qed.
117 |
118 | Theorem in_ball_after_rotation : ∀ p m,
119 | p ∈ ball
120 | → is_rotation_matrix m
121 | → mat_vec_mul m p ∈ ball.
122 | Proof.
123 | intros * His Hrm.
124 | destruct p as (x, y, z).
125 | remember (V x y z) as p eqn:HP.
126 | remember (x² + y² + z²)%R as r eqn:Hr; symmetry in Hr.
127 | assert (Hos : p ∈ sphere (√ r)). {
128 | subst p; simpl; rewrite Rsqr_sqrt; [ easy | subst r ].
129 | apply nonneg_sqr_vec_norm.
130 | }
131 | pose proof on_sphere_after_rotation _ _ _ Hos Hrm as H.
132 | unfold ball in His.
133 | unfold sphere in H.
134 | unfold ball.
135 | subst p; simpl in *.
136 | now rewrite H, <- Hos.
137 | Qed.
138 |
139 | Theorem in_ball_after_rotate : ∀ p e,
140 | p ∈ ball
141 | → rotate e p ∈ ball.
142 | Proof.
143 | intros * His.
144 | apply in_ball_after_rotation; [ easy | ].
145 | apply rotate_is_rotation_matrix.
146 | Qed.
147 |
--------------------------------------------------------------------------------
/OrbitRepr.v:
--------------------------------------------------------------------------------
1 | (* Banach-Tarski paradox. *)
2 |
3 | From Stdlib Require Import Utf8 List.
4 | From Stdlib Require Import Reals Nsatz.
5 | Import ListNotations.
6 |
7 | Require Import Misc Words Normalize Reverse Matrix Pset Orbit.
8 | Require Import Partition.
9 |
10 | Class sel_model {A} := mkos { os_fun : A → A }.
11 |
12 | Definition orbit_by_seq_of e {os : sel_model} :=
13 | mkset (λ p, ∃ n, (mat_of_path (repeat e (S n)) * os_fun p)%vec = p).
14 |
15 | Definition D :=
16 | mkset
17 | (λ p, ∃ el p₁, same_orbit p p₁
18 | ∧ norm_list el ≠ [] ∧ (mat_of_path el * p₁)%vec = p₁).
19 |
20 | Arguments D : simpl never.
21 |
22 | Definition M {os : sel_model} :=
23 | mkset (λ p, p ∈ ball ∖ D ∧ p = os_fun p).
24 | Definition SS {os : sel_model} e :=
25 | mkset
26 | (λ p,
27 | p ∈ ball ∖ D ∧
28 | ∃ el el₁,
29 | norm_list el = e :: el₁ ∧ fold_right rotate (os_fun p) el = p).
30 | Definition G {os : sel_model} :=
31 | mkset (λ p, p ∈ (ball ∖ D) ∩ orbit_by_seq_of ạ⁻¹).
32 |
33 | Opaque M SS G.
34 |
35 | Definition rot e (E : set vector) :=
36 | mkset (λ p, rotate (negf e) p ∈ E).
37 |
38 | Theorem empty_set_not_full_set : ∀ f os, os = mkos _ f →
39 | ∀ e p, p ∈ M → p ∈ SS e → False.
40 | Proof.
41 | intros f os Hos e p He Hs; subst os.
42 | destruct He as (Hinf & He); simpl in He.
43 | destruct Hs as (Hjnf & el & el₁ & Hn & Hs); simpl in Hs.
44 | rewrite rotate_vec_mul in Hs.
45 | rewrite <- He in Hs.
46 | simpl in Hinf.
47 | destruct Hinf as (Hle1 & Hinf).
48 | apply Hinf; clear Hinf.
49 | exists el, p.
50 | now rewrite Hn.
51 | Qed.
52 |
53 | Theorem start_with_same : ∀ f os, os = mkos _ f →
54 | ∀ e₁ e₂ p, p ∈ SS e₁ → p ∈ SS e₂ → e₁ = e₂.
55 | Proof.
56 | intros f os Hos (ti, di) (tj, dj) p Hsi Hsj; subst os.
57 | destruct Hsi as (Hinf & eli & eli₁ & Hni & Hsi); simpl in Hsi.
58 | destruct Hsj as (Hjnf & elj & elj₁ & Hnj & Hsj); simpl in Hsj.
59 | rewrite rotate_vec_mul in Hsi, Hsj.
60 | eapply rotate_rev_path in Hsj.
61 | destruct ti, tj. {
62 | destruct di, dj; [ easy | exfalso | exfalso | easy ]. {
63 | eapply not_in_fixpoints_one_path; try eassumption. {
64 | intros el Hn.
65 | destruct Hinf as (Hps, Hnpd).
66 | intros H; apply Hnpd; clear Hnpd.
67 | rewrite rotate_vec_mul in H.
68 | now exists el, p.
69 | } {
70 | rewrite <- rev_path_norm_list, Hnj.
71 | now rewrite rev_path_cons, rev_path_single.
72 | }
73 | easy.
74 | }
75 | eapply not_in_fixpoints_one_path; try eassumption. {
76 | intros el Hn.
77 | destruct Hinf as (Hps, Hnpd).
78 | intros H; apply Hnpd.
79 | rewrite rotate_vec_mul in H.
80 | now exists el, p.
81 | } {
82 | rewrite <- rev_path_norm_list, Hnj.
83 | now rewrite rev_path_cons, rev_path_single.
84 | }
85 | easy.
86 | } {
87 | exfalso.
88 | eapply not_in_fixpoints_one_path; try eassumption. {
89 | intros el Hn.
90 | destruct Hinf as (Hps, Hnpd).
91 | intros H; apply Hnpd.
92 | rewrite rotate_vec_mul in H.
93 | now exists el, p.
94 | } {
95 | rewrite <- rev_path_norm_list, Hnj.
96 | now rewrite rev_path_cons, rev_path_single.
97 | }
98 | easy.
99 | } {
100 | exfalso.
101 | eapply not_in_fixpoints_one_path; try eassumption. {
102 | intros el Hn.
103 | destruct Hinf as (Hps, Hnpd).
104 | intros H; apply Hnpd.
105 | rewrite rotate_vec_mul in H.
106 | now exists el, p.
107 | } {
108 | rewrite <- rev_path_norm_list, Hnj.
109 | now rewrite rev_path_cons, rev_path_single.
110 | }
111 | easy.
112 | }
113 | destruct di, dj; [ easy | exfalso | exfalso | easy ]. {
114 | eapply not_in_fixpoints_one_path; try eassumption. {
115 | intros el Hn.
116 | destruct Hinf as (Hps, Hnpd).
117 | intros H; apply Hnpd.
118 | rewrite rotate_vec_mul in H.
119 | now exists el, p.
120 | } {
121 | rewrite <- rev_path_norm_list, Hnj.
122 | now rewrite rev_path_cons, rev_path_single.
123 | }
124 | easy.
125 | } {
126 | eapply not_in_fixpoints_one_path; try eassumption. {
127 | intros el Hn.
128 | destruct Hinf as (Hps, Hnpd).
129 | intros H; apply Hnpd.
130 | rewrite rotate_vec_mul in H.
131 | now exists el, p.
132 | } {
133 | rewrite <- rev_path_norm_list, Hnj.
134 | now rewrite rev_path_cons, rev_path_single.
135 | }
136 | easy.
137 | }
138 | Qed.
139 |
140 | Theorem not_start_with_rot :
141 | ∀ f, orbit_selector f
142 | → ∀ os, os = mkos _ f
143 | → ∀ e p, p ∈ SS e → p ∈ rot e (SS (negf e)) → False.
144 | Proof.
145 | intros f (Hoe, Ho) os Hos e p Hs Hr; simpl in Hr; subst os.
146 | destruct Hs as (Hnf & el & el₁ & Hn & Hs); simpl in Hs.
147 | destruct Hr as (Hrnf & elr & elr₁ & Hnr & Hsr); simpl in Hsr.
148 | assert (Hr : f p = f (rotate (negf e) p)). {
149 | apply Hoe.
150 | exists (negf e :: []).
151 | now rewrite <- rotate_vec_mul.
152 | }
153 | rewrite <- Hr in Hsr.
154 | rewrite rotate_vec_mul in Hsr.
155 | eapply rotate_rev_path in Hsr.
156 | rewrite <- rotate_vec_mul in Hsr.
157 | rewrite <- fold_right_single, <- fold_right_app in Hsr.
158 | rewrite <- Hsr in Hs.
159 | rewrite <- fold_right_app in Hs.
160 | rewrite rotate_vec_mul in Hs.
161 | rewrite rotate_rotate_norm in Hs.
162 | pose proof is_normal [] el (rev_path elr ++ [negf e]) as H.
163 | do 2 rewrite app_nil_l in H.
164 | rewrite <- H in Hs; clear H.
165 | rewrite <- is_normal in Hs.
166 | rewrite <- rev_path_norm_list in Hs.
167 | rewrite Hnr in Hs.
168 | rewrite <- rotate_rotate_norm in Hs.
169 | rewrite rev_path_cons in Hs.
170 | rewrite rev_path_single in Hs.
171 | rewrite <- app_assoc in Hs.
172 | simpl in Hs.
173 | rewrite negf_involutive in Hs.
174 | rewrite app_assoc in Hs.
175 | rewrite rotate_cancel_in in Hs.
176 | rewrite app_nil_r in Hs.
177 | destruct Hnf as (Hps, Hnpd).
178 | apply Hnpd.
179 | exists (norm_list el ++ rev_path elr₁), p.
180 | split; [ easy | ].
181 | split; [ | easy ].
182 | intros H.
183 | apply norm_list_app_is_nil in H. {
184 | apply -> rev_path_eq_eq in H.
185 | rewrite H, Hn in Hnr.
186 | revert Hnr; apply norm_list_no_start2.
187 | } {
188 | symmetry; apply norm_list_idemp.
189 | }
190 | rewrite <- rev_path_norm_list; eapply norm_list_is_cons in Hnr.
191 | now rewrite Hnr.
192 | Qed.
193 |
194 | Theorem decompose_2a_contrad_case :
195 | ∀ f, orbit_selector f
196 | → ∀ os, os = mkos _ f
197 | → ∀ p, p ∈ M ∪ SS ạ ∪ G
198 | → p ∈ rot ạ (SS ạ⁻¹ ∖ G)
199 | → False.
200 | Proof.
201 | intros * (Hoe, Ho) * Hos * Hi Hj.
202 | assert (Hfr : f (rotate ạ⁻¹ p) = f p). {
203 | apply Hoe; exists (ạ :: []); simpl.
204 | rewrite <- rotate_vec_mul; simpl.
205 | apply rotate_neg_rotate.
206 | }
207 | destruct Hj as (Hs & Hb); simpl in Hs, Hb; apply Hb; clear Hb.
208 | split; [ now destruct Hs | ].
209 | destruct Hi as [[Hi| Hi] | Hi]. {
210 | destruct Hs as (Hrnf & el & el₁ & Hn & Hr).
211 | destruct Hi as (Hnf & Hp); subst os; simpl in Hp.
212 | exists O; simpl.
213 | rewrite Hfr, <- Hp.
214 | apply mat_of_path_single.
215 | } {
216 | eapply not_start_with_rot in Hi; try eassumption; [ easy | ].
217 | now split.
218 | } {
219 | destruct Hi as (Hnf, Hoo).
220 | destruct Hoo as (n, Hoo).
221 | unfold orbit_by_seq_of.
222 | remember S as g; subst os; simpl in Hoo; simpl; subst g.
223 | rewrite Hfr; simpl.
224 | exists (S n).
225 | rewrite mat_of_path_cons, mat_vec_mul_assoc, Hoo.
226 | rewrite <- mat_of_path_single.
227 | unfold mat_of_path; simpl.
228 | now rewrite mat_mul_id_r.
229 | }
230 | Qed.
231 |
232 | Theorem r_decomposed_5 :
233 | ∀ f, orbit_selector f
234 | → ∀ os, os = mkos _ f
235 | → is_partition (ball ∖ D) [M; SS ạ; SS ạ⁻¹; SS ḅ; SS ḅ⁻¹].
236 | Proof.
237 | intros f (Hoe, Ho) os Hos; subst os.
238 | split. {
239 | intros p.
240 | split. {
241 | intros Hnf.
242 | unfold set_union_list; simpl; unfold set_union.
243 | destruct (vec_eq_dec p (f p)) as [H₁| H₁]; [ left; now split | ].
244 | right.
245 | pose proof Ho p as H.
246 | destruct H as (el, Hel).
247 | remember (norm_list el) as el₁ eqn:Hel₁; symmetry in Hel₁.
248 | destruct (list_nil_app_dec el₁) as [H₂| (e & el₂ & H₂)]; subst el₁. {
249 | rewrite rotate_rotate_norm, H₂ in Hel.
250 | now rewrite mat_vec_mul_id in Hel.
251 | }
252 | destruct e as (t, d); destruct t, d. {
253 | left; split; [ easy | ].
254 | exists (rev_path el), (rev_path el₂).
255 | rewrite rotate_vec_mul.
256 | split; [ | now apply rotate_rev_path ].
257 | now rewrite <- rev_path_norm_list, H₂, rev_path_app.
258 | } {
259 | right; left; split; [ easy | ].
260 | exists (rev_path el), (rev_path el₂).
261 | rewrite rotate_vec_mul.
262 | split; [ | now apply rotate_rev_path ].
263 | now rewrite <- rev_path_norm_list, H₂, rev_path_app.
264 | } {
265 | right; right; left; split; [ easy | ].
266 | exists (rev_path el), (rev_path el₂).
267 | rewrite rotate_vec_mul.
268 | split; [ | now apply rotate_rev_path ].
269 | now rewrite <- rev_path_norm_list, H₂, rev_path_app.
270 | } {
271 | right; right; right; left; split; [ easy | ].
272 | exists (rev_path el), (rev_path el₂).
273 | rewrite rotate_vec_mul.
274 | split; [ | now apply rotate_rev_path ].
275 | now rewrite <- rev_path_norm_list, H₂, rev_path_app.
276 | }
277 | }
278 | intros Hul.
279 | unfold set_union_list in Hul; simpl in Hul; unfold set_union in Hul.
280 | destruct Hul as [Hul| [Hul| [Hul| [Hul| [Hul| Hul]]]]]. {
281 | destruct Hul as (Hnf, Hul); simpl in Hul.
282 | now apply Hnf.
283 | } {
284 | destruct Hul as (Hnf, Hul); simpl in Hul.
285 | now apply Hnf.
286 | } {
287 | destruct Hul as (Hnf, Hul); simpl in Hul.
288 | now apply Hnf.
289 | } {
290 | destruct Hul as (Hnf, Hul); simpl in Hul.
291 | now apply Hnf.
292 | } {
293 | destruct Hul as (Hnf, Hul); simpl in Hul.
294 | now apply Hnf.
295 | }
296 | easy.
297 | }
298 | intros i j Hij p.
299 | split; [ | easy ].
300 | intros (Hi, Hj).
301 | destruct i; [ simpl in Hi | ]. {
302 | destruct j; [ exfalso; now apply Hij | clear Hij ].
303 | destruct Hi as (Hinf & Hi); simpl in Hi.
304 | destruct j. {
305 | eapply empty_set_not_full_set; [ easy | | eassumption ].
306 | easy.
307 | }
308 | destruct j. {
309 | eapply empty_set_not_full_set; [ easy | | eassumption ].
310 | easy.
311 | }
312 | destruct j. {
313 | eapply empty_set_not_full_set; [ easy | | eassumption ].
314 | easy.
315 | }
316 | destruct j; [ | now destruct j ].
317 | eapply empty_set_not_full_set; [ easy | | eassumption ].
318 | easy.
319 | }
320 | destruct i; [ simpl in Hi | ]. {
321 | destruct j; [ clear Hij | ]. {
322 | eapply empty_set_not_full_set; [ easy | eassumption | eassumption ].
323 | }
324 | destruct j; [ exfalso; now apply Hij | clear Hij ].
325 | destruct j; [ simpl in Hj | ]. {
326 | eapply start_with_same in Hi; [ | easy | eassumption ].
327 | easy.
328 | }
329 | destruct j; [ simpl in Hj | ]. {
330 | eapply start_with_same in Hi; [ | easy | eassumption ].
331 | easy.
332 | }
333 | destruct j; [ simpl in Hj | now destruct j ].
334 | eapply start_with_same in Hi; [ | easy | eassumption ].
335 | easy.
336 | }
337 | destruct i; [ simpl in Hi | ]. {
338 | destruct j; [ clear Hij | ]. {
339 | eapply empty_set_not_full_set; [ easy | | ]; eassumption.
340 | }
341 | destruct j; [ simpl in Hj | ]. {
342 | eapply start_with_same in Hi; [ | easy | eassumption ].
343 | easy.
344 | }
345 | destruct j; [ exfalso; now apply Hij | clear Hij ].
346 | destruct j; [ simpl in Hj | ]. {
347 | eapply start_with_same in Hi; [ | easy | eassumption ].
348 | easy.
349 | }
350 | destruct j; [ simpl in Hj | now destruct j ].
351 | eapply start_with_same in Hi; [ | easy | eassumption ].
352 | easy.
353 | }
354 | destruct i; [ simpl in Hi | ]. {
355 | destruct j; [ clear Hij | ]. {
356 | eapply empty_set_not_full_set; [ easy | | ]; eassumption.
357 | }
358 | destruct j; [ simpl in Hj | ]. {
359 | eapply start_with_same in Hi; [ | easy | eassumption ].
360 | easy.
361 | }
362 | destruct j; [ simpl in Hj | ]. {
363 | eapply start_with_same in Hi; [ | easy | eassumption ].
364 | easy.
365 | }
366 | destruct j; [ exfalso; now apply Hij | clear Hij ].
367 | destruct j; [ simpl in Hj | now destruct j ].
368 | eapply start_with_same in Hi; [ | easy | eassumption ].
369 | easy.
370 | }
371 | destruct i; [ simpl in Hi | now destruct i ].
372 | destruct j; [ clear Hij | ]. {
373 | eapply empty_set_not_full_set; [ easy | | ]; eassumption.
374 | }
375 | destruct j; [ simpl in Hj | ]. {
376 | eapply start_with_same in Hi; [ | easy | eassumption ].
377 | easy.
378 | }
379 | destruct j; [ simpl in Hj | ]. {
380 | eapply start_with_same in Hi; [ | easy | eassumption ].
381 | easy.
382 | }
383 | destruct j; [ simpl in Hj | ]. {
384 | eapply start_with_same in Hi; [ | easy | eassumption ].
385 | easy.
386 | }
387 | destruct j; [ exfalso; now apply Hij | now destruct j ].
388 | Qed.
389 |
390 | Theorem r_decomposed_4 :
391 | ∀ f, orbit_selector f
392 | → ∀ os, os = mkos _ f
393 | → is_partition (ball ∖ D)
394 | [M ∪ SS ạ ∪ G; SS ạ⁻¹ ∖ G; SS ḅ; SS ḅ⁻¹].
395 | Proof.
396 | intros f HoeHo os Hos.
397 | pose proof r_decomposed_5 f HoeHo os Hos as H.
398 | destruct HoeHo as (Hoe, Ho).
399 | eapply is_partition_group_first_2_together in H.
400 | apply is_partition_union_sub; [ easy | ].
401 | intros p bm; subst os.
402 | destruct bm as (Ha & n & Hr); remember S as g; simpl in Hr; subst g.
403 | split; [ easy | ].
404 | exists (repeat ạ⁻¹ (S n)), (repeat ạ⁻¹ n).
405 | rewrite rotate_vec_mul.
406 | split; [ apply norm_list_repeat | easy ].
407 | Qed.
408 |
409 | Theorem no_fixpoint_after_rotate : ∀ p e, p ∉ D → rotate e p ∉ D.
410 | Proof.
411 | intros * His Hr; apply His; clear His.
412 | unfold D in Hr; simpl in Hr.
413 | unfold D; simpl.
414 | destruct Hr as (el & p₁ & Hso & Hn & Hr).
415 | exists el, p₁.
416 | split; [ | easy ].
417 | destruct Hso as (el₁ & Hso).
418 | exists (el₁ ++ [e]).
419 | now rewrite mat_of_path_app, mat_vec_mul_assoc, mat_of_path_single.
420 | Qed.
421 |
422 | Theorem r_decomposed_2 :
423 | ∀ f, orbit_selector f
424 | → ∀ os, os = mkos _ f
425 | → ∀ e,
426 | is_partition (ball ∖ D) [SS e; rot e (SS (negf e))].
427 | Proof.
428 | intros f (Hoe, Ho) os Hos e; subst os.
429 | split. {
430 | intros p.
431 | split. {
432 | intros Hnf.
433 | unfold set_union_list; simpl; unfold set_union.
434 | pose proof Ho p as H.
435 | apply same_orbit_sym in H.
436 | destruct H as (el, Hel).
437 | remember (norm_list el) as el₁ eqn:Hel₁; symmetry in Hel₁.
438 | destruct el₁ as [| e₁]. {
439 | rewrite rotate_rotate_norm, Hel₁ in Hel; simpl in Hel.
440 | rewrite mat_vec_mul_id in Hel.
441 | clear Hel₁.
442 | right; left.
443 | unfold rot.
444 | split. {
445 | split. {
446 | destruct Hnf as (His, _).
447 | now apply in_ball_after_rotate.
448 | }
449 | destruct Hnf as (Hps, Hnpd).
450 | now apply no_fixpoint_after_rotate.
451 | }
452 | exists (negf e :: []), [].
453 | split; [ easy | simpl ].
454 | assert (H : f p = f (rotate (negf e) p)). {
455 | apply Hoe.
456 | exists (negf e :: []).
457 | apply mat_of_path_single.
458 | }
459 | now rewrite <- H, Hel.
460 | }
461 | destruct (free_elem_dec e e₁) as [H₁| H₁]; [ subst e₁ | ]. {
462 | left; split; [ easy | ].
463 | exists el, el₁.
464 | rewrite rotate_vec_mul.
465 | now split.
466 | }
467 | right; left.
468 | unfold rot.
469 | split. {
470 | split. {
471 | destruct Hnf as (His, _).
472 | now apply in_ball_after_rotate.
473 | }
474 | destruct Hnf.
475 | now apply no_fixpoint_after_rotate.
476 | }
477 | assert (H : f p = f (rotate (negf e) p)). {
478 | apply Hoe.
479 | exists (negf e :: []).
480 | apply mat_of_path_single.
481 | }
482 | simpl; rewrite <- H.
483 | exists (negf e :: el), (e₁ :: el₁); simpl.
484 | rewrite rotate_vec_mul.
485 | rewrite Hel₁, Hel.
486 | destruct (letter_opp_dec (negf e) e₁) as [H₂| H₂]; [ | now split ].
487 | exfalso.
488 | apply letter_opp_negf in H₂.
489 | now apply H₁, negf_eq_eq.
490 | }
491 | intros Hul.
492 | destruct Hul as [(H, _)| [(H, _)| Hul]]; [ easy | | easy ].
493 | split. {
494 | destruct H as (His, _).
495 | apply in_ball_after_rotate with (e := e) in His.
496 | now rewrite rotate_rotate_neg in His.
497 | } {
498 | destruct H as (Hs, Hnp).
499 | apply no_fixpoint_after_rotate with (e := e) in Hnp.
500 | now rewrite rotate_rotate_neg in Hnp.
501 | }
502 | }
503 | intros i j Hij p.
504 | split; [ | easy ].
505 | intros (Hi, Hj).
506 | destruct i; [ simpl in Hi | ]. {
507 | destruct j; [ exfalso; now apply Hij | clear Hij ].
508 | destruct j; [ | now destruct j ].
509 | simpl in Hj.
510 | eapply not_start_with_rot in Hi; try eassumption; [ | easy ].
511 | easy.
512 | }
513 | destruct i; [ simpl in Hi | now destruct i ].
514 | destruct j; [ simpl in Hj; clear Hij | ].
515 | eapply not_start_with_rot in Hj; try eassumption; [ | easy ]. {
516 | easy.
517 | } {
518 | destruct j; [ now apply Hij | now destruct j ].
519 | }
520 | Qed.
521 |
522 | Add Parametric Morphism {A} : (@List.nth (set A))
523 | with signature eq ==> eq ==> set_eq ==> set_eq
524 | as nth_set_morph.
525 | Proof.
526 | intros i l a b Hab.
527 | revert i.
528 | induction l as [| y]; intros; [ destruct i; apply Hab | ].
529 | destruct i; simpl; [ easy | apply IHl ].
530 | Qed.
531 |
532 | Theorem rev_path_repeat : ∀ e n, rev_path (repeat e n) = repeat (negf e) n.
533 | Proof.
534 | intros e n.
535 | induction n; [ easy | simpl ].
536 | rewrite rev_path_cons, rev_path_single, IHn.
537 | apply app_repeat_diag.
538 | Qed.
539 |
540 | Theorem r_decomposed_2_a :
541 | ∀ f, orbit_selector f
542 | → ∀ os, os = mkos _ f
543 | → is_partition (ball ∖ D) [M ∪ SS ạ ∪ G; rot ạ (SS ạ⁻¹ ∖ G)].
544 | Proof.
545 | intros f (Hoe, Ho) os Hos.
546 | split. {
547 | intros p.
548 | assert (Hfr : f (rotate ạ⁻¹ p) = f p). {
549 | apply Hoe; exists (ạ :: []).
550 | rewrite mat_of_path_single.
551 | apply rotate_neg_rotate.
552 | }
553 | split. {
554 | intros Hnf.
555 | unfold set_union_list; simpl; unfold set_union.
556 | pose proof Ho p as H.
557 | apply same_orbit_sym in H.
558 | destruct H as (el, Hel).
559 | remember (norm_list el) as el₁ eqn:Hel₁; symmetry in Hel₁.
560 | destruct el₁ as [| e₁]. {
561 | rewrite rotate_rotate_norm, Hel₁ in Hel; simpl in Hel.
562 | rewrite mat_vec_mul_id in Hel.
563 | clear el Hel₁.
564 | left; left; left.
565 | split; [ easy | subst os; now symmetry ].
566 | }
567 | destruct e₁ as (t, d); destruct t. {
568 | destruct d. {
569 | destruct (EM (p ∈ G)) as [HB| HB]; [ left; now right | ].
570 | right; left; simpl.
571 | split. {
572 | split. {
573 | destruct Hnf as (His, Hnf).
574 | split; [ now apply in_ball_after_rotate | ].
575 | now apply no_fixpoint_after_rotate.
576 | }
577 | subst os; simpl.
578 | rewrite Hfr.
579 | exists (ạ⁻¹ :: el), (norm_list el).
580 | split; [ now simpl; rewrite Hel₁ | ].
581 | now simpl; rewrite rotate_vec_mul; f_equal.
582 | }
583 | simpl; intros (Haf & n & Hoo); apply HB; clear HB.
584 | split; [ easy | ].
585 | unfold orbit_by_seq_of in Hoo |-*; simpl.
586 | remember S as g; subst os; simpl in Hoo |-*; subst g.
587 | rewrite Hfr in Hoo; simpl in Hoo.
588 | apply f_equal with (f := rotate (FE la false)) in Hoo.
589 | rewrite <- rotate_vec_mul in Hoo; simpl in Hoo.
590 | do 2 rewrite rotate_rotate_neg in Hoo.
591 | rewrite rotate_vec_mul in Hoo.
592 | destruct n; [ | now exists n ].
593 | simpl in Hoo.
594 | rewrite mat_vec_mul_id in Hoo.
595 | rewrite Hoo in Hel.
596 | destruct Hnf as (His & Hoh).
597 | exfalso; apply Hoh.
598 | exists el, p.
599 | now rewrite Hel₁.
600 | }
601 | left; left; right.
602 | split; [ easy | ].
603 | exists el, el₁; subst os.
604 | rewrite rotate_vec_mul.
605 | now split.
606 | }
607 | right; left.
608 | split; simpl. {
609 | split. {
610 | destruct Hnf as (His & Hnf).
611 | split; [ now apply in_ball_after_rotate | ].
612 | now apply no_fixpoint_after_rotate.
613 | }
614 | subst os; simpl; rewrite Hfr.
615 | exists (ạ⁻¹ :: el), (norm_list el).
616 | split; [ now simpl; rewrite Hel₁ | ].
617 | now simpl; rewrite rotate_vec_mul; f_equal.
618 | }
619 | intros (Hnf₂, Hoo).
620 | subst os; simpl in Hoo.
621 | unfold orbit_by_seq_of in Hoo; simpl in Hoo.
622 | rewrite Hfr in Hoo.
623 | destruct Hoo as (n, Hr).
624 | apply f_equal with (f := rotate (FE la false)) in Hr.
625 | rewrite <- rotate_vec_mul in Hr; simpl in Hr.
626 | do 2 rewrite rotate_rotate_neg in Hr.
627 | rewrite rotate_vec_mul in Hr.
628 | destruct n. {
629 | simpl in Hr; rewrite mat_vec_mul_id in Hr.
630 | rewrite Hr in Hel.
631 | destruct Hnf as (His, Hoh).
632 | now apply Hoh; exists el, p; rewrite Hel₁.
633 | }
634 | apply rotate_rev_path in Hr.
635 | rewrite <- Hr in Hel.
636 | rewrite <- mat_vec_mul_assoc in Hel.
637 | rewrite <- mat_of_path_app in Hel.
638 | destruct Hnf as (His, Hoh).
639 | apply Hoh.
640 | exists (el ++ rev_path (repeat ạ⁻¹ (S n))), p.
641 | split; [ easy | ].
642 | split; [ | easy ].
643 | replace el with ([] ++ el) by easy.
644 | rewrite <- app_assoc, <- is_normal, Hel₁, app_nil_l.
645 | rewrite rev_path_repeat.
646 | remember norm_list as g; remember S as h; simpl; subst g h.
647 | rewrite app_of_cons, app_assoc.
648 | intros H.
649 | eapply norm_list_app_is_nil in H. {
650 | simpl in H.
651 | apply f_equal with (f := rev_path) in H.
652 | rewrite rev_path_involutive in H.
653 | rewrite <- app_repeat_diag in H.
654 | now rewrite rev_path_app in H; simpl in H.
655 | } {
656 | unfold app; rewrite <- Hel₁; symmetry.
657 | apply norm_list_idemp.
658 | }
659 | symmetry; apply norm_list_repeat.
660 | }
661 | intros HE.
662 | simpl in HE.
663 | destruct HE as [[[HE| HE]| HE]| [HE| HE]]; try now destruct HE.
664 | destruct HE as (((His & Hoo) & HE) & HB).
665 | split. {
666 | apply in_ball_after_rotate with (e := ạ) in His.
667 | now rewrite rotate_rotate_neg in His.
668 | } {
669 | apply no_fixpoint_after_rotate with (e := ạ) in Hoo.
670 | now rewrite rotate_rotate_neg in Hoo.
671 | }
672 | }
673 | intros i j Hij p.
674 | assert (Hfr : f (rotate ạ⁻¹ p) = f p). {
675 | apply Hoe; exists (ạ :: []).
676 | rewrite mat_of_path_single.
677 | apply rotate_neg_rotate.
678 | }
679 | split; [ | easy ].
680 | intros (Hi, Hj).
681 | destruct i; [ simpl in Hi | ]. {
682 | destruct j; [ exfalso; now apply Hij | clear Hij ].
683 | destruct j; [ simpl in Hj | now destruct j ].
684 | eapply decompose_2a_contrad_case; unfold set_union; try eassumption.
685 | easy.
686 | }
687 | destruct i; [ simpl in Hi | now destruct i ].
688 | destruct j. {
689 | eapply decompose_2a_contrad_case; unfold set_union; try eassumption.
690 | easy.
691 | }
692 | destruct j; [ now apply Hij | now destruct j ].
693 | Qed.
694 |
695 | Theorem r_decomposed_2_b :
696 | ∀ f, orbit_selector f
697 | → ∀ os, os = mkos _ f
698 | → is_partition (ball ∖ D) [SS ḅ; rot ḅ (SS ḅ⁻¹)].
699 | Proof.
700 | intros.
701 | eapply r_decomposed_2; eassumption.
702 | Qed.
703 |
704 | Theorem rot_set_map_mul : ∀ e E,
705 | (rot e E = set_map (mat_vec_mul (mat_of_elem e)) E)%S.
706 | Proof.
707 | intros; intros v.
708 | split; intros H. {
709 | exists (rotate (negf e) v).
710 | split; [ easy | unfold rotate ].
711 | rewrite <- mat_vec_mul_assoc.
712 | now rewrite mat_of_elem_mul_negf_r, mat_vec_mul_id.
713 | }
714 | destruct H as (u & Hu & Hv).
715 | rewrite <- Hv; simpl; unfold rotate.
716 | rewrite <- mat_vec_mul_assoc.
717 | now rewrite mat_of_elem_mul_negf_l, mat_vec_mul_id.
718 | Qed.
719 |
--------------------------------------------------------------------------------
/Partition.v:
--------------------------------------------------------------------------------
1 | (* Banach-Tarski paradox. *)
2 |
3 | From Stdlib Require Import Utf8 List Arith Compare_dec Setoid.
4 | Import ListNotations.
5 |
6 | Require Import Misc Pset.
7 |
8 | Definition is_partition {A} (E : set A) (Ep : list (set A)) :=
9 | (E = ⋃ Ep)%S ∧
10 | ∀ i j, i ≠ j → (Ep.[i] ∩ Ep.[j] = ∅)%S.
11 |
12 | Theorem is_partition_group_first_2_together :
13 | ∀ A (F : set A) P₁ P₂ Pl,
14 | is_partition F (P₁ :: P₂ :: Pl)
15 | → is_partition F (P₁ ∪ P₂ :: Pl).
16 | Proof.
17 | intros * Hp.
18 | destruct Hp as (Hu & Hi).
19 | split. {
20 | unfold set_union_list, set_union, set_eq in Hu |-*.
21 | simpl in Hu |-*.
22 | intros x.
23 | pose proof Hu x as H₁.
24 | destruct H₁ as (H₁ & H₂).
25 | split; intros H. {
26 | apply H₁ in H.
27 | destruct H as [H| H]; [ left; now left | ].
28 | destruct H as [H| H]; [ left; now right | ].
29 | now right.
30 | }
31 | apply H₂.
32 | destruct H as [[H| H]| H]; [ now left | right; now left | ].
33 | right; now right.
34 | }
35 | intros i j Hij.
36 | destruct i. {
37 | unfold set_inter, set_eq; simpl.
38 | intros x.
39 | split; [ | easy ].
40 | intros (H₁, H₂).
41 | destruct j; [ now apply Hij | clear Hij ].
42 | destruct H₁ as [H₁| H₁]. {
43 | eapply Hi with (i := O) (j := S (S j)); [ easy | ].
44 | unfold set_inter; simpl.
45 | split; eassumption.
46 | }
47 | eapply Hi with (i := 1%nat) (j := S (S j)); [ easy | ].
48 | unfold set_inter; simpl.
49 | split; eassumption.
50 | }
51 | unfold set_inter, set_union, set_eq; simpl.
52 | intros x.
53 | split; [ | easy ].
54 | intros (H₁ & H₂).
55 | destruct j. {
56 | destruct H₂ as [H₂| H₂]. {
57 | eapply Hi with (i := O) (j := S (S i)); [ easy | ].
58 | unfold set_inter; simpl.
59 | split; eassumption.
60 | }
61 | eapply Hi with (i := 1%nat) (j := S (S i)); [ easy | ].
62 | unfold set_inter; simpl.
63 | split; eassumption.
64 | }
65 | apply Hi with (i := S (S i)) (j := S (S j)) (x := x); [ | easy ].
66 | intros H; apply Hij.
67 | now apply Nat.succ_inj.
68 | Qed.
69 |
70 | Theorem is_partition_union_sub :
71 | ∀ A (F : set A) P₁ P₂ Pl (B : set A),
72 | is_partition F (P₁ :: P₂ :: Pl)
73 | → (B ⊂ P₂)%S
74 | → is_partition F (P₁ ∪ B :: P₂ ∖ B :: Pl)%S.
75 | Proof.
76 | intros A F P₁ P₂ Pl B Hp HB.
77 | destruct Hp as (Hu & Hi).
78 | split. {
79 | unfold set_union_list, set_union, set_sub, set_eq in Hu |-*.
80 | simpl in Hu |-*.
81 | intros x.
82 | split; intros H. {
83 | pose proof Hu x as H₁.
84 | destruct H₁ as (H₁ & H₂).
85 | pose proof H₁ H as H₃.
86 | destruct H₃ as [H₃| H₃]; [ left; now left | ].
87 | destruct H₃ as [H₃| H₃]; [ | right; now right ].
88 | destruct (EM (x ∈ B)) as [H₄| H₄]; [ left; now right | ].
89 | right; left; now split.
90 | }
91 | apply Hu.
92 | destruct H as [[H₁| H₁]| [H₁| H₁]]; [ now left | | | ]. {
93 | right; left; now apply HB.
94 | } {
95 | right; left; now destruct H₁.
96 | } {
97 | right; now right.
98 | }
99 | }
100 | intros i j Hij.
101 | destruct i. {
102 | unfold set_inter, set_union, set_sub, set_eq; simpl.
103 | intros x.
104 | split; [ | easy ].
105 | intros (H₁, H₂).
106 | destruct j; [ now apply Hij | clear Hij ].
107 | destruct H₁ as [H₁| H₁]. {
108 | eapply Hi with (i := O) (j := S j); [ easy | ].
109 | unfold set_inter; simpl.
110 | split; [ eassumption | ].
111 | destruct j; [ now destruct H₂ | easy ].
112 | }
113 | eapply Hi with (i := 1%nat) (j := S j). {
114 | destruct j; [ now destruct H₂ | easy ].
115 | }
116 | unfold set_inter; simpl.
117 | split; [ apply HB; eassumption | ].
118 | destruct j; [ now destruct H₂ | easy ].
119 | }
120 | unfold set_inter, set_union, set_sub, set_eq; simpl.
121 | intros x.
122 | split; [ | easy ].
123 | intros (H₁ & H₂).
124 | destruct j. {
125 | destruct H₂ as [H₂| H₂]. {
126 | eapply Hi with (i := O) (j := S i); [ easy | ].
127 | unfold set_inter; simpl.
128 | split; [ eassumption | ].
129 | destruct i; [ now destruct H₁ | easy ].
130 | }
131 | eapply Hi with (i := 1%nat) (j := S i). {
132 | destruct i; [ | easy ].
133 | now destruct H₁.
134 | }
135 | unfold set_inter; simpl.
136 | split; [ apply HB; eassumption | ].
137 | destruct i; [ now apply HB | easy ].
138 | }
139 | apply Hi with (i := S i) (j := S j) (x := x). {
140 | intros H; now apply Hij.
141 | }
142 | unfold set_inter; simpl.
143 | split. {
144 | destruct i; [ now destruct H₁ | easy ].
145 | } {
146 | destruct j; [ now destruct H₂ | easy ].
147 | }
148 | Qed.
149 |
150 | Theorem partition_union :
151 | ∀ A (F₁ F₂ : set A) P₁ P₂,
152 | (F₁ ∩ F₂ = ∅)%S
153 | → is_partition F₁ P₁
154 | → is_partition F₂ P₂
155 | → is_partition (F₁ ∪ F₂) (P₁ ++ P₂).
156 | Proof.
157 | intros * HFF HF₁ HF₂.
158 | destruct HF₁ as (HF₁ & HP₁).
159 | destruct HF₂ as (HF₂ & HP₂).
160 | split. {
161 | rewrite set_union_list_app.
162 | transitivity (F₁ ∪ ⋃ P₂). {
163 | intros x.
164 | split; intros H. {
165 | destruct H as [H| H]; [ now left | right ].
166 | now apply HF₂.
167 | }
168 | destruct H as [H| H]; [ now left | right ].
169 | now apply HF₂.
170 | }
171 | split; intros H. {
172 | destruct H as [H| H]; [ left | now right ].
173 | now apply HF₁.
174 | } {
175 | destruct H as [H| H]; [ left | now right ].
176 | now apply HF₁.
177 | }
178 | }
179 | intros * Hij.
180 | unfold set_inter, set_eq; simpl.
181 | intros x.
182 | split; intros H; [ | easy ].
183 | destruct H as (H₁, H₂).
184 | rewrite nth_set_app in H₁, H₂.
185 | destruct (lt_dec i (length P₁)) as [H₃| H₃]. {
186 | destruct (lt_dec j (length P₁)) as [H₄| H₄]. {
187 | eapply HP₁; [ eassumption | split; eassumption ].
188 | }
189 | eapply HFF.
190 | split. {
191 | apply HF₁.
192 | eapply nth_set_union_list; eassumption.
193 | }
194 | destruct (lt_dec (j - length P₁) (length P₂)) as [H₅| H₅]. {
195 | apply HF₂.
196 | eapply nth_set_union_list; eassumption.
197 | } {
198 | apply Nat.nlt_ge in H₅.
199 | now rewrite nth_overflow in H₂.
200 | }
201 | }
202 | apply Nat.nlt_ge in H₃.
203 | destruct (lt_dec j (length P₁)) as [H₄| H₄]. {
204 | apply HFF with x.
205 | split. {
206 | apply HF₁.
207 | eapply nth_set_union_list; eassumption.
208 | }
209 | destruct (lt_dec (i - length P₁) (length P₂)) as [H₅| H₅]. {
210 | apply HF₂.
211 | eapply nth_set_union_list; eassumption.
212 | }
213 | apply Nat.nlt_ge in H₅.
214 | now rewrite nth_overflow in H₁.
215 | }
216 | apply Nat.nlt_ge in H₄.
217 | eapply HP₂; [ | split; [ apply H₁ | apply H₂] ].
218 | intros H.
219 | apply Nat.add_cancel_l with (p := length P₁) in H.
220 | rewrite Nat.add_sub_assoc in H; [ | easy ].
221 | rewrite Nat.add_sub_assoc in H; [ | easy ].
222 | rewrite Nat.add_comm, Nat.add_sub in H.
223 | rewrite Nat.add_comm, Nat.add_sub in H.
224 | easy.
225 | Qed.
226 |
227 | Theorem is_partition_single : ∀ A (E : set A), is_partition E [E].
228 | Proof.
229 | intros.
230 | split; [ symmetry; now eapply set_union_empty_r | ].
231 | intros * Hij.
232 | destruct i. {
233 | destruct j; [ exfalso; now apply Hij | ].
234 | destruct j. {
235 | split; [ now intros (_, H) | easy ].
236 | }
237 | split; [ now intros (_, H) | easy ].
238 | }
239 | split; [ intros (H, _) | easy ].
240 | now destruct i.
241 | Qed.
242 |
243 | Add Parametric Morphism {A} : (@is_partition A)
244 | with signature set_eq ==> eq ==> iff
245 | as is_partition_morph.
246 | Proof.
247 | intros E F HEF SL.
248 | unfold is_partition.
249 | rewrite <- HEF.
250 | now split.
251 | Qed.
252 |
253 | Theorem is_partition_sub : ∀ A (E F : set A),
254 | F ⊂ E
255 | → is_partition E [F; E ∖ F].
256 | Proof.
257 | intros * HF.
258 | split. {
259 | simpl; rewrite set_union_empty_r.
260 | intros v; split; intros H. {
261 | now destruct (EM (v ∈ F)) as [Hi| Hni]; [ left | right ].
262 | }
263 | now destruct H as [H| H]; [ apply HF | destruct H ].
264 | }
265 | intros i j Hij.
266 | destruct i. {
267 | destruct j; [ easy | ].
268 | destruct j. {
269 | intros v.
270 | now split; intros Hv; [ simpl in Hv | ].
271 | }
272 | simpl; rewrite match_id.
273 | apply set_inter_empty_r.
274 | }
275 | destruct j. {
276 | destruct i. {
277 | intros v.
278 | now split; intros Hv; [ simpl in Hv | ].
279 | }
280 | simpl; rewrite match_id.
281 | apply set_inter_empty_l.
282 | }
283 | destruct i. {
284 | destruct j; [ easy | ].
285 | simpl; rewrite match_id.
286 | apply set_inter_empty_r.
287 | }
288 | destruct j. {
289 | simpl; rewrite match_id.
290 | apply set_inter_empty_l.
291 | }
292 | simpl; do 2 rewrite match_id.
293 | apply set_inter_empty_l.
294 | Qed.
295 |
--------------------------------------------------------------------------------
/Pset.v:
--------------------------------------------------------------------------------
1 | (* Sets as A → Prop *)
2 |
3 | From Stdlib Require Import Utf8 List Relations Arith Compare_dec Setoid.
4 | Require Import Misc.
5 |
6 | Record set A := mkset { setp : A → Prop }.
7 | Arguments mkset [A] _.
8 | Arguments setp [A] _ _.
9 |
10 | Definition empty_set {A} := mkset (λ _ : A, False).
11 |
12 | Notation "x ∈ S" := (setp S x) (at level 60).
13 | Notation "x ∉ S" := (¬ setp S x) (at level 60).
14 | Notation "∅" := (empty_set).
15 |
16 | Definition set_inter {A} (S₁ S₂ : set A) :=
17 | mkset (λ x, x ∈ S₁ ∧ x ∈ S₂).
18 | Definition set_union {A} (S₁ S₂ : set A) :=
19 | mkset (λ x, x ∈ S₁ ∨ x ∈ S₂).
20 | Definition set_union_list {A} (Si : list (set A)) :=
21 | fold_right set_union ∅ Si.
22 | Definition set_sub {A} (S₁ S₂ : set A) :=
23 | mkset (λ x, x ∈ S₁ ∧ x ∉ S₂).
24 | Definition set_incl {A} (S₁ S₂ : set A) :=
25 | ∀ x, x ∈ S₁ → x ∈ S₂.
26 |
27 | Arguments set_inter : simpl never.
28 | Arguments set_union : simpl never.
29 | Arguments set_sub : simpl never.
30 | Arguments set_incl : simpl never.
31 |
32 | Declare Scope set_scope.
33 | Delimit Scope set_scope with S.
34 |
35 | Definition set_eq {A} (S₁ S₂ : set A) := ∀ x, x ∈ S₁ ↔ x ∈ S₂.
36 |
37 | Notation "E₁ = E₂" := (set_eq E₁ E₂) : set_scope.
38 | Notation "E₁ ≠ E₂" := (¬ set_eq E₁ E₂) : set_scope.
39 | Notation "E₁ '∩' E₂" := (set_inter E₁ E₂) (at level 40).
40 | Notation "E₁ '∪' E₂" := (set_union E₁ E₂) (at level 50).
41 | Notation "E₁ '∖' E₂" := (set_sub E₁ E₂) (at level 50).
42 | Notation "E₁ '⊂' E₂" := (set_incl E₁ E₂) (at level 60).
43 | Notation "'⋃' Es" := (set_union_list Es) (at level 55).
44 | Notation "E .[ i ]" := (List.nth i E ∅)
45 | (at level 1, format "'[' E '[' .[ i ] ']' ']'").
46 |
47 | Definition set_map {A B} (f : A → B) s := mkset (λ v, ∃ u, u ∈ s ∧ f u = v).
48 |
49 | Add Parametric Morphism {A B f} : (@set_map A B f)
50 | with signature set_eq ==> set_eq
51 | as set_map_morph.
52 | Proof.
53 | intros E F HEF b.
54 | split; intros H. {
55 | destruct H as (a & Ha & Hf).
56 | now exists a; split; [ apply HEF in Ha | ].
57 | }
58 | destruct H as (a & Ha & Hf).
59 | now exists a; split; [ apply HEF in Ha | ].
60 | Qed.
61 |
62 | Theorem set_map_inter_distr : ∀ A B E F (f : A → B),
63 | FinFun.Injective f
64 | → (set_map f (E ∩ F) = set_map f E ∩ set_map f F)%S.
65 | Proof.
66 | intros * Hinj b.
67 | split; intros H. {
68 | destruct H as (a & (HaE & HaF) & Hf); simpl.
69 | split; exists a; easy.
70 | }
71 | simpl in H; simpl.
72 | destruct H as ((ae & HE & Hae) & (af & HF & Haf)).
73 | rewrite <- Haf in Hae.
74 | specialize (Hinj _ _ Hae); subst af.
75 | now exists ae.
76 | Qed.
77 |
78 | Theorem set_map_union_distr : ∀ A B E F (f : A → B),
79 | (set_map f (E ∪ F) = set_map f E ∪ set_map f F)%S.
80 | Proof.
81 | intros; intros b.
82 | split; intros H. {
83 | now destruct H as (a & [Hae| Haf] & Hf); [ left | right ]; exists a.
84 | }
85 | destruct H as [(a & Hae & Hf)| (a & Haf & Hf)]. {
86 | now exists a; split; [ left | ].
87 | } {
88 | now exists a; split; [ right | ].
89 | }
90 | Qed.
91 |
92 | Theorem set_map_empty : ∀ A B (f : A → B), (set_map f ∅ = ∅)%S.
93 | Proof.
94 | intros; intros b.
95 | now split; intros H; destruct H.
96 | Qed.
97 |
98 | Theorem in_set_map : ∀ A B x E (f : A → B),
99 | x ∈ E → f x ∈ set_map f E.
100 | Proof. now intros * Hx; exists x. Qed.
101 |
102 | Theorem set_map_in : ∀ A B x E (f : A → B),
103 | FinFun.Injective f
104 | → f x ∈ set_map f E
105 | → x ∈ E.
106 | Proof.
107 | intros * Hinj Hx.
108 | destruct Hx as (a & Ha & Hfa).
109 | now apply Hinj in Hfa; subst a.
110 | Qed.
111 |
112 | Theorem set_eq_refl A : reflexive (set A) set_eq.
113 | Proof. now intros P x; split. Qed.
114 |
115 | Theorem set_eq_sym A : symmetric (set A) set_eq.
116 | Proof.
117 | intros P₁ P₂ HPP x.
118 | destruct (HPP x) as (H₁, H₂).
119 | split; intros H; [ apply H₂, H | apply H₁, H ].
120 | Qed.
121 |
122 | Theorem set_eq_trans A : transitive (set A) set_eq.
123 | Proof.
124 | intros P₁ P₂ P₃ H₁₂ H₂₃ x.
125 | destruct (H₁₂ x) as (H₁, H₂).
126 | destruct (H₂₃ x) as (H₃, H₄).
127 | split; intros H; [ apply H₃, H₁, H | apply H₂, H₄, H ].
128 | Qed.
129 |
130 | Add Parametric Relation A : (set A) (@set_eq A)
131 | reflexivity proved by (set_eq_refl A)
132 | symmetry proved by (set_eq_sym A)
133 | transitivity proved by (set_eq_trans A)
134 | as set_eq_rel.
135 |
136 | Theorem eq_set_eq : ∀ A (x y : set A), x = y → (x = y)%S.
137 | Proof. intros; now subst x. Qed.
138 |
139 | Theorem included_trans A : transitive _ (@set_incl A).
140 | Proof.
141 | intros E F G HEF HFG x Hx.
142 | apply HFG, HEF, Hx.
143 | Qed.
144 |
145 | Add Parametric Morphism {A} : (@set_inter A)
146 | with signature set_eq ==> set_eq ==> set_eq
147 | as set_inter_morph.
148 | Proof.
149 | intros E E' HE F F' HF.
150 | unfold set_inter; intros p.
151 | split; intros (H₁, H₂). {
152 | now split; [ apply HE | apply HF ].
153 | } {
154 | now split; [ apply HE | apply HF ].
155 | }
156 | Qed.
157 |
158 | Add Parametric Morphism {A} : (@set_union A)
159 | with signature set_eq ==> set_eq ==> set_eq
160 | as set_union_morph.
161 | Proof.
162 | intros E E' HE F F' HF.
163 | intros p.
164 | split. {
165 | intros [H₁| H₂]; [ left; apply HE, H₁ | right; apply HF, H₂ ].
166 | } {
167 | intros [H₁| H₂]; [ left; apply HE, H₁ | right; apply HF, H₂ ].
168 | }
169 | Qed.
170 |
171 | Add Parametric Morphism {A} : (@set_sub A)
172 | with signature set_eq ==> set_eq ==> set_eq
173 | as set_sub_morph.
174 | Proof.
175 | intros E E' HE F F' HF.
176 | unfold set_sub; intros p.
177 | split; intros (H₁, H₂). {
178 | split; [ now apply HE | intros H; now apply H₂, HF ].
179 | } {
180 | split; [ now apply HE | intros H; now apply H₂, HF ].
181 | }
182 | Qed.
183 |
184 | Add Parametric Morphism {A} : (@set_incl A)
185 | with signature set_eq ==> set_eq ==> iff
186 | as included_morph.
187 | Proof.
188 | intros E F HEF E' F' HE'F'.
189 | split; intros HEE' x HF; apply HE'F', HEE', HEF, HF.
190 | Qed.
191 |
192 | Theorem set_eq_equiv {A} : ∀ (E F : set A),
193 | (E = F)%S
194 | → ∀ p, p ∈ E ↔ p ∈ F.
195 | Proof. intros s * HEF; apply HEF. Qed.
196 |
197 | Theorem set_union_empty_r : ∀ A (F : set A),
198 | (F ∪ ∅ = F)%S.
199 | Proof.
200 | intros * x.
201 | split; intros H; [ | now left ].
202 | now destruct H.
203 | Qed.
204 |
205 | Theorem set_inter_empty_l : ∀ A (F : set A),
206 | (∅ ∩ F = ∅)%S.
207 | Proof.
208 | intros * x.
209 | split; intros H; [ now destruct H as (H, _) | easy ].
210 | Qed.
211 |
212 | Theorem set_inter_empty_r : ∀ A (F : set A),
213 | (F ∩ ∅ = ∅)%S.
214 | Proof.
215 | intros * x.
216 | split; intros H; [ now destruct H as (_, H) | easy ].
217 | Qed.
218 |
219 | Theorem set_inter_comm : ∀ A (E F : set A),
220 | (E ∩ F = F ∩ E)%S.
221 | Proof.
222 | intros; intros x.
223 | split; intros H; destruct H as (HE & HF); now split.
224 | Qed.
225 |
226 | Theorem set_union_comm : ∀ A (E F : set A),
227 | (E ∪ F = F ∪ E)%S.
228 | Proof.
229 | intros; intros x.
230 | now split; intros [HE| HF]; [ right | left | right | left ] .
231 | Qed.
232 |
233 | Theorem set_inter_assoc : ∀ A (E F G : set A),
234 | (E ∩ (F ∩ G) = (E ∩ F) ∩ G)%S.
235 | Proof.
236 | intros; intros x.
237 | split; intros H. {
238 | destruct H as (HE & (HF & HG)).
239 | split; [ now split | easy ].
240 | } {
241 | destruct H as ((HE & HF) & HG).
242 | split; [ easy | now split ].
243 | }
244 | Qed.
245 |
246 | Theorem set_union_assoc : ∀ A (E F G : set A),
247 | (E ∪ (F ∪ G) = (E ∪ F) ∪ G)%S.
248 | Proof.
249 | intros; intros x.
250 | split; intros H. {
251 | destruct H as [H| [H| H]]. {
252 | left; now left.
253 | } {
254 | left; now right.
255 | } {
256 | now right.
257 | }
258 | }
259 | destruct H as [[H| H]| H]; [ now left | | ]. {
260 | right; now left.
261 | } {
262 | right; now right.
263 | }
264 | Qed.
265 |
266 | Theorem set_inter_shuffle0 : ∀ A (E F G : set A),
267 | (E ∩ F ∩ G = E ∩ G ∩ F)%S.
268 | Proof.
269 | intros; intros x.
270 | split; intros H. {
271 | destruct H as ((HE & HF) & HG).
272 | split; [ now split | easy ].
273 | } {
274 | destruct H as ((HE & HF) & HG).
275 | split; [ now split | easy ].
276 | }
277 | Qed.
278 |
279 | Theorem set_union_list_app : ∀ A (P₁ P₂ : list (set A)),
280 | (⋃ (P₁ ++ P₂) = (⋃ P₁) ∪ (⋃ P₂))%S.
281 | Proof.
282 | intros.
283 | revert P₁.
284 | induction P₂ as [| Q]; intros. {
285 | rewrite app_nil_r; simpl.
286 | now rewrite set_union_empty_r.
287 | }
288 | rewrite cons_comm_app, app_assoc; simpl.
289 | rewrite IHP₂.
290 | unfold set_union_list; simpl; rewrite set_union_assoc.
291 | intros x.
292 | split; intros H. {
293 | destruct H as [H| H]; [ left | now right ].
294 | unfold set_union_list in H.
295 | rewrite fold_right_app in H.
296 | simpl in H.
297 | clear - H.
298 | induction P₁ as [| R P₁]. {
299 | simpl in H; simpl.
300 | destruct H as [H| H]; [ now right | easy ].
301 | }
302 | simpl in H.
303 | destruct H as [H| H]; [ simpl; left; now left | ].
304 | apply IHP₁ in H.
305 | destruct H as [H| H]; [ simpl; left; now right | ].
306 | now right.
307 | }
308 | destruct H as [H| H]; [ left | now right ].
309 | unfold set_union_list.
310 | rewrite fold_right_app; simpl.
311 | clear - H.
312 | induction P₁ as [| R P₁]. {
313 | simpl in H; simpl; left.
314 | now destruct H.
315 | }
316 | simpl in H; simpl.
317 | destruct H. {
318 | destruct H; [ now left | right ].
319 | now apply IHP₁; left.
320 | }
321 | now right; apply IHP₁; right.
322 | Qed.
323 |
324 | Theorem nth_set_union_list : ∀ A (P : list (set A)) i x,
325 | i < length P → x ∈ P.[i] → x ∈ ⋃ P.
326 | Proof.
327 | intros A P i x Hi H.
328 | revert P H Hi.
329 | induction i; intros P H Hi. {
330 | destruct P as [| E P]; [ easy | now left ].
331 | }
332 | destruct P as [| E P]; [ easy | simpl in Hi ].
333 | apply Nat.succ_lt_mono in Hi.
334 | right; now apply IHi.
335 | Qed.
336 |
337 | Theorem nth_set_app : ∀ A (P₁ P₂ : list (set A)) i,
338 | (P₁ ++ P₂).[i] =
339 | if lt_dec i (length P₁) then P₁.[i] else P₂.[i-length P₁].
340 | Proof.
341 | intros.
342 | unfold set_union, set_eq; simpl; intros.
343 | destruct (lt_dec i (length P₁)) as [H₁| H₁]. {
344 | now rewrite app_nth1.
345 | }
346 | rewrite app_nth2; [ easy | now apply Nat.nlt_ge ].
347 | Qed.
348 |
349 | Theorem set_union_inter_self : ∀ A (E : set A) EL,
350 | E ⊂ ⋃ EL
351 | → (E = ⋃ map (set_inter E) EL)%S.
352 | Proof.
353 | intros * HEL x.
354 | split; intros Hx. {
355 | generalize Hx; intros Hxl.
356 | apply HEL in Hxl.
357 | clear -Hx Hxl.
358 | induction EL as [| E₁ EL]; intros; [ easy | ].
359 | destruct Hxl as [Hxl| Hxl]; [ left; now split | ].
360 | right; now apply IHEL.
361 | }
362 | clear -Hx.
363 | induction EL as [| E₁ EL]; intros; [ easy | ].
364 | destruct Hx as [(Hx, _)| Hx]; [ easy | ].
365 | apply IHEL, Hx.
366 | Qed.
367 |
368 | Theorem set_inter_union_distr_r : ∀ A (E F G : set A),
369 | ((E ∪ F) ∩ G = (E ∩ G) ∪ (F ∩ G))%S.
370 | Proof.
371 | intros * x.
372 | split; intros H. {
373 | now destruct H as ([HE| HF] & HG); [ left | right ].
374 | }
375 | destruct H as [(HE, HG)| (HF, HG)]. {
376 | now split; [ left | ].
377 | } {
378 | now split; [ right | ].
379 | }
380 | Qed.
381 |
382 | Add Parametric Morphism {A} : (@setp A)
383 | with signature set_eq ==> eq ==> iff
384 | as setp_morph.
385 | Proof.
386 | intros E F HEF x.
387 | apply HEF.
388 | Qed.
389 |
390 | Theorem set_map_union_list_distr : ∀ A B EL (f : A → B),
391 | (set_map f (⋃ EL) = ⋃ (map (set_map f) EL))%S.
392 | Proof.
393 | intros.
394 | induction EL as [| E EL]; [ apply set_map_empty | simpl ].
395 | now rewrite set_map_union_distr, IHEL.
396 | Qed.
397 |
398 | Theorem set_sub_empty_l : ∀ A (E : set A), (∅ ∖ E = ∅)%S.
399 | Proof.
400 | intros; intros a; now simpl; split; intros H.
401 | Qed.
402 |
403 | Theorem set_sub_sub_swap : ∀ A (E F G : set A),
404 | (E ∖ F ∖ G = E ∖ G ∖ F)%S.
405 | Proof.
406 | intros; intros x; split; intros Hx. {
407 | now destruct Hx as ((HE & HF) & HG).
408 | } {
409 | now destruct Hx as ((HE & HF) & HG).
410 | }
411 | Qed.
412 |
--------------------------------------------------------------------------------
/README.md:
--------------------------------------------------------------------------------
1 | # OVERVIEW
2 | Banach-Tarski paradox.
3 |
4 | # INSPIRATIONS
5 | - Stan Wagon: The Banach-Tarski Paradox, Cambridge University Press
6 | - Wikipedia: Banach–Tarski paradox
7 | - http://people.math.umass.edu/~weston/oldpapers/banach.pdf
8 |
9 | # AUTHOR
10 | Daniel de Rauglaudre
11 |
12 | # ROCQ VERSION
13 | $ rocq -v
14 | The Rocq Prover, version 9.0.0
15 | compiled with OCaml 4.14.1
16 |
--------------------------------------------------------------------------------
/Reverse.v:
--------------------------------------------------------------------------------
1 | (* Banach-Tarski paradox. *)
2 |
3 | From Stdlib Require Import Utf8 List Relations Arith Wf_nat Compare_dec.
4 | Import ListNotations.
5 | Require Import Misc Words Normalize.
6 |
7 | Definition rev_path el := map negf (rev el).
8 |
9 | Theorem rev_path_cons : ∀ e el,
10 | rev_path (e :: el) = rev_path el ++ rev_path [e].
11 | Proof.
12 | intros e el.
13 | unfold rev_path; simpl.
14 | now rewrite map_app.
15 | Qed.
16 |
17 | Theorem rev_path_app : ∀ el₁ el₂,
18 | rev_path (el₁ ++ el₂) = rev_path el₂ ++ rev_path el₁.
19 | Proof.
20 | intros el₁ el₂.
21 | revert el₁.
22 | induction el₂ as [| (t, d)]; intros; [ now rewrite app_nil_r | ].
23 | rewrite rev_path_cons, cons_comm_app, app_assoc, IHel₂.
24 | rewrite <- app_assoc; f_equal; simpl.
25 | clear el₂ IHel₂.
26 | induction el₁ as [| e₁]; [ easy | ].
27 | simpl; rewrite rev_path_cons; rewrite IHel₁.
28 | simpl; f_equal; symmetry.
29 | now rewrite rev_path_cons.
30 | Qed.
31 |
32 | Theorem rev_path_involutive : ∀ el, rev_path (rev_path el) = el.
33 | Proof.
34 | intros el.
35 | induction el as [| (t, d)]; [ easy | simpl ].
36 | rewrite rev_path_cons, rev_path_app.
37 | rewrite IHel; simpl; rewrite Bool.negb_involutive.
38 | easy.
39 | Qed.
40 |
41 | Theorem rev_path_single : ∀ e, rev_path [e] = [negf e].
42 | Proof. easy. Qed.
43 |
44 | Theorem rev_path_nil : rev_path [] = [].
45 | Proof. easy. Qed.
46 |
47 | Theorem rev_path_is_nil : ∀ el, rev_path el = [] → el = [].
48 | Proof.
49 | intros el Hr.
50 | destruct el as [| e]; [ easy | ].
51 | rewrite rev_path_cons, rev_path_single in Hr.
52 | now destruct (rev_path el).
53 | Qed.
54 |
55 | Theorem rev_path_eq_eq : ∀ el₁ el₂,
56 | rev_path el₁ = rev_path el₂ ↔ el₁ = el₂.
57 | Proof.
58 | intros el₁ el₂.
59 | split; [ | intros H; now subst ].
60 | intros Hr.
61 | revert el₂ Hr.
62 | induction el₁ as [| e₁]; intros. {
63 | rewrite rev_path_nil in Hr.
64 | symmetry in Hr; apply rev_path_is_nil in Hr.
65 | now destruct Hr.
66 | }
67 | rewrite rev_path_cons, rev_path_single in Hr.
68 | destruct el₂ as [| e₂]. {
69 | rewrite rev_path_nil in Hr.
70 | now destruct (rev_path el₁).
71 | }
72 | rewrite rev_path_cons, rev_path_single in Hr.
73 | apply app_inj_tail in Hr.
74 | destruct Hr as (Hr, Hn).
75 | apply IHel₁ in Hr.
76 | apply negf_eq_eq in Hn.
77 | now subst el₁ e₁.
78 | Qed.
79 |
80 | Theorem norm_list_rev_path : ∀ el,
81 | norm_list el = el → norm_list (rev_path el) = rev_path el.
82 | Proof.
83 | intros el Hel.
84 | induction el as [| e] using rev_ind; [ easy | ].
85 | rewrite rev_path_app; simpl.
86 | generalize Hel; intros Hn.
87 | apply norm_list_app_diag in Hn.
88 | rewrite IHel; [ | easy ].
89 | remember (rev_path el) as el₁ eqn:Hel₁.
90 | symmetry in Hel₁.
91 | destruct el₁ as [| e₁]; [ easy | ].
92 | destruct (letter_opp_dec (negf e) e₁) as [H₁| H₁]; [ | easy ].
93 | exfalso.
94 | apply letter_opp_sym, letter_opp_negf in H₁.
95 | rewrite negf_involutive in H₁; subst e₁.
96 | rewrite <- rev_path_involutive in Hel₁.
97 | rewrite rev_path_cons, rev_path_single in Hel₁.
98 | simpl in Hel₁.
99 | apply -> rev_path_eq_eq in Hel₁.
100 | rewrite Hel₁ in Hel.
101 | rewrite <- app_assoc in Hel; simpl in Hel.
102 | revert Hel.
103 | apply norm_list_no_consec2.
104 | Qed.
105 |
106 | Theorem rev_path_norm_list : ∀ el,
107 | rev_path (norm_list el) = norm_list (rev_path el).
108 | Proof.
109 | intros el.
110 | remember (length el) as len eqn:Hlen.
111 | symmetry in Hlen.
112 | revert el Hlen.
113 | induction len as (len, IHlen) using lt_wf_rec; intros.
114 | destruct len. {
115 | apply length_zero_iff_nil in Hlen; now subst el.
116 | }
117 | destruct (norm_list_dec el) as [H₁| H₁]. {
118 | generalize H₁; intros H₂.
119 | apply norm_list_rev_path in H₂.
120 | rewrite H₁, H₂.
121 | easy.
122 | }
123 | destruct H₁ as (el₁, (t, (d, (el₂, Hs)))).
124 | generalize Hs; intros H.
125 | rewrite H, norm_list_cancel_in.
126 | rewrite rev_path_app, rev_path_cons, rev_path_cons.
127 | do 2 rewrite rev_path_single.
128 | do 2 rewrite <- app_assoc; simpl.
129 | rewrite Bool.negb_involutive.
130 | rewrite norm_list_cancel_in.
131 | rewrite <- rev_path_app.
132 | apply IHlen with (m := length (el₁ ++ el₂)); [ | easy ].
133 | rewrite <- Hlen, H; simpl.
134 | do 2 rewrite length_app; simpl.
135 | apply Nat.add_lt_mono_l.
136 | etransitivity; eapply Nat.lt_succ_diag_r.
137 | Qed.
138 |
139 | Theorem norm_list_app_is_nil : ∀ el₁ el₂,
140 | el₁ = norm_list el₁
141 | → el₂ = norm_list el₂
142 | → norm_list (el₁ ++ el₂) = []
143 | → el₂ = rev_path el₁.
144 | Proof.
145 | intros el₁ el₂ Hel₁ Hel₂ Hn.
146 | symmetry in Hel₁, Hel₂.
147 | remember (length (el₁ ++ el₂)) as len eqn:Hlen.
148 | symmetry in Hlen.
149 | rewrite <- Hel₁, <- Hel₂, rev_path_norm_list.
150 | revert el₁ el₂ Hel₁ Hel₂ Hn Hlen.
151 | induction len as (len, IHlen) using lt_wf_rec; intros.
152 | destruct (norm_list_dec (el₁ ++ el₂)) as [H₁| H₁]. {
153 | rewrite H₁ in Hn.
154 | apply app_eq_nil in Hn.
155 | destruct Hn; now subst el₁ el₂.
156 | }
157 | destruct H₁ as (el₃ & t & d & el₄ & H₁).
158 | rewrite H₁, length_app, Nat.add_comm in Hlen.
159 | destruct len; [ easy | ].
160 | destruct len; [ easy | simpl in Hlen ].
161 | do 2 apply -> Nat.succ_inj_wd in Hlen.
162 | rewrite Nat.add_comm, <- length_app in Hlen.
163 | assert (H₂ : len < S (S len)). {
164 | transitivity (S len); apply Nat.lt_succ_diag_r.
165 | }
166 | rewrite <- Hel₁, <- Hel₂ in H₁.
167 | apply norm_list_app_split in H₁.
168 | destruct H₁ as (H₃, H₄).
169 | rewrite Hel₁ in H₃; rewrite H₃ in Hel₁.
170 | apply norm_list_app_diag in Hel₁.
171 | rewrite Hel₂ in H₄; rewrite H₄ in Hel₂.
172 | apply norm_list_cons in Hel₂.
173 | rewrite H₃, H₄ in Hn.
174 | rewrite <- app_assoc, <- cons_comm_app in Hn.
175 | rewrite norm_list_cancel_in in Hn.
176 | pose proof IHlen len H₂ el₃ el₄ Hel₁ Hel₂ Hn Hlen as H.
177 | rewrite Hel₂, <- rev_path_norm_list, Hel₁ in H.
178 | rewrite H₃, H₄, H, rev_path_app.
179 | easy.
180 | Qed.
181 |
182 | Theorem rev_path_length : ∀ el, length (rev_path el) = length el.
183 | Proof.
184 | intros.
185 | induction el as [| e el]; [ easy | simpl ].
186 | rewrite rev_path_cons, rev_path_single.
187 | rewrite length_app; simpl.
188 | now rewrite Nat.add_1_r, IHel.
189 | Qed.
190 |
191 | Theorem rev_path_nth : ∀ el i,
192 | (i < length el)%nat
193 | → List.nth i (rev_path el) ạ = negf (List.nth (length el - S i) el ạ⁻¹).
194 | Proof.
195 | intros el i Hlen.
196 | revert i Hlen.
197 | induction el as [| e el]; intros; [ now simpl; rewrite match_id | ].
198 | rewrite rev_path_cons, rev_path_single.
199 | destruct (lt_dec i (length el)) as [Hi| Hi]. {
200 | rewrite app_nth1; [ | now rewrite rev_path_length ].
201 | rewrite IHel; simpl; [ f_equal | easy ].
202 | remember (length el - i)%nat as n eqn:Hn.
203 | symmetry in Hn.
204 | destruct n. {
205 | apply Nat.sub_0_le in Hn.
206 | apply Nat.lt_succ_r in Hn.
207 | now apply Nat.nle_gt in Hn.
208 | }
209 | f_equal; apply Nat.succ_inj.
210 | now rewrite <- Hn, <- Nat.sub_succ_l.
211 | }
212 | apply Nat.nlt_ge in Hi.
213 | simpl in Hlen; unfold lt in Hlen.
214 | apply Nat.succ_le_mono in Hlen.
215 | apply Nat.le_antisymm in Hi; [ | easy ].
216 | rewrite Hi.
217 | rewrite app_nth2; [ simpl | now rewrite rev_path_length; unfold ge ].
218 | now rewrite rev_path_length, <- Hi, Nat.sub_diag.
219 | Qed.
220 |
221 | Theorem nth_in_consec_split : ∀ A (n : nat) (l : list A) (d₁ d₂ : A),
222 | (S n < length l)%nat
223 | → ∃ l1 l2 : list A,
224 | l = l1 ++ List.nth n l d₁ :: List.nth (S n) l d₂ :: l2.
225 | Proof.
226 | intros * Hn.
227 | revert n Hn.
228 | induction l as [| x l]; intros; [ now apply Nat.nlt_0_r in Hn | ].
229 | simpl in Hn.
230 | apply Nat.succ_lt_mono in Hn.
231 | destruct n. {
232 | destruct l as [| y l]; [ now apply Nat.nlt_0_r in Hn | ].
233 | now exists [], l.
234 | }
235 | apply IHl in Hn.
236 | destruct Hn as (l1 & l2 & Hn).
237 | now exists (x :: l1), l2; simpl; f_equal.
238 | Qed.
239 |
240 | Theorem rev_norm_path_eq_path : ∀ el,
241 | norm_list el = el
242 | → rev_path el = el
243 | → el = [].
244 | Proof.
245 | intros * Hn Hr.
246 | destruct el as [| e₁ el]; [ easy | exfalso ].
247 | destruct (zerop (length el mod 2)) as [Hel| Hel]. {
248 | apply Nat.Div0.mod_divides in Hel.
249 | destruct Hel as (c, Hc).
250 | assert (Hlt : (c < length (e₁ :: el))%nat). {
251 | simpl; rewrite Hc; simpl.
252 | rewrite Nat.add_0_r.
253 | apply Nat.lt_succ_r, Nat.le_add_r.
254 | }
255 | pose proof rev_path_nth (e₁ :: el) c Hlt as H.
256 | rewrite Hr in H; simpl in H.
257 | symmetry in H.
258 | replace (length el - c)%nat with c in H. {
259 | destruct c; [ now apply no_fixpoint_negf in H | ].
260 | simpl in Hlt.
261 | apply Nat.succ_lt_mono in Hlt.
262 | erewrite nth_indep in H; [ now apply no_fixpoint_negf in H | ].
263 | rewrite Hc; simpl; rewrite Nat.add_0_r.
264 | apply Nat.lt_succ_r, Nat.le_add_r.
265 | }
266 | rewrite Hc; simpl.
267 | now rewrite Nat.add_0_r, Nat.add_sub.
268 | }
269 | assert (He : (length (e₁ :: el) mod 2 = 0)%nat). {
270 | simpl.
271 | rewrite <- Nat.add_1_r.
272 | rewrite <- Nat.Div0.add_mod_idemp_l.
273 | remember (length el mod 2) as m eqn:Hm.
274 | destruct m; [ easy | ].
275 | destruct m; [ easy | ].
276 | assert (H : (2 ≠ 0)%nat) by easy.
277 | apply (Nat.mod_upper_bound (length el)) in H.
278 | rewrite <- Hm in H.
279 | do 2 apply Nat.succ_lt_mono in H.
280 | now apply Nat.nlt_0_r in H.
281 | }
282 | apply Nat.Div0.mod_divides in He.
283 | destruct He as (c, Hc).
284 | destruct c; [ easy | ].
285 | assert (Hlt : (S c < length (e₁ :: el))%nat). {
286 | rewrite Hc; simpl; rewrite Nat.add_0_r.
287 | apply Nat.lt_succ_r; rewrite Nat.add_comm.
288 | apply Nat.le_add_r.
289 | }
290 | generalize Hlt; intros H.
291 | apply rev_path_nth in H.
292 | rewrite Hr in H.
293 | remember (e₁ :: el) as el₂; simpl in H.
294 | rewrite Hc in H; simpl in H.
295 | rewrite Nat.add_0_r, Nat.add_sub in H; subst el₂.
296 | rename H into Hlen.
297 | pose proof nth_in_consec_split free_elem c (e₁ :: el) ạ⁻¹ ạ Hlt.
298 | destruct H as (l₁ & l₂ & Hll).
299 | rewrite Hlen, <- Hn in Hll.
300 | now apply norm_list_no_consec in Hll.
301 | Qed.
302 |
303 | Theorem rev_path_eq_path : ∀ el,
304 | rev_path (norm_list el) = norm_list el
305 | → norm_list el = [].
306 | Proof.
307 | intros el Hel.
308 | remember (norm_list el) as el₁ eqn:Hel₁.
309 | assert (H : norm_list el₁ = el₁) by (subst el₁; apply norm_list_idemp).
310 | clear el Hel₁.
311 | rename el₁ into el; rename H into Hn.
312 | now apply rev_norm_path_eq_path.
313 | Qed.
314 |
315 | Theorem norm_list_app_diag_is_nil : ∀ el,
316 | norm_list (el ++ el) = []
317 | → norm_list el = [].
318 | Proof.
319 | intros el Hel.
320 | rewrite norm_list_normal_l in Hel.
321 | rewrite norm_list_normal_r in Hel.
322 | apply norm_list_app_is_nil in Hel; try now rewrite norm_list_idemp.
323 | now apply rev_path_eq_path.
324 | Qed.
325 |
--------------------------------------------------------------------------------
/RnCountable.v:
--------------------------------------------------------------------------------
1 | (* Banach-Tarski paradox. *)
2 |
3 | From Stdlib Require Import Utf8 Arith.
4 | From Stdlib Require Import Reals Psatz.
5 |
6 | Require Import MiscReals Countable.
7 |
8 | Notation "'ℝ'" := R.
9 | Notation "x '≤' y" := (Rle x y) : R_scope.
10 |
11 | Definition ter_bin_of_frac_part x n :=
12 | if Rlt_dec (frac_part (x * 3 ^ n)) (1 / 3) then false else true.
13 |
14 | Fixpoint partial_sum3_aux k (u : nat → bool) pow i :=
15 | match k with
16 | | 0 => 0%R
17 | | S k' =>
18 | if u i then (pow / 3 + partial_sum3_aux k' u (pow / 3) (S i))%R
19 | else partial_sum3_aux k' u (pow / 3)%R (S i)
20 | end.
21 |
22 | Definition partial_sum3 u k := partial_sum3_aux k u 1%R 0.
23 |
24 | (* Σ (i=0,c-1) 3^(c-1-i)ui *)
25 | Fixpoint n_partial_sum3 (u : ℕ → bool) c :=
26 | match c with
27 | | O => O
28 | | S c' => (3 * n_partial_sum3 u c' + Nat.b2n (u c'))%nat
29 | end.
30 |
31 | Definition b2r b := INR (Nat.b2n b).
32 |
33 | Theorem partial_sum3_aux_le_half_pow : ∀ u k pow pow2 i,
34 | (0 ≤ pow)%R
35 | → pow2 = (pow / 2)%R
36 | → (partial_sum3_aux k u pow i ≤ pow2)%R.
37 | Proof.
38 | intros * Hpow Hpow2; subst pow2.
39 | revert pow i Hpow.
40 | induction k; intros; simpl; [ lra | ].
41 | destruct (u i). {
42 | apply Rplus_le_reg_l with (r := (- (pow / 3))%R).
43 | rewrite <- Rplus_assoc, Rplus_opp_l, Rplus_0_l.
44 | eapply Rle_trans; [ apply IHk; lra | lra ].
45 | }
46 | eapply Rle_trans; [ apply IHk; lra | lra ].
47 | Qed.
48 |
49 | Theorem partial_sum3_aux_succ : ∀ u n pow i,
50 | partial_sum3_aux (S n) u pow i =
51 | (partial_sum3_aux n u pow i +
52 | INR (Nat.b2n (u (i + n)%nat)) * pow / 3 ^ S n)%R.
53 | Proof.
54 | intros.
55 | revert pow i.
56 | induction n; intros. {
57 | simpl; rewrite Rplus_0_r, Rplus_0_l, Rmult_1_r, Nat.add_0_r.
58 | destruct (u i); simpl; lra.
59 | }
60 | remember (S n) as sn; simpl; subst sn.
61 | remember (u i) as bi eqn:Hbi; symmetry in Hbi.
62 | destruct bi. {
63 | remember (3 ^ S n)%R as sn3 eqn:Hsn3.
64 | rewrite IHn; simpl; rewrite Hbi.
65 | rewrite Rplus_assoc.
66 | do 2 apply Rplus_eq_compat_l.
67 | rewrite <- Nat.add_succ_comm.
68 | unfold Rdiv; subst sn3.
69 | rewrite Rinv_mult.
70 | now do 3 rewrite Rmult_assoc.
71 | } {
72 | remember (3 ^ S n)%R as sn3 eqn:Hsn3.
73 | rewrite IHn; simpl; rewrite Hbi.
74 | rewrite <- Nat.add_succ_comm.
75 | apply Rplus_eq_compat_l.
76 | unfold Rdiv; subst sn3.
77 | rewrite Rinv_mult.
78 | now do 3 rewrite Rmult_assoc.
79 | }
80 | Qed.
81 |
82 | Theorem partial_sum3_succ : ∀ u n,
83 | (partial_sum3 u (S n) =
84 | partial_sum3 u n + INR (Nat.b2n (u n)) / 3 ^ S n)%R.
85 | Proof.
86 | intros.
87 | unfold partial_sum3.
88 | rewrite partial_sum3_aux_succ.
89 | now rewrite Rmult_1_r.
90 | Qed.
91 |
92 | Theorem partial_sum3_aux_add : ∀ u k₁ k₂ pow i,
93 | partial_sum3_aux (k₁ + k₂) u pow i =
94 | (partial_sum3_aux k₁ u pow i +
95 | partial_sum3_aux k₂ u (pow / 3 ^ k₁) (i + k₁))%R.
96 | Proof.
97 | intros.
98 | revert k₂ pow i.
99 | induction k₁; intros. {
100 | now simpl; rewrite Rplus_0_l, Nat.add_0_r, Rdiv_1_r.
101 | }
102 | simpl.
103 | remember (u i) as bi eqn:Hbi; symmetry in Hbi.
104 | rewrite <- Nat.add_succ_comm.
105 | unfold Rdiv at 7.
106 | rewrite Rinv_mult.
107 | rewrite <- Rmult_assoc; do 2 rewrite fold_Rdiv.
108 | destruct bi; [ | apply IHk₁; lra ].
109 | rewrite Rplus_assoc.
110 | apply Rplus_eq_compat_l, IHk₁; lra.
111 | Qed.
112 |
113 | Theorem partial_sum3_aux_nonneg : ∀ u k pos i,
114 | (0 ≤ pos)%R
115 | → (0 ≤ partial_sum3_aux k u pos i)%R.
116 | Proof.
117 | intros * Hpos.
118 | revert pos i Hpos.
119 | induction k; intros; simpl; [ lra | ].
120 | destruct (u i); [ | apply IHk; lra ].
121 | apply Rplus_le_le_0_compat; [ lra | apply IHk; lra ].
122 | Qed.
123 |
124 | Theorem partial_sum3_upper_bound : ∀ u n k,
125 | (partial_sum3 u k ≤ partial_sum3 u n + / (2 * 3 ^ n))%R.
126 | Proof.
127 | intros.
128 | unfold partial_sum3.
129 | destruct (le_dec k n) as [Hkn| Hkn]. {
130 | remember (n - k)%nat as nk eqn:Hnk.
131 | assert (Hn : (n = k + nk)%nat). {
132 | now subst nk; rewrite Nat.add_comm, Nat.sub_add.
133 | }
134 | subst n.
135 | rewrite partial_sum3_aux_add, Nat.add_0_l, Rplus_assoc.
136 | eapply Rplus_le_reg_l; rewrite Rplus_opp_l.
137 | rewrite <- Rplus_assoc, Rplus_opp_l, Rplus_0_l.
138 | apply Rplus_le_le_0_compat. {
139 | apply partial_sum3_aux_nonneg.
140 | unfold Rdiv; rewrite Rmult_1_l.
141 | eapply Rmult_le_reg_l; [ | rewrite Rmult_0_r, Rinv_r ]. {
142 | apply pow_lt; lra.
143 | } {
144 | lra.
145 | } {
146 | apply pow_nonzero; lra.
147 | }
148 | }
149 | rewrite Rinv_mult.
150 | apply Rmult_le_pos; [ lra | ].
151 | eapply Rmult_le_reg_l; [ | rewrite Rmult_0_r, Rinv_r ]. {
152 | apply pow_lt; lra.
153 | } {
154 | lra.
155 | } {
156 | apply pow_nonzero; lra.
157 | }
158 | }
159 | apply Nat.nle_gt in Hkn.
160 | remember (k - n)%nat as nk eqn:Hnk.
161 | assert (Hn : (k = n + nk)%nat). {
162 | subst nk; rewrite Nat.add_comm, Nat.sub_add; [ easy | ].
163 | now apply Nat.lt_le_incl.
164 | }
165 | subst k; clear Hnk Hkn; rename nk into k.
166 | rewrite partial_sum3_aux_add, Nat.add_0_l.
167 | apply Rplus_le_compat_l.
168 | revert n.
169 | induction k; intros; simpl. {
170 | rewrite Rinv_mult.
171 | apply Rmult_le_pos; [ lra | ].
172 | eapply Rmult_le_reg_l; [ | rewrite Rmult_0_r, Rinv_r ]. {
173 | apply pow_lt; lra.
174 | } {
175 | lra.
176 | } {
177 | apply pow_nonzero; lra.
178 | }
179 | }
180 | remember (u n) as b eqn:Hb; symmetry in Hb.
181 | destruct b. {
182 | apply Rplus_le_reg_l with (r := (- (1 / 3 ^ n / 3))%R).
183 | rewrite <- Rplus_assoc, Rplus_opp_l, Rplus_0_l.
184 | field_simplify; [ | apply pow_nonzero; lra ].
185 | apply partial_sum3_aux_le_half_pow. {
186 | unfold Rdiv; rewrite Rmult_1_l.
187 | apply Rmult_le_pos; [ | lra ].
188 | eapply Rmult_le_reg_l; [ | rewrite Rmult_0_r, Rinv_r ]. {
189 | apply pow_lt; lra.
190 | } {
191 | lra.
192 | } {
193 | apply pow_nonzero; lra.
194 | }
195 | }
196 | unfold Rdiv.
197 | rewrite Rinv_mult; lra.
198 | }
199 | replace (1 / 3 ^ n / 3)%R with (1 / (3 ^ S n))%R. {
200 | eapply Rle_trans; [ apply IHk | ].
201 | apply Rinv_le_contravar. {
202 | apply Rmult_lt_0_compat; [ lra | apply pow_lt; lra ].
203 | } {
204 | apply Rmult_le_compat_l; [ lra | ].
205 | apply Rle_pow; [ lra | apply Nat.le_succ_diag_r ].
206 | }
207 | }
208 | simpl; unfold Rdiv.
209 | rewrite Rinv_mult; lra.
210 | Qed.
211 |
212 | Theorem partial_sum3_aux_shift_seq : ∀ u k pow i,
213 | partial_sum3_aux (S k) u pow i =
214 | ((pow * b2r (u i) + partial_sum3_aux k (λ n, u (S n)) pow i) / 3)%R.
215 | Proof.
216 | intros.
217 | set (v := λ n, u (S n)).
218 | revert pow i.
219 | induction k; intros; [ simpl; destruct (u i); unfold b2r; simpl; lra | ].
220 | rewrite partial_sum3_aux_succ.
221 | rewrite IHk.
222 | rewrite partial_sum3_aux_succ.
223 | set (x := partial_sum3_aux k v pow i).
224 | unfold v; rewrite <- Nat.add_succ_comm; simpl.
225 | set (y := INR (Nat.b2n (u (S (i + k))))).
226 | field_simplify; [ easy | | ]; apply pow_nonzero; lra.
227 | Qed.
228 |
229 | Theorem n_partial_sum3_succ2 : ∀ u n,
230 | n_partial_sum3 u (S n) =
231 | (3 ^ n * Nat.b2n (u O) + n_partial_sum3 (λ n, u (S n)) n)%nat.
232 | Proof.
233 | intros.
234 | set (v n := u (S n)).
235 | induction n; [ now simpl; do 2 rewrite Nat.add_0_r | ].
236 | remember (S n) as sn; simpl; subst sn.
237 | rewrite IHn; simpl.
238 | set (m := n_partial_sum3 v n).
239 | subst v.
240 | do 3 rewrite Nat.add_0_r.
241 | ring.
242 | Qed.
243 |
244 | Theorem n_partial_sum3_succ : ∀ u n,
245 | n_partial_sum3 u (S n) = (3 * n_partial_sum3 u n + Nat.b2n (u n))%nat.
246 | Proof. easy. Qed.
247 |
248 | Theorem partial_sum3_n_partial_sum3 : ∀ u n,
249 | (3 ^ n * partial_sum3 u n = INR (n_partial_sum3 u n))%R.
250 | Proof.
251 | intros.
252 | unfold partial_sum3.
253 | induction n; [ simpl; lra | ].
254 | rewrite partial_sum3_aux_succ, n_partial_sum3_succ.
255 | rewrite plus_INR, mult_INR; simpl.
256 | replace (2 + 1)%R with 3%R by lra.
257 | rewrite Rmult_plus_distr_l.
258 | rewrite Rmult_assoc.
259 | rewrite IHn.
260 | replace (1 + 1 + 1) with 3 by lra.
261 | apply Rplus_eq_compat_l.
262 | rewrite Rmult_comm.
263 | unfold Rdiv.
264 | rewrite Rmult_1_r.
265 | rewrite Rmult_assoc, Rinv_l, Rmult_1_r; [ easy | ].
266 | apply Rmult_integral_contrapositive.
267 | split; [ lra | ].
268 | apply pow_nonzero; lra.
269 | Qed.
270 |
271 | Theorem le_partial_sum3_lt_n_partial_sum3 : ∀ u r n,
272 | (r ≤ partial_sum3 u (S n) + / (2 * 3 ^ S n))%R
273 | → (r * 3 ^ n < INR (n_partial_sum3 u n) + 1)%R.
274 | Proof.
275 | intros * Hr2.
276 | apply Rmult_le_compat_r with (r := (3 ^ n)%R) in Hr2; [ | apply pow_le; lra ].
277 | rewrite Rmult_plus_distr_r in Hr2.
278 | rewrite Rinv_mult in Hr2.
279 | simpl in Hr2.
280 | rewrite Rinv_mult in Hr2.
281 | rewrite <- Rmult_assoc in Hr2.
282 | rewrite Rmult_assoc in Hr2.
283 | rewrite Rinv_l in Hr2; [ | apply pow_nonzero; lra ].
284 | rewrite Rmult_1_r in Hr2.
285 | rewrite <- Rinv_mult in Hr2.
286 | setoid_rewrite Rmult_comm in Hr2 at 2.
287 | rewrite partial_sum3_succ in Hr2.
288 | rewrite Rmult_plus_distr_l in Hr2.
289 | unfold Rdiv in Hr2; simpl in Hr2.
290 | rewrite Rinv_mult in Hr2.
291 | rewrite <- Rmult_assoc, Rmult_shuffle0 in Hr2.
292 | rewrite <- Rmult_assoc in Hr2.
293 | rewrite Rmult_assoc, Rmult_shuffle0 in Hr2.
294 | rewrite <- Rmult_assoc in Hr2.
295 | rewrite Rinv_r in Hr2; [ | apply pow_nonzero; lra ].
296 | rewrite Rmult_1_l in Hr2.
297 | rewrite partial_sum3_n_partial_sum3 in Hr2.
298 | destruct (u n); simpl in Hr2; lra.
299 | Qed.
300 |
301 | Theorem Int_part_n_partial_sum3 : ∀ u r n,
302 | (∀ k, (partial_sum3 u k ≤ r)%R)
303 | → (∀ b, (∀ k, (partial_sum3 u k ≤ b)%R) → (r ≤ b)%R)
304 | → Int_part (r * 3 ^ n) = Z.of_nat (n_partial_sum3 u n).
305 | Proof.
306 | intros * Hr1 Hr2.
307 | specialize (Hr1 (S n)).
308 | assert (H : (r ≤ partial_sum3 u (S n) + / (2 * 3 ^ S n))%R). {
309 | apply Hr2, partial_sum3_upper_bound.
310 | }
311 | clear Hr2; rename H into Hr2.
312 | rewrite (Int_part_interv (Z.of_nat (n_partial_sum3 u n))); [ easy | ].
313 | rewrite plus_IZR, <- INR_IZR_INZ; simpl.
314 | split. {
315 | revert u r Hr1 Hr2.
316 | induction n; intros. {
317 | unfold partial_sum3 in Hr1, Hr2; simpl in Hr1, Hr2; simpl.
318 | destruct (u O); simpl; lra.
319 | }
320 | unfold partial_sum3 in Hr1, Hr2.
321 | rewrite partial_sum3_aux_shift_seq in Hr1, Hr2.
322 | rewrite Rmult_1_l in Hr1, Hr2.
323 | rewrite n_partial_sum3_succ2.
324 | remember (u O) as b eqn:Hb; symmetry in Hb.
325 | unfold b2r in Hr1, Hr2.
326 | destruct b. {
327 | remember (S n) as sn; simpl in Hr1, Hr2; subst sn.
328 | simpl; rewrite Nat.mul_1_r.
329 | set (v n := u (S n)) in *.
330 | rewrite plus_INR.
331 | apply Rplus_le_reg_l with (r := (- INR (3 ^ n))%R).
332 | rewrite <- Rplus_assoc, Rplus_opp_l, Rplus_0_l.
333 | rewrite Rplus_comm.
334 | rewrite pow_INR; simpl.
335 | replace (2 + 1)%R with 3%R by lra.
336 | replace (- 3 ^ n)%R with ((- 1) * 3 ^ n)%R by lra.
337 | replace (1 + 1 + 1) with 3 by lra.
338 | replace (3 ^ n) with (1 * 3 ^ n) at 2 by lra.
339 | rewrite fold_Rminus.
340 | rewrite <- Rmult_assoc.
341 | rewrite <- Rmult_minus_distr_r.
342 | apply IHn; [ unfold partial_sum3; lra | ].
343 | unfold partial_sum3.
344 | set (x := partial_sum3_aux (S n) v 1 0) in *.
345 | apply Rplus_le_reg_r with (r := 1%R).
346 | replace (r * 3 - 1 + 1)%R with (r * 3)%R by lra.
347 | remember 3%R as three.
348 | rewrite Rplus_comm, <- Rplus_assoc; subst three.
349 | apply Rmult_le_reg_r with (r := (/ 3)%R); [ lra | ].
350 | rewrite Rmult_assoc, Rinv_r; [ | lra ].
351 | rewrite Rmult_1_r.
352 | rewrite Rmult_plus_distr_r.
353 | rewrite fold_Rdiv.
354 | rewrite <- Rinv_mult.
355 | now rewrite <- Rmult_assoc in Hr2; rewrite Rmult_shuffle0.
356 | }
357 | remember (S n) as sn; simpl in Hr1, Hr2; subst sn.
358 | simpl; rewrite Nat.mul_0_r, Nat.add_0_l.
359 | rewrite Rplus_0_l in Hr1, Hr2.
360 | set (v n := u (S n)) in *.
361 | rewrite <- Rmult_assoc.
362 | apply IHn; [ unfold partial_sum3; lra | ].
363 | unfold partial_sum3.
364 | set (x := partial_sum3_aux (S n) v 1 0) in *.
365 | apply Rmult_le_reg_r with (r := (/ 3)%R); [ lra | ].
366 | rewrite Rmult_assoc, Rinv_r; [ | lra ].
367 | rewrite Rmult_1_r.
368 | rewrite Rmult_plus_distr_r.
369 | rewrite fold_Rdiv.
370 | rewrite <- Rinv_mult.
371 | now rewrite <- Rmult_assoc in Hr2; rewrite Rmult_shuffle0.
372 | }
373 | now apply le_partial_sum3_lt_n_partial_sum3.
374 | Qed.
375 |
376 | Theorem IZR_Int_part_mult_pow_succ : ∀ u r n,
377 | (∀ k, (partial_sum3 u k ≤ r)%R)
378 | → (∀ b, (∀ k, (partial_sum3 u k ≤ b)%R) → (r ≤ b)%R)
379 | → IZR (Int_part (r * 3 ^ S n)) =
380 | (3 * IZR (Int_part (r * 3 ^ n)) + INR (Nat.b2n (u n)))%R.
381 | Proof.
382 | intros * Hr1 Hr2.
383 | rewrite Int_part_n_partial_sum3 with (u := u); [ | easy | easy ].
384 | rewrite Int_part_n_partial_sum3 with (u := u); [ | easy | easy ].
385 | do 2 rewrite <- INR_IZR_INZ.
386 | rewrite n_partial_sum3_succ.
387 | rewrite plus_INR, mult_INR.
388 | now replace (INR 3) with 3%R by (simpl; lra).
389 | Qed.
390 |
391 | Theorem Int_part_eq_partial_sum3 : ∀ u r n,
392 | (∀ k : nat, (partial_sum3 u k ≤ r)%R)
393 | → (∀ b : R, (∀ k : nat, (partial_sum3 u k ≤ b)%R) → (r ≤ b)%R)
394 | → IZR (Int_part (r * 3 ^ n)) = (partial_sum3 u n * 3 ^ n)%R.
395 | Proof.
396 | intros * Hk1 Hk2.
397 | induction n. {
398 | unfold partial_sum3; simpl.
399 | do 2 rewrite Rmult_1_r.
400 | specialize (Hk1 O); simpl in Hk1.
401 | unfold partial_sum3 in Hk1; simpl in Hk1.
402 | assert (H : ∀ k, (partial_sum3 u k ≤ 1 / 2)%R). {
403 | intros k; apply partial_sum3_aux_le_half_pow; lra.
404 | }
405 | specialize (Hk2 (1 / 2)%R H).
406 | replace 0%R with (IZR 0) by easy.
407 | apply IZR_eq, Int_part_interv; simpl; lra.
408 | }
409 | rewrite partial_sum3_succ.
410 | rewrite Rmult_plus_distr_r.
411 | unfold Rdiv; rewrite Rmult_assoc.
412 | rewrite Rinv_l; [ | apply pow_nonzero; lra ].
413 | rewrite Rmult_1_r.
414 | remember (r * 3 ^ S n)%R as x; simpl; subst x.
415 | rewrite <- Rmult_assoc, Rmult_shuffle0, <- IHn.
416 | setoid_rewrite Rmult_comm at 3.
417 | now apply IZR_Int_part_mult_pow_succ.
418 | Qed.
419 |
420 | Theorem ter_bin_of_frac_part_surj : ∀ u : nat → bool,
421 | ∃ r : R, (0 ≤ r < 1)%R ∧ (∀ n, ter_bin_of_frac_part r n = u n).
422 | Proof.
423 | intros.
424 | set (E x := ∃ k, partial_sum3 u k = x).
425 | assert (Hb : bound E). {
426 | exists (1 / 2)%R; subst E; simpl.
427 | intros r (k & H); subst r.
428 | apply partial_sum3_aux_le_half_pow; lra.
429 | }
430 | assert (He : ∃ r, E r). {
431 | exists 0%R; subst E; simpl.
432 | now exists O; unfold partial_sum3.
433 | }
434 | destruct (completeness E Hb He) as (r & Hr1 & Hr2).
435 | assert (Hr3 : (∀ k, partial_sum3 u k ≤ r)%R). {
436 | unfold is_upper_bound, E in Hr1; simpl in Hr1.
437 | now intros k; apply Hr1; exists k.
438 | }
439 | assert (Hr4 : (∀ b, (∀ k, partial_sum3 u k ≤ b) → (r ≤ b))%R). {
440 | unfold is_upper_bound, E in Hr2; simpl in Hr2.
441 | intros b H; apply Hr2; intros x (k, Hx); subst x.
442 | apply H.
443 | }
444 | assert (Hh : (r ≤ 1 / 2)%R). {
445 | apply Hr2; unfold E; simpl.
446 | intros x (k & H); subst x.
447 | apply partial_sum3_aux_le_half_pow; lra.
448 | }
449 | exists r; clear Hb He; simpl.
450 | split. {
451 | split; [ | lra ].
452 | apply Hr1; unfold E; simpl.
453 | now exists O; unfold partial_sum3.
454 | }
455 | intros n.
456 | clear E Hr1 Hr2.
457 | unfold ter_bin_of_frac_part; symmetry.
458 | destruct (Rlt_dec (frac_part (r * 3 ^ n)) (1 / 3)) as [H1| H1]. {
459 | unfold frac_part in H1.
460 | rewrite (Int_part_eq_partial_sum3 u) in H1; [ | easy | easy ].
461 | unfold Rminus in H1.
462 | rewrite Ropp_mult_distr_l in H1.
463 | rewrite <- Rmult_plus_distr_r in H1.
464 | rewrite fold_Rminus in H1.
465 | apply Rmult_lt_compat_r with (r := (/ 3 ^ n)%R) in H1. {
466 | rewrite Rmult_assoc in H1.
467 | rewrite Rinv_r in H1; [ | apply pow_nonzero; lra ].
468 | rewrite Rmult_1_r in H1.
469 | unfold Rdiv in H1.
470 | rewrite Rmult_assoc in H1.
471 | rewrite <- Rinv_mult in H1.
472 | replace (3 * 3 ^ n)%R with (3 ^ S n)%R in H1 by easy.
473 | fold (Rdiv 1 (3 ^ S n)) in H1.
474 | specialize (Hr3 (S n)).
475 | rewrite partial_sum3_succ in Hr3.
476 | destruct (u n); [ exfalso | easy ].
477 | simpl in Hr3, H1; lra.
478 | }
479 | apply Rinv_0_lt_compat, pow_lt; lra.
480 | }
481 | apply Rnot_lt_le in H1.
482 | unfold frac_part in H1.
483 | rewrite (Int_part_eq_partial_sum3 u) in H1; [ | easy | easy ].
484 | unfold Rminus in H1.
485 | rewrite Ropp_mult_distr_l in H1.
486 | rewrite <- Rmult_plus_distr_r in H1.
487 | rewrite fold_Rminus in H1.
488 | apply Rmult_le_compat_r with (r := (/ 3 ^ n)%R) in H1. {
489 | rewrite Rmult_assoc in H1.
490 | rewrite Rinv_r in H1; [ | apply pow_nonzero; lra ].
491 | rewrite Rmult_1_r in H1.
492 | unfold Rdiv in H1; rewrite Rmult_1_l in H1.
493 | rewrite <- Rinv_mult in H1.
494 | replace (3 * 3 ^ n)%R with (3 ^ S n)%R in H1 by easy.
495 | specialize (partial_sum3_upper_bound u (S n)); intros H.
496 | apply Hr4 in H.
497 | rewrite partial_sum3_succ in H.
498 | destruct (u n); [ easy | exfalso ].
499 | simpl in H1, H.
500 | unfold Rdiv in H.
501 | rewrite Rmult_0_l, Rplus_0_r in H.
502 | rewrite Rinv_mult in H.
503 | set (s := partial_sum3 u n) in H1, H.
504 | set (t := (/ (3 * 3 ^ n))%R) in H1, H.
505 | enough (0 < t)%R by lra; subst t.
506 | apply Rinv_0_lt_compat.
507 | apply Rmult_lt_0_compat; [ lra | apply pow_lt; lra ].
508 | }
509 | apply Rlt_le.
510 | apply Rinv_0_lt_compat.
511 | apply pow_lt; lra.
512 | Qed.
513 |
514 | Definition id {A} (a : A) := a.
515 |
516 | Theorem id_nat : ∀ e : ℕ, ∃ x : ℕ, id x = e.
517 | Proof. now intros; exists e. Qed.
518 |
519 | Theorem Cantor_ℕ_ℝ : ∀ f : nat → R, ∃ x : R, ∀ n : nat, x ≠ f n.
520 | Proof.
521 | specialize
522 | (Cantor_gen ℕ ℕ ℝ (λ x, (0 ≤ x < 1)%R) id ter_bin_of_frac_part id_nat
523 | ter_bin_of_frac_part_surj).
524 | intros H f.
525 | specialize (H f).
526 | destruct H as (x, H); exists x.
527 | intros n; apply H.
528 | Qed.
529 |
--------------------------------------------------------------------------------
/Words.v:
--------------------------------------------------------------------------------
1 | (* Banach-Tarski paradox. *)
2 |
3 | From Stdlib Require Import Utf8 List Relations Wf_nat.
4 | Import ListNotations.
5 | Require Import Misc.
6 |
7 | Inductive letter := la | lb.
8 |
9 | Inductive free_elem := FE : letter → bool → free_elem.
10 |
11 | Notation "'ạ'" := (FE la false).
12 | Notation "'ạ⁻¹'" := (FE la true).
13 | Notation "'ḅ'" := (FE lb false).
14 | Notation "'ḅ⁻¹'" := (FE lb true).
15 |
16 | Definition negf '(FE t d) := FE t (negb d).
17 |
18 | Theorem letter_dec : ∀ l1 l2 : letter, {l1 = l2} + {l1 ≠ l2}.
19 | Proof.
20 | intros.
21 | now destruct l1, l2; try (now left); right.
22 | Defined.
23 |
24 | Theorem free_elem_dec : ∀ e₁ e₂ : free_elem, { e₁ = e₂ } + { e₁ ≠ e₂ }.
25 | Proof.
26 | intros.
27 | destruct e₁ as (t₁, d₁).
28 | destruct e₂ as (t₂, d₂).
29 | destruct (letter_dec t₁ t₂) as [H₁| H₁]; [ subst t₂ | ]. {
30 | destruct (Bool.bool_dec d₁ d₂) as [H₂| H₂]; [ subst d₂ | ]. {
31 | now left.
32 | } {
33 | right; intros H; apply H₂.
34 | now injection H.
35 | }
36 | } {
37 | right; intros H; apply H₁.
38 | now injection H.
39 | }
40 | Qed.
41 |
42 | Theorem letter_dec_diag : ∀ t, letter_dec t t = left (eq_refl _).
43 | Proof.
44 | intros t.
45 | destruct (letter_dec t t) as [p| p]; [ | exfalso; now apply p ].
46 | destruct t; refine (match p with eq_refl => eq_refl end).
47 | Qed.
48 |
49 | Definition letter_opp '(FE l₁ d₁) '(FE l₂ d₂) :=
50 | if letter_dec l₁ l₂ then
51 | if Bool.bool_dec d₁ d₂ then False else True
52 | else False.
53 |
54 | Theorem letter_opp_dec : ∀ e₁ e₂,
55 | {letter_opp e₁ e₂} + {not (letter_opp e₁ e₂)}.
56 | Proof.
57 | intros.
58 | destruct e₁ as (x₁, d₁).
59 | destruct e₂ as (x₂, d₂); simpl.
60 | destruct (letter_dec x₁ x₂) as [Hx| Hx]; [ | now right ].
61 | destruct (Bool.bool_dec d₁ d₂) as [Hd| Hd]; [ | left; constructor ].
62 | now right.
63 | Defined.
64 |
65 | Theorem letter_opp_inv : ∀ x d, letter_opp (FE x d) (FE x (negb d)).
66 | Proof.
67 | intros.
68 | unfold letter_opp.
69 | now rewrite letter_dec_diag, bool_dec_negb_r.
70 | Qed.
71 |
72 | Theorem letter_opp_iff : ∀ x₁ d₁ x₂ d₂,
73 | letter_opp (FE x₁ d₁) (FE x₂ d₂)
74 | ↔ x₁ = x₂ ∧ d₂ = negb d₁.
75 | Proof.
76 | intros x₁ d₁ x₂ d₂.
77 | split; intros H. {
78 | unfold letter_opp in H.
79 | destruct (letter_dec x₁ x₂) as [H₁| H₁]; [ | easy ].
80 | split; [ easy | ].
81 | destruct (Bool.bool_dec d₁ d₂) as [H₂| H₂]; [ easy | ].
82 | now apply neq_negb, not_eq_sym.
83 | } {
84 | destruct H; subst x₂ d₂.
85 | apply letter_opp_inv.
86 | }
87 | Qed.
88 |
89 | Theorem letter_opp_negf : ∀ e₁ e₂, letter_opp e₁ e₂ ↔ e₁ = negf e₂.
90 | Proof.
91 | intros.
92 | destruct e₁ as (t₁, d₁).
93 | destruct e₂ as (t₂, d₂).
94 | split; intros H. {
95 | apply letter_opp_iff in H.
96 | destruct H; subst t₂ d₂; simpl.
97 | now rewrite Bool.negb_involutive.
98 | } {
99 | injection H; intros; subst; simpl.
100 | now rewrite letter_dec_diag, bool_dec_negb_l.
101 | }
102 | Qed.
103 |
104 | Theorem no_fixpoint_negf : ∀ e, negf e ≠ e.
105 | Proof.
106 | intros * H.
107 | destruct e as (t, d).
108 | injection H.
109 | apply Bool.no_fixpoint_negb.
110 | Qed.
111 |
112 | Theorem negf_involutive : ∀ e, negf (negf e) = e.
113 | Proof.
114 | intros (t, d); simpl.
115 | now rewrite Bool.negb_involutive.
116 | Qed.
117 |
118 | Theorem letter_opp_negf_r : ∀ e, letter_opp e (negf e).
119 | Proof.
120 | intros.
121 | apply letter_opp_negf.
122 | now rewrite negf_involutive.
123 | Qed.
124 |
125 | Theorem letter_opp_sym : ∀ e₁ e₂, letter_opp e₁ e₂ → letter_opp e₂ e₁.
126 | Proof.
127 | intros * H.
128 | apply letter_opp_negf in H.
129 | subst e₁.
130 | apply letter_opp_negf_r.
131 | Qed.
132 |
133 | Theorem negf_eq_eq : ∀ e₁ e₂, negf e₁ = negf e₂ → e₁ = e₂.
134 | Proof.
135 | intros e₁ e₂ Hn.
136 | destruct e₁ as (t₁, d₁).
137 | destruct e₂ as (t₂, d₂).
138 | simpl in Hn.
139 | injection Hn; intros H₁ H₂; subst.
140 | now apply negb_eq_eq in H₁; subst d₁.
141 | Qed.
142 |
--------------------------------------------------------------------------------
/interface.sh:
--------------------------------------------------------------------------------
1 | #!/bin/sh
2 | # Print interfaces of a .v file,
3 | # i.e. all but Proof..Qed or Proof..Defined
4 | grep -v 'Proof.*Qed' $* | sed -e '/Proof/,/\(Qed\|Defined\)/d'
5 |
--------------------------------------------------------------------------------