This commit is contained in:
Thomas Baruchel 2023-02-11 16:02:22 +01:00
parent 6a9c177d2f
commit 8d067d4ae7
3 changed files with 1122 additions and 31 deletions

View File

@ -1287,6 +1287,7 @@ Lemma tm_step_factor4_1mod4 : forall (n' : nat) (w z y : list bool),
tm_step n' = w ++ z ++ y tm_step n' = w ++ z ++ y
-> length z = 4 -> length w mod 4 = 1 -> length z = 4 -> length w mod 4 = 1
-> exists (x : bool), firstn 2 z = [x;x]. -> exists (x : bool), firstn 2 z = [x;x].
Proof.
intros n' w z y. intros G0 G1 G2. intros n' w z y. intros G0 G1 G2.
assert (U: 4 * (length w / 4) <= length w). assert (U: 4 * (length w / 4) <= length w).
apply Nat.mul_div_le. easy. apply Nat.mul_div_le. easy.

View File

@ -308,8 +308,14 @@ Lemma tm_step_palindromic_length_8 :
forall (n : nat) (hd a tl : list bool), forall (n : nat) (hd a tl : list bool),
tm_step n = hd ++ a ++ (rev a) ++ tl tm_step n = hd ++ a ++ (rev a) ++ tl
-> length a = 4 -> length a = 4
-> length (hd ++ a) mod 4 = 0 -> (
\/ nth_error hd (length hd - 3) <> nth_error tl 2. length (hd ++ a) mod 4 = 0
/\ exists b, a = [b; negb b; negb b; b]
) \/ (
length (hd ++ a) mod 4 = 2
/\ (exists b, a = [b; negb b; b; negb b])
/\ nth_error hd (length hd - 3) <> nth_error tl 2
).
Proof. Proof.
intros n hd a tl. intros H I. intros n hd a tl. intros H I.
@ -450,30 +456,44 @@ Proof.
rewrite <- app_assoc in H. rewrite <- app_assoc in H.
assert ({last (b3::hd) false=b1} + {~ last (b3::hd) false=b1}). assert ({last (b3::hd) false=b1} + {~ last (b3::hd) false=b1}).
apply bool_dec. destruct H1. rewrite e0 in H. apply bool_dec. destruct H1. rewrite e0 in H.
assert (b2 = b). destruct b2; destruct b1; destruct b.
reflexivity. contradiction n0. reflexivity. reflexivity.
contradiction n1. reflexivity. contradiction n1. reflexivity.
reflexivity. contradiction n0. reflexivity. reflexivity. rewrite H1 in H.
assert (R: b1 = negb b). destruct b; destruct b1;
reflexivity || ( rewrite H1 in n0; contradiction n0; reflexivity).
(* un cas à étudier (* un cas à étudier
(??? T) | F T T F || F T T F | ??? impossible à gauche (??? T) | F T T F || F T T F | ??? impossible à gauche
*) *)
destruct M. left. assumption. (* ici on postule le modulo = 2 *) destruct M. left.
split. assumption.
rewrite e. rewrite R. rewrite H1. exists b. reflexivity.
(* suite *)
rewrite app_removelast_last with (l := b3::hd) (d := false) in Q. rewrite app_removelast_last with (l := b3::hd) (d := false) in Q.
rewrite last_length in Q. rewrite Nat.even_succ in Q. apply odd_mod4 in Q. rewrite last_length in Q. rewrite Nat.even_succ in Q. apply odd_mod4 in Q.
rewrite app_removelast_last with (l := b3::hd) (d := false) in H1. rewrite app_removelast_last with (l := b3::hd) (d := false) in H2.
destruct Q. destruct Q.
assert (exists x, firstn 2 [b1;b;b1;b1] = [x;x]). generalize H2. assert (exists x, firstn 2 [b1;b;b1;b1] = [x;x]). generalize H3.
assert (length [b1;b;b1;b1] = 4). reflexivity. generalize H3. assert (length [b1;b;b1;b1] = 4). reflexivity. generalize H4.
replace ( replace (
removelast (b3 :: hd) ++ removelast (b3 :: hd) ++
[b1] ++ b :: b1 :: b1 :: b2 :: b2 :: b1 :: b1 :: b :: b4 :: tl ) [b1] ++ b :: b1 :: b1 :: b :: b :: b1 :: b1 :: b :: b4 :: tl )
with ( with (
removelast (b3 :: hd) ++ [b1;b;b1;b1] removelast (b3 :: hd) ++ [b1;b;b1;b1]
++ (b2 :: b2 :: b1 :: b1 :: b :: b4 :: tl )) in H. ++ (b :: b :: b1 :: b1 :: b :: b4 :: tl )) in H.
generalize H. apply tm_step_factor4_1mod4. reflexivity. generalize H. apply tm_step_factor4_1mod4. reflexivity.
destruct H3. inversion H3. rewrite <- H6 in H5. rewrite H5 in n1. destruct H4. inversion H4. rewrite <- H7 in H6. rewrite H6 in n1.
contradiction n1. reflexivity. contradiction n1. reflexivity.
rewrite app_length in H1. rewrite app_length in H1. rewrite app_length in H2. rewrite app_length in H2.
rewrite <- Nat.add_assoc in H1. rewrite <- Nat.add_mod_idemp_l in H1. rewrite <- Nat.add_assoc in H2. rewrite <- Nat.add_mod_idemp_l in H2.
rewrite H2 in H1. inversion H1. easy. easy. easy. rewrite H3 in H2. inversion H2.
easy. easy. easy.
(* Deux cas (* Deux cas
(T F) | F T T F || F T T F | (F T) cube (T F) | F T T F || F T T F | (F T) cube
@ -483,23 +503,35 @@ Proof.
(* un sous-cas : (* un sous-cas :
(T F) | F T T F || F T T F | (T F) impossible à droite (T F) | F T T F || F T T F | (T F) impossible à droite
*) *)
destruct M. left. assumption. (* ici on postule le modulo = 2 *)
rewrite e in H1. assert (b2 = b). destruct b2; destruct b1; destruct b.
reflexivity. contradiction n0. reflexivity. reflexivity.
contradiction n1. reflexivity. contradiction n1. reflexivity.
reflexivity. contradiction n0. reflexivity. reflexivity. rewrite H1 in H.
assert (R: b1 = negb b). destruct b; destruct b1;
reflexivity || ( rewrite H1 in n0; contradiction n0; reflexivity).
destruct M. left.
split. assumption. (* ici on postule le modulo = 2 *)
rewrite e. rewrite R. rewrite H1. exists b. reflexivity.
rewrite e in H2.
assert (length (((b3 :: hd) ++ [b; b1; b1; b2]) ++ [b2]) mod 4 = 3). assert (length (((b3 :: hd) ++ [b; b1; b1; b2]) ++ [b2]) mod 4 = 3).
rewrite app_length. rewrite <- Nat.add_mod_idemp_l. rewrite H1. rewrite app_length. rewrite <- Nat.add_mod_idemp_l. rewrite H2.
reflexivity. easy. reflexivity. easy.
replace( replace(
removelast (b3 :: hd) ++ [last (b3 :: hd) false] ++ removelast (b3 :: hd) ++ [last (b3 :: hd) false] ++
b :: b1 :: b1 :: b2 :: b2 :: b1 :: b1 :: b :: b1 :: tl ) b :: b1 :: b1 :: b :: b :: b1 :: b1 :: b :: b1 :: tl )
with ( with (
(((b3 :: hd) ++ [b; b1; b1; b2]) ++ [b2]) ++ [b1; b1;b;b1] ++ tl) in H. (((b3 :: hd) ++ [b; b1; b1; b]) ++ [b]) ++ [b1; b1;b;b1] ++ tl) in H.
assert (exists x, skipn 2 [b1;b1;b;b1] = [x;x]). generalize H2. rewrite H1 in H3.
assert (length [b1;b1;b;b1] = 4). reflexivity. generalize H3. assert (exists x, skipn 2 [b1;b1;b;b1] = [x;x]). generalize H3.
assert (length [b1;b1;b;b1] = 4). reflexivity. generalize H4.
generalize H. apply tm_step_factor4_3mod4. generalize H. apply tm_step_factor4_3mod4.
destruct H3. inversion H3. rewrite <- H6 in H5. rewrite H5 in n1. destruct H4. inversion H4. rewrite <- H7 in H6. rewrite H6 in n1.
contradiction n1. reflexivity. contradiction n1. reflexivity.
symmetry. rewrite app_assoc. rewrite <- app_removelast_last. symmetry. rewrite app_assoc. rewrite <- app_removelast_last.
rewrite <- app_assoc. rewrite <- app_assoc. reflexivity. easy. rewrite <- app_assoc. rewrite <- app_assoc. reflexivity. easy.
@ -601,8 +633,28 @@ reflexivity. contradiction n0. reflexivity. reflexivity. rewrite H1 in H.
apply bool_dec. destruct H1. rewrite e0 in H. apply bool_dec. destruct H1. rewrite e0 in H.
(* problème à gauche *) (* problème à gauche *)
destruct M. left. assumption. (* ici on postule le modulo = 2 *) destruct M.
(* problème avec n1 !!! *)
assert (T1: b0 = b1).
replace (removelast (b3 :: hd) ++
[b0] ++ b1 :: b0 :: b1 :: b2 :: b2 :: b1 :: b0 :: b1 :: b4 :: tl)
with (((removelast (b3 :: hd) ++ [b0]) ++ [b1; b0; b1; b2; b2])
++ [b1; b0; b1; b4] ++ tl) in H.
assert (T2: exists (x : bool), firstn 2 [b1;b0;b1;b4] = [x;x]).
assert (length ((b3 :: hd) ++ [b; b0; b1; b2; b2]) mod 4 = 1).
replace ((b3 :: hd) ++ [b; b0; b1; b2; b2])
with (((b3 :: hd) ++ [b; b0; b1; b2]) ++ [b2]).
rewrite app_length. rewrite <- Nat.add_mod_idemp_l.
rewrite H1. reflexivity. easy. rewrite <- app_assoc. reflexivity.
generalize H2. assert (length ([b1;b0;b1;b4]) = 4). reflexivity.
generalize H3. generalize H. rewrite e. rewrite <- e0 at 1.
rewrite <- app_removelast_last. apply tm_step_factor4_1mod4.
easy. inversion T2. inversion H2. reflexivity.
rewrite <- app_assoc. rewrite <- app_assoc. reflexivity.
rewrite T1 in n1. contradiction n1. reflexivity.
(* ici on postule le modulo = 2 *)
rewrite app_removelast_last with (l := b3::hd) (d := false) in Q. rewrite app_removelast_last with (l := b3::hd) (d := false) in Q.
rewrite last_length in Q. rewrite Nat.even_succ in Q. apply odd_mod4 in Q. rewrite last_length in Q. rewrite Nat.even_succ in Q. apply odd_mod4 in Q.
destruct Q. destruct Q.
@ -635,8 +687,27 @@ reflexivity. contradiction n0. reflexivity. reflexivity. rewrite H1 in H.
(F T) | T F T F || F T F T | (T F) 4n+2 a/rev a/a possible ? (F T) | T F T F || F T F T | (T F) 4n+2 a/rev a/a possible ?
*) *)
(* régler d'abord la question du modulo au centre *) (* régler d'abord la question du modulo au centre *)
destruct M. left. assumption. (* ici on postule le modulo = 2 *) destruct M.
assert (T1: b0 = b1).
replace (removelast (b3 :: hd) ++ [last (b3::hd) false]
++ b1 :: b0 :: b1 :: b2 :: b2 :: b1 :: b0 :: b1 :: b4 :: tl)
with (((removelast (b3 :: hd) ++ [last (b3::hd) false])
++ [b1; b0; b1; b2; b2])
++ [b1; b0; b1; b4] ++ tl) in H.
assert (T2: exists (x : bool), firstn 2 [b1;b0;b1;b4] = [x;x]).
assert (length ((b3 :: hd) ++ [b; b0; b1; b2; b2]) mod 4 = 1).
replace ((b3 :: hd) ++ [b; b0; b1; b2; b2])
with (((b3 :: hd) ++ [b; b0; b1; b2]) ++ [b2]).
rewrite app_length. rewrite <- Nat.add_mod_idemp_l.
rewrite H1. reflexivity. easy. rewrite <- app_assoc. reflexivity.
generalize H2. assert (length ([b1;b0;b1;b4]) = 4). reflexivity.
generalize H3. generalize H. rewrite e.
rewrite <- app_removelast_last. apply tm_step_factor4_1mod4.
easy. inversion T2. inversion H2. reflexivity.
rewrite <- app_assoc. rewrite <- app_assoc. reflexivity.
rewrite T1 in n1. contradiction n1. reflexivity.
(* on suppose maintenant le modulo = 2 *)
assert (last (b3::hd) false = b1). assert (last (b3::hd) false = b1).
destruct (last (b3::hd) false); destruct b1; destruct b0; destruct (last (b3::hd) false); destruct b1; destruct b0;
reflexivity || contradiction n2 || contradiction n1; reflexivity. reflexivity || contradiction n2 || contradiction n1; reflexivity.
@ -851,10 +922,6 @@ Le lemme repeating_patterns se base sur les huit premiers termes de TM :
rewrite <- Nat.add_0_l at 1. apply Nat.add_le_mono. rewrite <- Nat.add_0_l at 1. apply Nat.add_le_mono.
apply Nat.le_0_l. apply Nat.le_refl. apply Nat.le_0_l. apply Nat.le_refl.
(* inutile
assert (V: nth_error (b4 :: b6 :: b9 :: tl) 2 = Some b9). reflexivity.
*)
(* analyse finale *) (* analyse finale *)
assert ({b8=b9} + {~ b8=b9}). apply bool_dec. destruct H7. rewrite e0 in H. assert ({b8=b9} + {~ b8=b9}). apply bool_dec. destruct H7. rewrite e0 in H.
@ -938,7 +1005,7 @@ Le lemme repeating_patterns se base sur les huit premiers termes de TM :
rewrite Nat.add_lt_mono_r with (p := 8). rewrite Nat.add_lt_mono_r with (p := 8).
rewrite <- Nat.add_sub_swap. rewrite Nat.add_sub. rewrite <- Nat.add_sub_swap. rewrite Nat.add_sub.
rewrite <- Nat.add_assoc. replace (2+8) with 10. rewrite <- Nat.add_assoc. replace (2+8) with 10.
replace (8) with (2^3) at 1. rewrite <- Nat.pow_add_r. replace 8 with (2^3) at 1. rewrite <- Nat.pow_add_r.
rewrite Nat.add_sub_assoc. rewrite Nat.add_sub_swap. simpl. rewrite Nat.add_sub_assoc. rewrite Nat.add_sub_swap. simpl.
rewrite <- tm_size_power2. rewrite H'. unfold lh. rewrite <- tm_size_power2. rewrite H'. unfold lh.
rewrite app_length. rewrite app_length. rewrite <- Nat.add_assoc. rewrite app_length. rewrite app_length. rewrite <- Nat.add_assoc.
@ -1193,7 +1260,15 @@ Le lemme repeating_patterns se base sur les huit premiers termes de TM :
rewrite <- length_zero_iff_nil. rewrite Y''. easy. rewrite <- length_zero_iff_nil. rewrite Y''. easy.
(* fin de la preuve, on a b8 <> b9 *) (* fin de la preuve, on a b8 <> b9 *)
right. rewrite U. simpl. injection. inversion H7. rewrite H9 in n6. right.
split. assumption.
split. rewrite H4. rewrite e.
assert (b0 = negb b1). destruct b0; destruct b1.
contradiction n1. reflexivity. reflexivity. reflexivity.
contradiction n1. reflexivity. rewrite H7. exists b1.
reflexivity.
rewrite U. simpl. injection. inversion H7. rewrite H9 in n6.
contradiction n6. reflexivity. contradiction n6. reflexivity.
rewrite removelast_firstn_len. rewrite removelast_firstn_len. simpl. rewrite removelast_firstn_len. rewrite removelast_firstn_len. simpl.
@ -1262,11 +1337,17 @@ Proof.
rewrite <- app_assoc in H. rewrite <- app_assoc in H.
pose (tl' := rev (firstn (length a - 4) a) ++ tl). fold tl' in H. pose (tl' := rev (firstn (length a - 4) a) ++ tl). fold tl' in H.
assert (K: length (hd' ++ a') mod 4 = 0 assert (K: (
\/ nth_error hd' (length hd' - 3) <> nth_error tl' 2). length (hd' ++ a') mod 4 = 0
/\ exists b, a' = [b; negb b; negb b; b]
) \/ (
length (hd' ++ a') mod 4 = 2
/\ (exists b, a' = [b; negb b; b; negb b])
/\ nth_error hd' (length hd' - 3) <> nth_error tl' 2
)).
fold a' in H1. generalize H1. generalize H. fold a' in H1. generalize H1. generalize H.
apply tm_step_palindromic_length_8. apply tm_step_palindromic_length_8.
destruct K. destruct K. destruct H2.
unfold hd' in H2. unfold a' in H2. rewrite <- app_assoc in H2. unfold hd' in H2. unfold a' in H2. rewrite <- app_assoc in H2.
rewrite firstn_skipn in H2. assumption. rewrite firstn_skipn in H2. assumption.
@ -1280,6 +1361,7 @@ Proof.
apply Nat.succ_lt_mono in I. apply Nat.succ_lt_mono in I. apply Nat.succ_lt_mono in I. apply Nat.succ_lt_mono in I.
inversion I. inversion I.
destruct H2 as [H2' H2'']. destruct H2'' as [H2'' H2].
unfold hd' in H2. unfold tl' in H2. unfold hd' in H2. unfold tl' in H2.
rewrite nth_error_app2 in H2. rewrite nth_error_app1 in H2. rewrite nth_error_app2 in H2. rewrite nth_error_app1 in H2.
rewrite app_length in H2. rewrite Nat.add_comm in H2. rewrite app_length in H2. rewrite Nat.add_comm in H2.

File diff suppressed because it is too large Load Diff