diff --git a/thys/DPRM_Theorem/DPRM.thy b/thys/DPRM_Theorem/DPRM.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/DPRM.thy @@ -0,0 +1,88 @@ +section \Proof of the DPRM theorem\ + +theory DPRM + imports "Machine_Equations/Machine_Equation_Equivalence" +begin + +definition is_recenum :: "nat set \ bool" where + "is_recenum A = + (\ p :: program. + \ n :: nat. + \ a :: nat. \ ic. ic = initial_config n a \ is_valid_initial ic p a \ + (a \ A) = (\ q::nat. terminates ic p q))" + +theorem DPRM: "is_recenum A \ is_dioph_set A" +proof - + assume "is_recenum A" + hence "(\ p :: program. + \ n :: nat. \ a :: nat. + \ ic. ic = initial_config n a \ is_valid_initial ic p a \ + (a \ A) = (\ q::nat. terminates ic p q))" using is_recenum_def by auto + then obtain p n where p: + "\ a :: nat. + \ ic. ic = initial_config n a \ is_valid_initial ic p a \ + (a \ A) = (\ q::nat. terminates ic p q)" by auto + + interpret rm: register_machine p "Suc n" apply (unfold_locales) + proof - + from p have + "\ ic. ic = initial_config n 0 \ is_valid_initial ic p 0 \ + (0 \ A) = (\ q::nat. terminates ic p q)" by auto + then obtain ic where ic: "ic = initial_config n 0" and is_val: "is_valid_initial ic p 0" by auto + + show "length p > 0" + using is_val unfolding is_valid_initial_def is_valid_def by auto + + have "length (snd ic) = Suc n" + unfolding ic initial_config_def by auto + + moreover have "snd ic \ []" + using is_val unfolding is_valid_initial_def is_valid_def tape_check_initial.simps by auto + + ultimately show "Suc n > 0" + by auto + + show "program_register_check p (Suc n)" + using is_val unfolding is_valid_initial_def is_valid_def using \length (snd ic) = Suc n\ + by auto + qed + + + have equiv: "a \ A \ register_machine.rm_equations p (Suc n) a" for a + proof - + from p have "\ ic. ic = initial_config n a \ is_valid_initial ic p a \ + (a \ A) = (\ q::nat. terminates ic p q)" by auto + then obtain ic where ic: "ic = initial_config n a \ is_valid_initial ic p a \ + (a \ A) = (\ q::nat. terminates ic p q)" by blast + + have ic_def: "ic = initial_config n a" using ic by auto + hence n_is_length: "Suc n = length (snd ic)" using initial_config_def[of "n" "a"] by auto + have is_val: "is_valid_initial ic p a" using ic by auto + have "(a \ A) = (\q. terminates ic p q)" using ic by auto + moreover have "(\q. terminates ic p q) = register_machine.rm_equations p (Suc n) a" + using is_val n_is_length rm.conclusion_4_5 + by auto + + ultimately show ?thesis by auto + qed + + hence A_characterization: "A = {a::nat. register_machine.rm_equations p (Suc n) a}" by auto + + have eq_dioph: "\P\<^sub>1 P\<^sub>2. \a. register_machine.rm_equations p (Suc n) (peval A a) + = (\v. ppeval P\<^sub>1 a v = ppeval P\<^sub>2 a v)" for A + using rm.rm_dioph[of A] using is_dioph_rel_def[of "rm.rm_equations_relation A"] + unfolding rm.rm_equations_relation_def by(auto simp: unary_eval) + + have "\P\<^sub>1 P\<^sub>2. \b. register_machine.rm_equations p (Suc n) (peval (Param 0) (\x. b)) + = (\v. ppeval P\<^sub>1 (\x. b) v = ppeval P\<^sub>2 (\x. b) v)" + using eq_dioph[of "Param 0"] by blast + + hence "\P1 P2. \a. register_machine.rm_equations p (Suc n) a + = (\v. ppeval P1 (\x. a) v = ppeval P2 (\x. a) v)" + by auto + + thus ?thesis + unfolding A_characterization is_dioph_set_def by simp +qed + +end diff --git a/thys/DPRM_Theorem/Diophantine/Alpha_Sequence.thy b/thys/DPRM_Theorem/Diophantine/Alpha_Sequence.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Diophantine/Alpha_Sequence.thy @@ -0,0 +1,80 @@ +subsection \Diophantine description of alpha function\ + +theory Alpha_Sequence + imports Modulo_Divisibility "Exponentiation" +begin + +text \The alpha function is diophantine\ +definition alpha ("[_ = \ _ _]" 1000) + where "[X = \ B N] \ (TERNARY (\b n x. b > 3 \ x = Exp_Matrices.\ b n) B N X)" + +lemma alpha_dioph[dioph]: + fixes B N X + defines "D \ [X = \ B N]" + shows "is_dioph_rel D" +proof - + define r s t u v w x y where param_defs: + "r == (Param 0)" "s == (Param 1)" "t == (Param 2)" "u == (Param 3)" "v == (Param 4)" + "w == (Param 5)" "x == (Param 6)" "y == (Param 7)" + define B' X' N' where pushed_defs: "B' == (push_param B 8)" "X' == (push_param X 8)" + "N' == (push_param N 8)" + + (* eqns. (3.41) – (3.55) from YuMat lecture notes (p. 36) where "a" = X and "c" = N here *) + (* don't confuse capital X with x == (param 6) here, this is unhappy notation *) + define DR1 where "DR1 \ B' [>] (Const 3) [\] (Const 1 [+] B' [*] u [*] t [=] u[^2] [+] t[^2])" + define DR2 where "DR2 \ (Const 1 [+] B' [*] s [*] r [=] s[^2] [+] r[^2]) [\] r [<] s" + define DR3 where "DR3 \ (DVD (u[^2]) s) [\] (v [+] (Const 2) [*] r [=] B' [*] s)" + define DR4 where "DR4 \ (MOD B' v w) [\] (MOD (Const 2) u w) [\] (Const 2) [<] w" + define DR5 where "DR5 \ (Const 1 [+] w [*] x [*] y [=] x[^2] [+] y[^2])" + define DR6 where "DR6 \ (Const 2) [*] X' [<] u [\] (Const 2) [*] X' [<] v + [\] (MOD x v X') [\] (Const 2) [*] N' [<] u [\] MOD x u N'" + + define DR where "DR \ [\8] DR1 [\] DR2 [\] DR3 [\] DR4 [\] DR5 [\] DR6" + (* definition would otherwise be too long to handle for Isabelle *) + + note DR_defs = DR1_def DR2_def DR3_def DR4_def DR5_def DR6_def + + have "is_dioph_rel DR" + unfolding DR_def DR_defs + by (auto simp: dioph) + + moreover have "eval D a = eval DR a" for a + proof - + define x_ev b n where evaled_defs: "x_ev \ peval X a" "b \ peval B a" "n \ peval N a" + have h: "eval D a = (\r s t u v w x y::nat. Exp_Matrices.alpha_equations x_ev b n r s t u v w x y)" + unfolding D_def alpha_def evaled_defs defs using alpha_equivalence by simp + + show ?thesis + proof (rule) + assume "eval D a" + then obtain r s t u v w x y :: nat where "Exp_Matrices.alpha_equations x_ev b n r s t u v w x y" + using h by auto + then show "eval DR a" + unfolding evaled_defs Exp_Matrices.alpha_equations_def + unfolding DR_def DR_defs defs param_defs apply (auto simp: sq_p_eval) + apply (rule exI[of _ "[r, s, t, u, v, w, x, y]"]) + unfolding pushed_defs by (auto simp add: push_push[where ?n = 8] push_list_eval) + next + assume "eval DR a" + then show "eval D a" + (* simplify "eval DR a" and undo all the pushes *) + unfolding DR_def DR_defs defs param_defs apply (auto simp: sq_p_eval) + unfolding pushed_defs apply (auto simp add: push_push[where ?n = 8] push_list_eval) + (* simplify "eval D a" using the above fact h *) + unfolding h Exp_Matrices.alpha_equations_def evaled_defs + subgoal for ks + apply (rule exI[of _ "ks!0"]) apply (rule exI[of _ "ks!1"]) apply (rule exI[of _ "ks!2"]) + apply (rule exI[of _ "ks!3"]) apply (rule exI[of _ "ks!4"]) apply (rule exI[of _ "ks!5"]) + apply (rule exI[of _ "ks!6"]) apply (rule exI[of _ "ks!7"]) + by simp_all + done + qed + qed + + ultimately show ?thesis + by (auto simp: is_dioph_rel_def) +qed + +declare alpha_def[defs] + +end \ No newline at end of file diff --git a/thys/DPRM_Theorem/Diophantine/Assignments.thy b/thys/DPRM_Theorem/Diophantine/Assignments.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Diophantine/Assignments.thy @@ -0,0 +1,249 @@ +subsection \Variable Assignments\ + +text \The following theory defines manipulations of variable assignments and proves elementary facts about + these. Such preliminary results will later be necesary to e.g. prove that conjunction is + diophantine.\ + +theory Assignments + imports Parametric_Polynomials +begin + +definition shift :: "nat list \ nat \ assignment" where + "shift l a \ \i. l ! (i + a)" + +definition push :: "assignment \ nat \ assignment" where + "push a n i = (if i = 0 then n else a (i-1))" + +definition push_list :: "assignment \ nat list \ nat \ nat" where + "push_list a ns i = (if i < length ns then (ns!i) else a (i - length ns))" + +lemma push0: "push a n 0 = n" + by (auto simp: push_def) + +lemma push_list_empty: "push_list a [] = a" + unfolding push_list_def by auto + +lemma push_list_singleton: "push_list a [n] = push a n" + unfolding push_list_def push_def by auto + +lemma push_list_eval: "i < length ns \ push_list a ns i = ns!i" + unfolding push_list_def by auto + +lemma push_list1: "push (push_list a ns) n = push_list a (n # ns)" + unfolding push_def push_list_def by fastforce + +lemma push_list2_aux: "(push_list (push a n) ns) i = push_list a (ns @ [n]) i" + unfolding push_def push_list_def by (auto simp: nth_append) + +lemma push_list2: "(push_list (push a n) ns) = push_list a (ns @ [n])" + unfolding push_list2_aux by auto + +fun pull_param :: "ppolynomial \ ppolynomial \ ppolynomial" where + "pull_param (ppolynomial.Param 0) repl = repl" | + "pull_param (ppolynomial.Param (Suc n)) _ = (ppolynomial.Param n)" | + "pull_param (D1 \<^bold>+ D2) repl = (pull_param D1 repl) \<^bold>+ (pull_param D2 repl)" | + "pull_param (D1 \<^bold>- D2) repl = (pull_param D1 repl) \<^bold>- (pull_param D2 repl)" | + "pull_param (D1 \<^bold>* D2) repl = (pull_param D1 repl) \<^bold>* (pull_param D2 repl)" | + "pull_param P repl = P" + +(* FROM PREVIOUS FILE *) + +fun var_set :: "ppolynomial \ nat set" where + "var_set (ppolynomial.Const c) = {}" | + "var_set (ppolynomial.Param x) = {}" | + "var_set (ppolynomial.Var x) = {x}" | + "var_set (D1 \<^bold>+ D2) = var_set D1 \ var_set D2" | + "var_set (D1 \<^bold>- D2) = var_set D1 \ var_set D2" | + "var_set (D1 \<^bold>* D2) = var_set D1 \ var_set D2" + +definition disjoint_var :: "ppolynomial \ ppolynomial \ bool" where + "disjoint_var P Q = (var_set P \ var_set Q = {})" + +named_theorems disjoint_vars + +lemma disjoint_var_sym: "disjoint_var P Q = disjoint_var Q P" + unfolding disjoint_var_def by auto + +lemma disjoint_var_sum[disjoint_vars]: "disjoint_var (P1 \<^bold>+ P2) Q = (disjoint_var P1 Q \ disjoint_var P2 Q)" + unfolding disjoint_var_def by auto + +lemma disjoint_var_diff[disjoint_vars]: "disjoint_var (P1 \<^bold>- P2) Q = (disjoint_var P1 Q \ disjoint_var P2 Q)" + unfolding disjoint_var_def by auto + +lemma disjoint_var_prod[disjoint_vars]: "disjoint_var (P1 \<^bold>* P2) Q = (disjoint_var P1 Q \ disjoint_var P2 Q)" + unfolding disjoint_var_def by auto + +lemma aux_var_set: + assumes "\i \ var_set P. x i = y i" + shows "ppeval P a x = ppeval P a y" + using assms by (induction P, auto) + +text \First prove that disjoint variable sets allow the unification into one variable assignment\ +definition zip_assignments :: "ppolynomial \ ppolynomial \ assignment \ assignment \ assignment" + where "zip_assignments P Q v w i = (if i \ var_set P then v i else w i)" + +lemma help_eval_zip_assignments1: + shows "ppeval P1 a (\i. if i \ var_set P1 \ var_set P2 then v i else w i) + = ppeval P1 a (\i. if i \ var_set P1 then v i else w i)" + using aux_var_set by auto + +lemma help_eval_zip_assignments2: + shows "ppeval P2 a (\i. if i \ var_set P1 \ var_set P2 then v i else w i) + = ppeval P2 a (\i. if i \ var_set P2 then v i else w i)" + using aux_var_set by auto + +lemma eval_zip_assignments1: + fixes v w + assumes "disjoint_var P Q" + defines "x \ zip_assignments P Q v w" + shows "ppeval P a v = ppeval P a x" + using assms + apply (induction P arbitrary: x) + unfolding x_def zip_assignments_def + using help_eval_zip_assignments1 help_eval_zip_assignments2 + by (auto simp add: disjoint_vars) + +lemma eval_zip_assignments2: + fixes v w + assumes "disjoint_var P Q" + defines "x \ zip_assignments P Q v w" + shows "ppeval Q a w = ppeval Q a x" + using assms + apply (induction Q arbitrary: P x) + unfolding x_def zip_assignments_def + using disjoint_var_sym disjoint_vars + by (auto simp: disjoint_var_def) (smt (z3) inf_commute)+ + +lemma zip_assignments_correct: + assumes "ppeval P1 a v = ppeval P2 a v" and "ppeval Q1 a w = ppeval Q2 a w" + and "disjoint_var (P1 \<^bold>+ P2) (Q1 \<^bold>+ Q2)" + defines "x \ zip_assignments (P1 \<^bold>+ P2) (Q1 \<^bold>+ Q2) v w" + shows "ppeval P1 a x = ppeval P2 a x" and "ppeval Q1 a x = ppeval Q2 a x" +proof - + from assms(3) have "disjoint_var P1 (Q1 \<^bold>+ Q2)" + by (auto simp: disjoint_var_sum) + moreover have "ppeval P1 a x = ppeval P1 a (zip_assignments P1 (Q1 \<^bold>+ Q2) v w)" + unfolding x_def zip_assignments_def using help_eval_zip_assignments1 by auto + ultimately have p1: "ppeval P1 a x = ppeval P1 a v" + using eval_zip_assignments1[of "P1"] by auto + + from assms(3) have "disjoint_var P2 (Q1 \<^bold>+ Q2)" + by (auto simp: disjoint_var_sum) + moreover have "ppeval P2 a x = ppeval P2 a (zip_assignments P2 (Q1 \<^bold>+ Q2) v w)" + unfolding x_def zip_assignments_def using help_eval_zip_assignments2 by auto + ultimately have p2: "ppeval P2 a x = ppeval P2 a v" + using eval_zip_assignments1[of "P2"] by auto + + from p1 p2 show "ppeval P1 a x = ppeval P2 a x" + using assms(1) by auto +next + have "disjoint_var (P1 \<^bold>+ P2) Q1" + using assms(3) disjoint_var_sum disjoint_var_sym by auto + moreover have "ppeval Q1 a x = ppeval Q1 a (zip_assignments (P1 \<^bold>+ P2) Q1 v w)" + unfolding x_def zip_assignments_def using help_eval_zip_assignments1 by auto + ultimately have q1: "ppeval Q1 a x = ppeval Q1 a w" + using eval_zip_assignments2[of _ "Q1"] by auto + + from assms(3) have "disjoint_var (P1 \<^bold>+ P2) Q2" + using assms(3) disjoint_var_sum disjoint_var_sym by auto + moreover have "ppeval Q2 a x = ppeval Q2 a (zip_assignments (P1 \<^bold>+ P2) Q2 v w)" + unfolding x_def zip_assignments_def using help_eval_zip_assignments2 by auto + ultimately have q2: "ppeval Q2 a x = ppeval Q2 a w" + using eval_zip_assignments2[of _ "Q2"] by auto + + from q1 q2 show "ppeval Q1 a x = ppeval Q2 a x" + using assms(2) by auto +qed + +lemma disjoint_var_unifies: + assumes "\v1. ppeval P1 a v1 = ppeval P2 a v1" and "\v2. ppeval Q1 a v2 = ppeval Q2 a v2" + and "disjoint_var (P1 \<^bold>+ P2) (Q1 \<^bold>+ Q2)" + shows "\v. ppeval P1 a v = ppeval P2 a v \ ppeval Q1 a v = ppeval Q2 a v" + using assms zip_assignments_correct by (auto) metis + +(* Next prove that one can always find a polynomial with disjoint variables given some other polys*) +text \A function to manipulate variables in ppolynomials\ +fun push_var :: "ppolynomial \ nat \ ppolynomial" where + "push_var (ppolynomial.Var x) n = ppolynomial.Var (x + n)" | + "push_var (D1 \<^bold>+ D2) n = push_var D1 n \<^bold>+ push_var D2 n" | + "push_var (D1 \<^bold>- D2) n = push_var D1 n \<^bold>- push_var D2 n" | + "push_var (D1 \<^bold>* D2) n = push_var D1 n \<^bold>* push_var D2 n" | + "push_var D n = D" + +lemma push_var_bound: "x \ var_set (push_var P (Suc n)) \ x > n" + by (induction P, auto) + +definition pull_assignment :: "assignment \ nat \ assignment" where + "pull_assignment v n = (\x. v (x+n))" + +lemma push_var_pull_assignment: + shows "ppeval (push_var P n) a v = ppeval P a (pull_assignment v n)" + by (induction P, auto simp: pull_assignment_def) + +lemma max_set: "finite A \ \x\A. x \ Max A" + using Max_ge by blast + + +(* FROM PREVIOUS FILE SIMPLEDIOPHANTINE_NAT *) + +fun push_param :: "polynomial \ nat \ polynomial" where + "push_param (Const c) n = Const c" | + "push_param (Param x) n = Param (x + n)" | + "push_param (Sum D1 D2) n = Sum (push_param D1 n) (push_param D2 n)" | + "push_param (NatDiff D1 D2) n = NatDiff (push_param D1 n) (push_param D2 n)" | + "push_param (Prod D1 D2) n = Prod (push_param D1 n) (push_param D2 n)" + +definition push_param_list :: "polynomial list \ nat \ polynomial list" where + "push_param_list s k \ map (\x. push_param x k) s" + + +lemma push_param0: "push_param P 0 = P" + by (induction P, auto) + +lemma push_push_aux: "peval (push_param P (Suc m)) (push a n) = peval (push_param P m) a" + by (induction P, auto simp: push_def) + +lemma push_push: + shows "length ns = n \ peval (push_param P n) (push_list a ns) = peval P a" +proof (induction ns arbitrary: n) + case Nil + then show ?case by (auto simp: push_list_empty push_param0) +next + case (Cons n ns) + thus ?case + using push_push_aux[where ?a = "push_list a ns"] + by (auto simp add: length_Cons push_list1) +qed + +lemma push_push_simp: + shows "peval (push_param P (length ns)) (push_list a ns) = peval P a" +proof (induction ns) + case Nil + then show ?case by (auto simp: push_list_empty push_param0) +next + case (Cons n ns) + thus ?case + using push_push_aux[where ?a = "push_list a ns"] + by (auto simp add: length_Cons push_list1) +qed + + +lemma push_push1: "peval (push_param P 1) (push a k) = peval P a" + using push_push[where ?ns = "[k]"] by (auto simp: push_list_singleton) + +lemma push_push_map: "length ns = n \ + list_eval (map (\x. push_param x n) ls) (push_list a ns) = list_eval ls a" + unfolding list_eval_def apply (induction ls, simp) + apply (induction ns, auto) + apply (metis length_map list.size(3) nth_equalityI push_push) + by (metis length_Cons length_map map_nth push_push) + +lemma push_push_map_i: "length ns = n \ i < length ls \ + peval (map (\x. push_param x n) ls ! i) (push_list a ns) = list_eval ls a i" + unfolding list_eval_def by (auto simp: push_push_map push_push) + +lemma push_push_map1: "i < length ls \ + peval (map (\x. push_param x 1) ls ! i) (push a n) = list_eval ls a i" + unfolding list_eval_def using push_push1 by (auto) + +end \ No newline at end of file diff --git a/thys/DPRM_Theorem/Diophantine/Binary_And.thy b/thys/DPRM_Theorem/Diophantine/Binary_And.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Diophantine/Binary_And.thy @@ -0,0 +1,150 @@ +subsection \Binary and is Diophantine\ + +theory Binary_And + imports Binary_Masking Binary_Orthogonal +begin + +lemma lm0244: "(a && b) \ a" +proof (induct a b rule: bitAND_nat.induct) + case (1 uu) + then show ?case by simp +next + case (2 v n) + then show ?case + apply (auto simp add: mult.commute) + by (smt One_nat_def add_cancel_left_right even_succ_div_two masks.elims(3) mod_Suc_le_divisor + mod_by_Suc_0 mod_mod_trivial mod_mult_self4 mult_numeral_1_right mult_zero_right + nonzero_mult_div_cancel_left not_mod2_eq_Suc_0_eq_0 numeral_2_eq_2 numeral_One + odd_two_times_div_two_succ zero_neq_numeral) +qed + +lemma lm0245: "(a && b) \ b" + by (subst bitAND_commutes) (simp add: lm0244) + +lemma bitAND_lt_left: "m && n \ m" + using lm0244 masks_leq by auto +lemma bitAND_lt_right: "m && n \ n" + using lm0245 masks_leq by auto + +lemmas bitAND_lt = bitAND_lt_right bitAND_lt_left + +lemma auxm3_lm0246: + shows "bin_carry a b k = bin_carry a b k mod 2" + using bin_carry_bounded by auto + +lemma auxm2_lm0246: + assumes "(\r< n.(nth_bit a r + nth_bit b r \ 1))" + shows "(nth_bit (a+b) n) = (nth_bit a n + nth_bit b n) mod 2" + using assms no_carry by auto + +lemma auxm1_lm0246: "a \ (a+b) \ (\n. nth_bit a n + nth_bit b n \ 1)" (is "?P \ ?Q") +proof- +{ + assume asm: "\?Q" + then obtain n where n1:"\(nth_bit a n + nth_bit b n \ 1)" + and n2:"\r < n. (nth_bit a r + nth_bit b r \ 1)" + using obtain_smallest by (auto dest: obtain_smallest) + hence ab1: "nth_bit a n =1 \ nth_bit b n = 1" using nth_bit_def by auto + hence "nth_bit (a+b) n = 0" using n2 auxm2_lm0246 by auto + hence "\?P" using masks_leq_equiv ab1 by auto (metis One_nat_def not_one_le_zero) +} then show "?P \ ?Q" by auto +qed + +lemma aux0_lm0246:"a \ (a+b) \(a+b)\ n = a \ n + b \ n" +proof- + show ?thesis using auxm1_lm0246 auxm2_lm0246 less_Suc_eq_le numeral_2_eq_2 by auto +qed + +lemma aux1_lm0246:"a\b \ (\n. nth_bit (b-a) n = nth_bit b n - nth_bit a n)" + using aux0_lm0246 masks_leq by auto (metis add_diff_cancel_left' le_add_diff_inverse) + +lemma lm0246:"(a - (a && b)) \ (b - (a && b))" + apply (subst ortho_mult_equiv) + apply (rule allI) subgoal for k + proof(cases "nth_bit a k = 0") + case True + have "nth_bit (a- (a && b)) k = 0" by (auto simp add: lm0244 aux1_lm0246 "True") + then show ?thesis by simp + next + case False + then show ?thesis proof(cases "nth_bit b k = 0") + case True + have "nth_bit (b- (a && b)) k = 0" by (auto simp add: lm0245 aux1_lm0246 "True") + then show ?thesis by simp + next + case False2: False + have "nth_bit a k = 1" using False nth_bit_def by auto + moreover have "nth_bit b k = 1" using False2 nth_bit_def by auto + ultimately have "nth_bit (b- (a && b)) k = 0" + by (auto simp add: lm0245 aux1_lm0246 bitAND_digit_mult) + then show ?thesis by simp + qed + qed +done + +lemma aux0_lm0247:"(nth_bit a k) * (nth_bit b k) \ 1" + using eq_iff nth_bit_def by fastforce + +lemma lm0247_masking_equiv: + fixes a b c :: nat + shows "(c = a && b) \ (c \ a \ c \ b \ (a - c) \ (b - c))" (is "?P \ ?Q") +proof (rule) + assume ?P + thus ?Q + apply (auto simp add: lm0244 lm0245) + using lm0246 orthogonal.simps by blast +next + assume Q: ?Q + have "(\k. nth_bit c k \ nth_bit a k \ nth_bit c k \ nth_bit b k)" + using Q masks_leq_equiv by auto + moreover have "(\k x. nth_bit x k \ 1)" + by (auto simp add: nth_bit_def) + ultimately have f0:"(\k. nth_bit c k \ ((nth_bit a k) * (nth_bit b k)))" + by (metis mult.right_neutral mult_0_right not_mod_2_eq_0_eq_1 nth_bit_def) + show "?Q \ ?P" + proof (rule ccontr) + assume contr:"c \ a && b" + have k_exists:"(\k. (nth_bit c k) < ((nth_bit a k) * (nth_bit b k)))" + using bitAND_mult_equiv by (meson f0 contr le_less) + then obtain k + where "(nth_bit c k) < ((nth_bit a k) * (nth_bit b k))" .. + hence abc_kth:"((nth_bit c k) = 0) \ ((nth_bit a k) = 1) \ ((nth_bit b k) = 1)" + using aux0_lm0247 less_le_trans + by (metis One_nat_def Suc_leI nth_bit_bounded less_le less_one one_le_mult_iff) + hence "(nth_bit (a - c) k) = 1 \ (nth_bit (b - c) k) = 1" + by (auto simp add: abc_kth aux1_lm0246 Q) + hence "\ ((a - c) \ (b - c))" + by (metis mult.left_neutral not_mod_2_eq_0_eq_1 ortho_mult_equiv) + then show False + using Q by blast + qed +qed + +definition binary_and ("[_ = _ && _]" 1000) + where "[A = B && C] \ (TERNARY (\a b c. a = b && c) A B C)" + +lemma binary_and_dioph[dioph]: + fixes A B C :: polynomial + defines "DR \ [A = B && C]" + shows "is_dioph_rel DR" +proof - + define DS where "DS \ (A [\] B) [\] (A [\] C) [\] (B [-] A) [\] (C [-] A)" + + have "eval DS a = eval DR a" for a + unfolding DS_def DR_def binary_and_def defs + by (simp add: lm0247_masking_equiv) + + moreover have "is_dioph_rel DS" + unfolding DS_def by (auto simp: dioph) + + ultimately show ?thesis + by (auto simp: is_dioph_rel_def) +qed + +declare binary_and_def[defs] + + +definition binary_and_attempt :: "polynomial \ polynomial \ polynomial" ("_ &? _") where + "A &? B \ Const 0" + +end \ No newline at end of file diff --git a/thys/DPRM_Theorem/Diophantine/Binary_Masking.thy b/thys/DPRM_Theorem/Diophantine/Binary_Masking.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Diophantine/Binary_Masking.thy @@ -0,0 +1,79 @@ +subsection \Binary masking is Diophantine\ + +theory Binary_Masking + imports Binary_Orthogonal +begin + +lemma lm0243_masks_binom_equiv: "(b \ c) \ odd (c choose b)" (is "?P \ ?Q") +proof- + consider (lt) "b < c" | (eq) "b = c" | (gt) "b > c" using nat_neq_iff by blast + then show ?thesis + proof(cases) + case lt + hence "\a. c = a + b" using less_imp_add_positive semiring_normalization_rules(24) by blast + then obtain a where a_def:"c = a + b" .. + have "a \ b \ b \ a + b" (is "?P \ ?Q") + proof + assume ?P + then show ?Q + using ortho_mult_equiv no_carry_mult_equiv masks_leq_equiv[of b "a+b"] + sum_digit_formula nth_bit_bounded + by auto (metis add.commute add.right_neutral lessI less_one mod_less + nat_less_le one_add_one plus_1_eq_Suc zero_le) + next + assume ?Q + have "?Q \ \k. a \ k + b \ k \ 1" + proof(rule ccontr) + assume "\(\k. a \ k + b \ k \ 1) " + then obtain k where k1:"\(a \ k + b \ k \ 1)" and k2:"\r r + b \ r \ 1" + by (auto dest: obtain_smallest) + have c1: "bin_carry a b k = 1" + using `?Q` masks_leq_equiv sum_digit_formula carry_bounded nth_bit_bounded k1 + by (metis add.commute add.left_neutral add_self_mod_2 less_one nat_less_le not_le) + then show False using carry_digit_impl[of a b k] k2 by auto + qed + then show ?P + using `?Q` ortho_mult_equiv no_carry_mult_equiv masks_leq_equiv[of b "a+b"] + sum_digit_formula nth_bit_bounded + by auto (metis add_le_same_cancel2 le_0_eq le_SucE) + qed + then show ?thesis using lm0241_ortho_binom_equiv a_def by auto + next + case eq + hence "odd (c choose b)" by simp + moreover have "b \ c" using digit_wise_equiv masks_leq_equiv eq by blast + ultimately show ?thesis by simp + next + case gt + hence "\ odd (c choose b)" by (simp add: binomial_eq_0) + moreover have "\ (b \ c)" using masks_leq_equiv masks_leq gt not_le by blast + ultimately show ?thesis by simp + qed +qed + +definition masking ("_ [\] _" 60) + where "P [\] Q \ (BINARY (\a b. a \ b) P Q)" + +lemma masking_dioph[dioph]: + fixes P Q + defines "DR \ (P [\] Q)" + shows "is_dioph_rel DR" +proof - + define P' Q' where pushed_def: "P' \ push_param P 1" "Q' \ push_param Q 1" + + define DS where "DS \ [\] [Param 0 = Q' choose P'] [\] ODD Param 0" + + have "eval DS a = eval DR a" for a + unfolding DS_def DR_def defs pushed_def masking_def + using push_push1 by (simp add: push0 lm0243_masks_binom_equiv) + + moreover have "is_dioph_rel DS" + unfolding DS_def by (simp add: dioph) + + ultimately show ?thesis + by (auto simp: is_dioph_rel_def) +qed + +declare masking_def[defs] + +end \ No newline at end of file diff --git a/thys/DPRM_Theorem/Diophantine/Binary_Orthogonal.thy b/thys/DPRM_Theorem/Diophantine/Binary_Orthogonal.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Diophantine/Binary_Orthogonal.thy @@ -0,0 +1,144 @@ +subsection \Binary orthogonality is Diophantine\ + +theory Binary_Orthogonal + imports Binomial_Coefficient "Digit_Expansions.Binary_Operations" "Lucas_Theorem.Lucas_Theorem" +begin + +lemma equiv_with_lucas: "nth_digit = Lucas_Theorem.nth_digit_general" + unfolding nth_digit_def Lucas_Theorem.nth_digit_general_def by simp + +lemma lm0241_ortho_binom_equiv:"(a \ b) \ odd ((a + b) choose b)" (is "?P \ ?Q") +proof + assume "?P" + hence dig0:"(\i. (nth_bit a i) * (nth_bit b i) = 0)" + using ortho_mult_equiv + by auto + hence "(\i. (nth_bit a i) * (nth_bit b i) \ 1)" + by presburger + hence dcons:"(\i. \(((nth_bit a i) = 1) \ ((nth_bit b i) = 1)))" + by auto + hence "(\i. (bin_carry a b i) = 0)" using no_carry_mult_equiv dig0 + by blast + hence dsum:"(\i. (nth_bit (a + b) i) = (nth_bit a i) + (nth_bit b i))" + by (metis One_nat_def add.commute add_cancel_right_left add_self_mod_2 + dig0 mult_is_0 not_mod2_eq_Suc_0_eq_0 nth_bit_def one_mod_two_eq_one + sum_digit_formula) + + have bdd_ab_exists:"(\p. (a + b) < 2^(Suc p))" + using aux1_lm0241_pow2_up_bound by auto + then obtain p where bdd_ab:"(a + b) < 2^(Suc p)" by blast + moreover from bdd_ab have "b < 2^(Suc p)" by auto + + ultimately have "((a + b) choose b) mod 2 = + (\i\p. ((nth_digit (a + b) i 2) choose (nth_digit b i 2))) mod 2" + using lucas_theorem[of "a+b" 2 p b] bdd_ab two_is_prime_nat + by (simp add: equiv_with_lucas) + + then have a_choose_b_digit_prod: "((a + b) choose b) mod 2 = + (\i\p. ((nth_bit (a + b) i) choose (nth_bit b i))) mod 2" + using nth_digit_base2_equiv + by (auto cong: prod.cong) + + have "(\i. ((nth_bit (a + b) i) choose (nth_bit b i) = 1))" + using aux2_lm0241_single_digit_binom[where ?a="nth_bit (a + b) i" + and ?b="nth_bit b i"] + by (metis add.commute add.right_neutral binomial_n_0 binomial_n_n dig0 dsum + mult_is_0) + hence f0:"1 = (\i\i. (a + b) \ i choose b \ i = 1\) + then show "?Q" using f0 by fastforce +next + assume "?Q" + hence a_choose_b_odd:"((a + b) choose b) mod 2 = 1" + using odd_iff_mod_2_eq_one by blast + + have bdd_ab_exists:"(\p. (a + b) < 2^(Suc p))" + using aux1_lm0241_pow2_up_bound by auto + then obtain p where bdd_ab:"(a + b) < 2^(Suc p)" by blast + moreover from bdd_ab have bdd_b: "b < 2^(Suc p)" by auto + + ultimately have "((a + b) choose b) mod 2 = + (\i\p. ((nth_digit (a + b) i 2) choose (nth_digit b i 2))) mod 2" + using lucas_theorem[of "a+b" 2 p b] bdd_ab two_is_prime_nat + by (simp add: equiv_with_lucas) + + then have a_choose_b_digit_prod: "((a + b) choose b) mod 2 = + (\i\p. ((nth_bit (a + b) i) choose (nth_bit b i))) mod 2" + using nth_digit_base2_equiv nth_digit_def + by (auto cong: prod.cong) + hence all_prod_one_mod2:"... = 1" using a_choose_b_odd by linarith + + have choose_bdd:"(\i. 1 \ (nth_bit (a + b) i) choose (nth_bit b i))" + using nth_bit_bounded + by (metis One_nat_def binomial_n_0 choose_one not_mod2_eq_Suc_0_eq_0 + nth_bit_def order_refl) + hence "1 \ (\i\p. ((nth_bit (a + b) i) choose (nth_bit b i)))" + using all_prod_one_mod2 by (meson prod_le_1 zero_le) + hence "(\i\p. ((nth_bit (a + b) i) choose (nth_bit b i))) mod 2 = + (\i\p. ((nth_bit (a + b) i) choose (nth_bit b i)))" + using all_prod_one_mod2 by linarith + hence "... = 1" + using all_prod_one_mod2 by linarith + hence sub_pq_one:"\i\p. (nth_bit (a + b) i) choose (nth_bit b i) = 1" + using + aux4_lm0241_prod_one[where ?n="p" and ?f="(\i. (nth_bit (a + b) i) choose (nth_bit b i))"] + choose_bdd by blast + + have "\r > p. (a + b) < 2^r" using bdd_ab + by(metis One_nat_def Suc_lessI lessI less_imp_add_positive numeral_2_eq_2 + power_strict_increasing_iff trans_less_add1) + hence "\r > p. r \ p \ (a + b) < 2^r" by auto + hence ab_digit_bdd:"\r > p. r \ p \ nth_bit (a + b) r = 0" + using nth_bit_def by simp + + have "\k > p. b < 2 ^ k" using bdd_b + by(metis One_nat_def Suc_lessI lessI less_imp_add_positive numeral_2_eq_2 + power_strict_increasing_iff trans_less_add1) + hence b_digit_bdd:"\k > p. k \ p \ nth_bit b k = 0" + using nth_bit_def + by (simp add: \\k>p. b < 2 ^ k\) + + from b_digit_bdd ab_digit_bdd aux3_lm0241_binom_bounds + have "\i. i \ p \ (nth_bit (a + b) i) choose (nth_bit b i) = 1" + by (simp add: le_less sub_pq_one) + + hence "\i. ((nth_bit (a + b) i) choose (nth_bit b i)) = 1" + using sub_pq_one not_less by (metis linear) + hence "\i. \(nth_bit a i = 1 \ nth_bit b i = 1)" using aux5_lm0241 by blast + hence "\i. ((nth_bit a i = 0 \ nth_bit b i = 1) \ + (nth_bit a i = 1 \ nth_bit b i = 0) \ + (nth_bit a i = 0 \ nth_bit b i = 0))" + by (auto simp add:nth_bit_def nth_digit_bounded; metis nat.simps(3)) + hence "\i. (nth_bit a i) * (nth_bit b i) = 0" using mult_is_0 by blast + then show "?P" using ortho_mult_equiv by blast +qed + +definition orthogonal (infix "[\]" 50) + where "P [\] Q \ (BINARY (\a b. a \ b) P Q)" + +lemma orthogonal_dioph[dioph]: + fixes P Q + defines "DR \ (P [\] Q)" + shows "is_dioph_rel DR" +proof - + define P' Q' where pushed_def: "P' \ push_param P 1" "Q' \ push_param Q 1" + + define DS where "DS \ [\] [Param 0 = (P' [+] Q') choose Q'] [\] ODD (Param 0)" + + have "eval DS a = eval DR a" for a + unfolding DS_def DR_def orthogonal_def pushed_def defs + using push_push1 lm0241_ortho_binom_equiv by (simp add: push0) + + moreover have "is_dioph_rel DS" + unfolding DS_def by (simp add: dioph) + + ultimately show ?thesis + by (auto simp: is_dioph_rel_def) +qed + +declare orthogonal_def[defs] + +end \ No newline at end of file diff --git a/thys/DPRM_Theorem/Diophantine/Binomial_Coefficient.thy b/thys/DPRM_Theorem/Diophantine/Binomial_Coefficient.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Diophantine/Binomial_Coefficient.thy @@ -0,0 +1,88 @@ +subsection \Binomial Coefficient is Diophantine\ + +theory Binomial_Coefficient + imports Digit_Function +begin + +lemma bin_coeff_diophantine: + shows "c = a choose b \ (\u.(u = 2^(Suc a) \ c = nth_digit ((u+1)^a) b u))" +proof- + have "(u + 1)^a = (\k\a. (a choose k) * u ^ k)" for u + using binomial[of u 1 a] by auto + moreover have "a choose k < 2 ^ Suc a" for k + using binomial_le_pow2[of a k] by (simp add: le_less_trans) + ultimately have "nth_digit (((2 ^ Suc a) + 1)^a) b (2 ^ Suc a) = a choose b" + using nth_digit_gen_power_series[of "\k.(a choose k)" a a b] by (simp add: atLeast0AtMost) + then show ?thesis by auto +qed + +definition binomial_coefficient ("[_ = _ choose _]" 1000) + where "[A = B choose C] \ (TERNARY (\a b c. a = b choose c) A B C)" + +lemma binomial_coefficient_dioph[dioph]: + fixes A B C :: polynomial + defines "DR \ [C = A choose B]" + shows "is_dioph_rel DR" +proof - + define A' B' C' where pushed_def: + "A' \ (push_param A 2)" "B' \ (push_param B 2)" "C' \ (push_param C 2)" + + (* Param 0 = u = 2^(a + 1), Param 1 = (u+1)^a *) + define DS where "DS \ [\2] [Param 0 = Const 2 ^ (A' [+] \<^bold>1)] + [\] [Param 1 = (Param 0 [+] \<^bold>1) ^ A'] + [\] [C' = Digit (Param 1) B' (Param 0)]" + + have "eval DS a = eval DR a" for a + proof - + have "eval DS a = (peval C a = nth_digit ((2 ^ Suc (peval A a) + 1)^ peval A a) + (peval B a) (2 ^ Suc (peval A a)))" + unfolding DS_def defs pushed_def apply (auto simp add: push_push) + apply (rule exI[of _ "[2 * 2 ^ peval A a, Suc (2 * 2 ^ peval A a) ^ peval A a]"]) + apply (auto simp add: push_push push_list_eval) + by (metis (mono_tags, lifting) Suc_lessI mult_pos_pos n_not_Suc_n + numeral_2_eq_2 one_eq_mult_iff pos2 zero_less_power) + + then show ?thesis + unfolding DR_def binomial_coefficient_def defs by (simp add: bin_coeff_diophantine) + qed + + moreover have "is_dioph_rel DS" + unfolding DS_def by (auto simp: dioph) + + ultimately show ?thesis + by (auto simp: is_dioph_rel_def) +qed + +declare binomial_coefficient_def[defs] + + +text \odd function is diophantine\ + +lemma odd_dioph_repr: + fixes a :: nat + shows "odd a \ (\x::nat. a = 2*x + 1)" + by (meson dvd_triv_left even_plus_one_iff oddE) + +definition odd_lift ("ODD _" [999] 1000) + where "ODD A \ (UNARY (odd) A)" + +lemma odd_dioph[dioph]: + fixes A + defines "DR \ (ODD A)" + shows "is_dioph_rel DR" +proof - + define DS where "DS \ [\] (push_param A 1) [=] Const 2 [*] Param 0 [+] Const 1" + + have "eval DS a = eval DR a" for a + unfolding DS_def DR_def odd_lift_def defs using push_push1 by (simp add: odd_dioph_repr push0) + + moreover have "is_dioph_rel DS" + unfolding DS_def by (auto simp: dioph) + + ultimately show ?thesis + by (auto simp: is_dioph_rel_def) +qed + +declare odd_lift_def[defs] + +end \ No newline at end of file diff --git a/thys/DPRM_Theorem/Diophantine/Digit_Function.thy b/thys/DPRM_Theorem/Diophantine/Digit_Function.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Diophantine/Digit_Function.thy @@ -0,0 +1,86 @@ +subsection \Digit function is Diophantine\ + +theory Digit_Function + imports Exponential_Relation "Digit_Expansions.Bits_Digits" +begin + +definition digit ("[ _ = Digit _ _ _]" [999] 1000) + where "[D = Digit AA K BASE] \ (QUATERNARY (\d a k b. b > 1 + \ d = nth_digit a k b) D AA K BASE)" +lemma mod_dioph2[dioph]: + fixes A B C + defines "D \ (MOD A B C)" + shows "is_dioph_rel D" +proof - + define A' B' C' where pushed_defs: "A' \ push_param A 2" "B' \ push_param B 2" "C' \ push_param C 2" + define DS where "DS \ [\2] (Param 0 [*] B' [+] C' [=] Param 1 [*] B' [+] A')" + + have "eval DS a = eval D a" for a + proof + show "eval DS a \ eval D a" + unfolding DS_def defs D_def mod_def + by auto (metis add.commute mod_mult_self1 push_push_simp pushed_defs(1) + pushed_defs(2) pushed_defs(3)) + show "eval D a \ eval DS a" + unfolding DS_def defs D_def mod_def + apply (auto simp: mod_repr) + subgoal for x y + apply (rule exI[of _ "[x, y]"]) + unfolding pushed_defs by (simp add: push_push[where ?n = 2] push_list_eval) + done + qed + + moreover have "is_dioph_rel DS" + unfolding DS_def by (simp add: dioph) + + ultimately show ?thesis + by (auto simp: is_dioph_rel_def) +qed + +lemma digit_dioph[dioph]: + fixes D A B K :: polynomial + defines "DR \ [D = Digit A K B]" + shows "is_dioph_rel DR" +proof - + define D' A' B' K' where pushed_defs: + "D' == (push_param D 4)" "A' == (push_param A 4)" "B' == (push_param B 4)" "K' == (push_param K 4)" + (* Param 2 = b^(k+1), Param 3 = b^k *) + define x y where param_defs: + "x == (Param 0)" "y == (Param 1)" + + define DS where "DS \ [\4] ( B' [>] Const 1 [\] + [(Param 2) = B' ^ (K' [+] Const 1)] [\] + [(Param 3) = B' ^ K'] [\] + A' [=] x [*] (Param 2) [+] D' [*] (Param 3) [+] y [\] + D' [<] B' [\] + y [<] (Param 3)) " + have "eval DS a = eval DR a" for a + proof + show "eval DS a \ eval DR a" + unfolding DS_def defs DR_def digit_def apply auto + unfolding pushed_defs push_push using pushed_defs push_push digit_gen_equiv by auto + + assume asm: "eval DR a" + then obtain x_val y_val where cond: "(peval A a) = x_val * (peval B a)^( (peval K a)+1) + + (peval D a) * (peval B a)^ (peval K a) + y_val + \ (peval D a) < (peval B a) + \ y_val < (peval B a)^ (peval K a)" + unfolding DS_def defs DR_def digit_def using digit_gen_equiv by auto metis + show "eval DS a" + using asm unfolding DS_def defs DR_def digit_def apply auto + apply (rule exI[of _ "[x_val, y_val, (peval B a) ^ ((peval K a) + 1), + (peval B a) ^ (peval K a)]"]) + unfolding pushed_defs using param_defs push_push push_list_def cond by auto+ + qed + + moreover have "is_dioph_rel DS" + unfolding DS_def by (simp add: dioph) + + ultimately show ?thesis + by (auto simp: is_dioph_rel_def) +qed + +declare digit_def[defs] + + +end \ No newline at end of file diff --git a/thys/DPRM_Theorem/Diophantine/Diophantine_Relations.thy b/thys/DPRM_Theorem/Diophantine/Diophantine_Relations.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Diophantine/Diophantine_Relations.thy @@ -0,0 +1,246 @@ +subsection \Diophantine Relations and Predicates\ + +theory Diophantine_Relations + imports Assignments +begin + +datatype relation = + NARY "nat list \ bool" "polynomial list" + | AND relation relation (infixl "[\]" 35) + | OR relation relation (infixl "[\]" 30) + | EXIST_LIST nat relation ("[\_] _" 10) + +fun eval :: "relation \ assignment \ bool" where + "eval (NARY R PL) a = R (map (\P. peval P a) PL)" + | "eval (AND D1 D2) a = (eval D1 a \ eval D2 a)" + | "eval (OR D1 D2) a = (eval D1 a \ eval D2 a)" + | "eval ([\n] D) a = (\ks::nat list. n = length ks \ eval D (push_list a ks))" + +definition is_dioph_rel :: "relation \ bool" where + "is_dioph_rel DR = (\P\<^sub>1 P\<^sub>2::ppolynomial. \a. (eval DR a) \ (\v. ppeval P\<^sub>1 a v = ppeval P\<^sub>2 a v))" + + +definition UNARY :: "(nat \ bool) \ polynomial \ relation" where + "UNARY R P = NARY (\l. R (l!0)) [P]" + +lemma unary_eval: "eval (UNARY R P) a = R (peval P a)" + unfolding UNARY_def by simp + +definition BINARY :: "(nat \ nat \ bool) \ polynomial \ polynomial \ relation" where + "BINARY R P\<^sub>1 P\<^sub>2 = NARY (\l. R (l!0) (l!1)) [P\<^sub>1, P\<^sub>2]" + +lemma binary_eval: "eval (BINARY R P\<^sub>1 P\<^sub>2) a = R (peval P\<^sub>1 a) (peval P\<^sub>2 a)" + unfolding BINARY_def by simp + +definition TERNARY :: "(nat \ nat \ nat \ bool) + \ polynomial \ polynomial \ polynomial \ relation" where + "TERNARY R P\<^sub>1 P\<^sub>2 P\<^sub>3 = NARY (\l. R (l!0) (l!1) (l!2)) [P\<^sub>1, P\<^sub>2, P\<^sub>3]" + +lemma ternary_eval: "eval (TERNARY R P\<^sub>1 P\<^sub>2 P\<^sub>3) a = R (peval P\<^sub>1 a) (peval P\<^sub>2 a) (peval P\<^sub>3 a)" + unfolding TERNARY_def by simp + +definition QUATERNARY :: "(nat \ nat \ nat \ nat \ bool) + \ polynomial \ polynomial \ polynomial \ polynomial \ relation" where + "QUATERNARY R P\<^sub>1 P\<^sub>2 P\<^sub>3 P\<^sub>4 = NARY (\l. R (l!0) (l!1) (l!2) (l!3)) [P\<^sub>1, P\<^sub>2, P\<^sub>3, P\<^sub>4]" + +definition EXIST :: "relation \ relation" ("[\] _" 10) where + "([\] D) = ([\1] D)" + +definition TRUE where "TRUE = UNARY ((=) 0) (Const 0)" + +text \Bounded constant all quantifier (i.e. recursive conjunction)\ +fun ALLC_LIST :: "nat list \ (nat \ relation) \ relation" ("[\ in _] _") where + "[\ in []] DF = TRUE" | + "[\ in (l # ls)] DF = (DF l [\] [\ in ls] DF)" + +lemma ALLC_LIST_eval_list_all: "eval ([\ in L] DF) a = list_all (\l. eval (DF l) a) L" + by (induction L, auto simp: TRUE_def UNARY_def) + +lemma ALLC_LIST_eval: "eval ([\ in L] DF) a = (\k (nat \ relation) \ relation" ("[\<_] _") where + "[\ [\ in [0..k 'a list" where + "concat [] = []" | + "concat (l # ls) = l @ concat ls" + +fun splits :: "'a list \ nat list \ 'a list list" where + "splits L [] = []" | + "splits L (n # ns) = (take n L) # (splits (drop n L) ns)" + +lemma split_concat: + "splits (map f (concat pls)) (map length pls) = map (map f) pls" + by (induction pls, auto) + +definition LARY :: "(nat list list \ bool) \ (polynomial list list) \ relation" where + "LARY R PLL = NARY (\l. R (splits l (map length PLL))) (concat PLL)" + +lemma LARY_eval: + fixes PLL :: "polynomial list list" + shows "eval (LARY R PLL) a = R (map (map (\P. peval P a)) PLL)" + unfolding LARY_def apply (induction PLL, simp) + subgoal for pl pls by (induction pl, auto simp: split_concat) + done + +lemma or_dioph: + assumes "is_dioph_rel A" and "is_dioph_rel B" + shows "is_dioph_rel (A [\] B)" +proof - + from assms obtain PA1 PA2 where PA: "\a. (eval A a) \ (\v. ppeval PA1 a v = ppeval PA2 a v)" + by (auto simp: is_dioph_rel_def) + from assms obtain PB1 PB2 where PB: "\a. (eval B a) \ (\v. ppeval PB1 a v = ppeval PB2 a v)" + by (auto simp: is_dioph_rel_def) + + (* OR means (A1 - A2) * (B1 - B2) = 0 + Rewrite as A1*B1 + A2*B2 = A1*B2 + A2*B1 to eliminate subtraction. *) + show ?thesis + unfolding is_dioph_rel_def + apply (rule exI[of _ "PA1 \<^bold>* PB1 \<^bold>+ PA2 \<^bold>* PB2"]) + apply (rule exI[of _ "PA1 \<^bold>* PB2 \<^bold>+ PA2 \<^bold>* PB1"]) + using PA PB by (auto) (metis crossproduct_eq add.commute)+ + (* For each subgoal, a different fact is unused. The warning is due to 'metis+' being used. *) +qed + +lemma exists_disjoint_vars: + fixes Q1 Q2 :: ppolynomial + fixes A :: relation + assumes "is_dioph_rel A" + shows "\P1 P2. disjoint_var (P1 \<^bold>+ P2) (Q1 \<^bold>+ Q2) + \ (\a. eval A a \ (\v. ppeval P1 a v = ppeval P2 a v))" +proof - + obtain P1 P2 where p_defs: "\a. eval A a \ (\v. ppeval P1 a v = ppeval P2 a v)" + using assms is_dioph_rel_def by auto + + define n::nat where "n \ Max (var_set (Q1 \<^bold>+ Q2))" + + define P1' P2' where p'_defs: "P1' \ push_var P1 (Suc n)" "P2' \ push_var P2 (Suc n)" + + have "disjoint_var (P1' \<^bold>+ P2') (Q1 \<^bold>+ Q2)" + proof - + have "finite (var_set (Q1 \<^bold>+ Q2))" + apply (induction Q1, auto) + by (induction Q2, auto)+ + + hence "\x \ var_set (Q1 \<^bold>+ Q2). x \ n" + unfolding n_def using Max.coboundedI by blast + + moreover have "\x \ var_set (P1' \<^bold>+ P2'). x > n" + unfolding p'_defs using push_var_bound by auto + + ultimately show ?thesis + unfolding disjoint_var_def by fastforce + qed + + moreover have "\a. eval A a \ (\v. ppeval P1' a v = ppeval P2' a v)" + unfolding p'_defs apply (auto simp add: p_defs push_var_pull_assignment pull_assignment_def) + subgoal for a v by (rule exI[of _ "\i. v (i - Suc n)"]) auto + done + + ultimately show ?thesis + by auto +qed + +(* Combine the two results to show that AND has a diophantine representation *) +lemma and_dioph: + assumes "is_dioph_rel A" and "is_dioph_rel B" + shows "is_dioph_rel (A [\] B)" +proof - + from assms(1) obtain PA1 PA2 where PA: "\a. (eval A a) \ (\v. ppeval PA1 a v = ppeval PA2 a v)" + by (auto simp: is_dioph_rel_def) + from assms(2) obtain PB1 PB2 where disj: "disjoint_var (PB1 \<^bold>+ PB2) (PA1 \<^bold>+ PA2)" + and PB: "(\a. eval B a \ (\v. ppeval PB1 a v = ppeval PB2 a v))" + using exists_disjoint_vars[of B] by blast + + from disjoint_var_unifies have unified: "\a. (eval (A [\] B) a) + \ (\v. ppeval PA1 a v = ppeval PA2 a v \ ppeval PB1 a v = ppeval PB2 a v)" + using PA PB disj disjoint_var_sym by simp blast + + (* AND means (A1 - A2)^2 + (B1 - B2)^2 = 0 + Rewrite as A1^2 + A2^2 + B1^2 + B2^2 = 2*A1*A2 + 2*B1*B2 to eliminate subtraction. *) + have h0: "p1 = p2 \ p1^2 + p2^2 = 2*p1*p2" for p1 p2 :: nat + apply (auto simp: algebra_simps power2_eq_square) + using crossproduct_eq by fastforce + + have "p1 = p2 \ q1 = q2 \ p1^2 + p2^2 + q1^2 + q2^2 = 2*p1*p2 + 2*q1*q2" for p1 p2 q1 q2 :: nat + proof (rule) + assume "p1 = p2 \ q1 = q2" + thus "p1\<^sup>2 + p2\<^sup>2 + q1\<^sup>2 + q2\<^sup>2 = 2 * p1 * p2 + 2 * q1 * q2" + by (auto simp: algebra_simps power2_eq_square) + next + assume "p1\<^sup>2 + p2\<^sup>2 + q1\<^sup>2 + q2\<^sup>2 = 2 * p1 * p2 + 2 * q1 * q2" + hence "(int p1)\<^sup>2 + (int p2)\<^sup>2 + (int q1)\<^sup>2 + (int q2)\<^sup>2 - 2 * int p1 * int p2 - 2 * int q1 * int q2 = 0" + by (auto) (smt (verit, best) mult_2 of_nat_add of_nat_mult power2_eq_square) + hence "(int p1 - int p2)^2 + (int q1 - int q2)^2 = 0" + by (simp add: power2_diff) + hence "int p1 = int p2" and "int q1 = int q2" + by (simp add: sum_power2_eq_zero_iff)+ + thus "p1 = p2 \ q1 = q2" + by auto + qed + + thus ?thesis + apply (simp only: is_dioph_rel_def) + apply (rule exI[of _ "PA1\<^bold>^\<^bold>2 \<^bold>+ PA2\<^bold>^\<^bold>2 \<^bold>+ PB1\<^bold>^\<^bold>2 \<^bold>+ PB2\<^bold>^\<^bold>2"]) + apply (rule exI[of _ "(ppolynomial.Const 2) \<^bold>* PA1 \<^bold>* PA2 \<^bold>+ (ppolynomial.Const 2) \<^bold>* PB1 \<^bold>* PB2"]) + apply (subst unified) + by (simp add: Sq_pp_def power2_eq_square) +qed + + +(* Some basic relations are diophantine *) +definition eq (infix "[=]" 50) where "eq Q R \ BINARY (=) Q R" +definition lt (infix "[<]" 50) where "lt Q R \ BINARY (<) Q R" +definition le (infix "[\]" 50) where "le Q R \ Q [<] R [\] Q [=] R" +definition gt (infix "[>]" 50) where "gt Q R \ R [<] Q" +definition ge (infix "[\]" 50) where "ge Q R \ Q [>] R [\] Q [=] R" + +named_theorems defs +lemmas [defs] = zero_p_def one_p_def eq_def lt_def le_def gt_def ge_def LARY_eval + UNARY_def BINARY_def TERNARY_def QUATERNARY_def ALLC_LIST_eval ALLC_eval + +named_theorems dioph +lemmas [dioph] = or_dioph and_dioph + +lemma true_dioph[dioph]: "is_dioph_rel TRUE" + unfolding TRUE_def UNARY_def is_dioph_rel_def by auto + +lemma eq_dioph[dioph]: "is_dioph_rel (Q [=] R)" + unfolding is_dioph_rel_def + apply (rule exI[of _ "convert Q"]) + apply (rule exI[of _ "convert R"]) + using convert_eval BINARY_def by (auto simp: eq_def) + +lemma lt_dioph[dioph]: "is_dioph_rel (Q [<] R)" + unfolding is_dioph_rel_def + apply (rule exI[of _ "(ppolynomial.Const 1) \<^bold>+ (ppolynomial.Var 0) \<^bold>+ convert Q"]) + apply (rule exI[of _ "convert R"]) + using convert_eval BINARY_def apply (auto simp: lt_def) + by (metis add.commute add.right_neutral less_natE) + +definition zero ("[0=] _" [60] 60) where[defs]: "zero Q \ \<^bold>0 [=] Q" +lemma zero_dioph[dioph]: "is_dioph_rel ([0=] Q)" + unfolding zero_def by (auto simp: eq_dioph) + +lemma gt_dioph[dioph]: "is_dioph_rel (Q [>] R)" + unfolding gt_def by (auto simp: lt_dioph) + +lemma le_dioph[dioph]: "is_dioph_rel (Q [\] R)" + unfolding le_def by (auto simp: lt_dioph eq_dioph or_dioph) + +lemma ge_dioph[dioph]: "is_dioph_rel (Q [\] R)" + unfolding ge_def by (auto simp: gt_dioph eq_dioph or_dioph) + +text \Bounded Constant All Quantifier, dioph rules\ + +lemma ALLC_LIST_dioph[dioph]: "list_all (is_dioph_rel \ DF) L \ is_dioph_rel ([\ in L] DF)" + by (induction L, auto simp add: dioph) + +lemma ALLC_dioph[dioph]: "\i is_dioph_rel ([\Existential quantification is Diophantine\ + +theory Existential_Quantifier + imports Diophantine_Relations +begin + +lemma exist_list_dioph[dioph]: + fixes D + assumes "is_dioph_rel D" + shows "is_dioph_rel ([\n] D)" +proof (induction n) + case 0 + then show ?case + using assms unfolding is_dioph_rel_def by (auto simp: push_list_empty) +next + case (Suc n) + + have h: "(\i. if i = 0 then v 0 else v i) = v" for v::assignment + by auto + + have "eval ([\Suc n] D) a = (\k::nat. eval ([\n] D) (push a k))" for a + apply (simp add: push_list2) + by (smt (z3) Zero_not_Suc add_Suc_right append_Nil2 length_Cons + length_append list.size(3) nat.inject rev_exhaust) + moreover from Suc is_dioph_rel_def obtain P\<^sub>1 P\<^sub>2 where + "\a. eval ([\n] D) a = (\v. ppeval P\<^sub>1 a v = ppeval P\<^sub>2 a v)" + by auto + ultimately have t1: "eval ([\Suc n] D) a = (\k::nat. (\v. ppeval P\<^sub>1 (push a k) v + = ppeval P\<^sub>2 (push a k) v))" for a + by simp + + define f :: "ppolynomial \ ppolynomial" where + "f \ \P. pull_param (push_var P 1) (Var 0)" + have "ppeval P (push a k) v = ppeval (f P) a (push v k)" for P a k v + apply (induction P, auto simp: push_def f_def) + by (metis (no_types, lifting) Suc_pred ppeval.simps(2) pull_param.simps(2)) + then have t2: "eval ([\Suc n] D) a = (\k::nat. (\v. ppeval (f P\<^sub>1) a (push v k) + = ppeval (f P\<^sub>2) a (push v k)))" for a + using t1 by auto + moreover have "(\k::nat. \v. ppeval P a (push v k) = ppeval Q a (push v k)) + \ (\v. ppeval P a v = ppeval Q a v)" for P Q a + unfolding push_def + apply auto + subgoal for v + apply (rule exI[of _ "v 0"]) + apply (rule exI[of _ "\i. v (i + 1)"]) + by (auto simp: h cong: if_cong) + done + ultimately have "eval ([\Suc n] D) a = (\v. ppeval (f P\<^sub>1) a v = ppeval (f P\<^sub>2) a v)" for a + by auto + + thus ?case + unfolding is_dioph_rel_def by auto +qed + +lemma exist_dioph[dioph]: + fixes D + assumes "is_dioph_rel D" + shows "is_dioph_rel ([\] D)" + unfolding EXIST_def using assms by (auto simp: exist_list_dioph) + +lemma exist_eval[defs]: + shows "eval ([\] D) a = (\k. eval D (push a k))" + unfolding EXIST_def apply (simp add: push_list_def) + by (metis length_Suc_conv list.exhaust list.size(3) nat.simps(3) push_list_singleton) + +end \ No newline at end of file diff --git a/thys/DPRM_Theorem/Diophantine/Exponential_Relation.thy b/thys/DPRM_Theorem/Diophantine/Exponential_Relation.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Diophantine/Exponential_Relation.thy @@ -0,0 +1,173 @@ +subsection \Exponentiation is a Diophantine Relation\ + +theory Exponential_Relation + imports Alpha_Sequence "Exponentiation" +begin + + +(* Exponentiation is diophantine *) +definition exp_equations :: "nat \ nat \ nat \ nat \ nat \ bool" where + "exp_equations p q r b m = (b = Exp_Matrices.\ (q + 4) (r + 1) + q * q + 2 \ + m + q^2 + 1 = b * q \ + p < m \ + (p + b * Exp_Matrices.\ b r) mod m = (q * Exp_Matrices.\ b r + + Exp_Matrices.\ b (r + 1)) mod m)" + +lemma exp_repr: + fixes p q r :: nat + shows "p = q^r \ ((q = 0 \ r = 0 \ p = 1) \ + (q = 0 \ 0 < r \ p = 0) \ + (q > 0 \ (\b m :: nat. exp_equations p q r b m)))" (is "?P \ ?Q") +proof + assume P: ?P + (* 3 disjunction elims applied to assumption; makes cases 1 and 2 trivial *) + consider (c1)"q = 0 \ r = 0 \ p = 1" | (c2) "q = 0 \ 0 < r \ p = 0" | (c3) "q > 0 \ + (\b m. b = Exp_Matrices.\ (q + 4) (r + 1) + q * q + 2 \ m = b * q - q * q - 1 \ + p < m \ p mod m = (q * Exp_Matrices.\ b r - (b * Exp_Matrices.\ b r - + Exp_Matrices.\ b (r + 1))) mod m)" using exp_alpha[of p q r] P by auto + then show ?Q using P + proof cases + case c1 + then show ?thesis by auto + next + case c2 + then show ?thesis by auto + next + case c3 + obtain b m where + b_def: "b = Exp_Matrices.\ (q + 4) (r + 1) + q * q + 2" and + "m = b * q - q * q - 1" and + "p < m" and + "int (p mod m) = (int q * Exp_Matrices.\ b r - (int b * Exp_Matrices.\ b r - + Exp_Matrices.\ b (r + 1))) mod int m" + using exp_alpha[of p q r] c3 by blast + then have "exp_equations p q r b m" unfolding exp_equations_def + apply(intro conjI, auto simp add: power2_eq_square) using mod_add_right_eq by smt + then show ?thesis using c3 by blast + qed +next + assume ?Q + then show ?P + proof (elim disjE) + (* eliminate disjs, first two trivial cases by auto *) + show "q = 0 \ r = 0 \ p = 1 \ p = q ^ r" by auto + show "q = 0 \ 0 < r \ p = 0 \ p = q ^ r" by auto + (* third, non-trivial case *) + assume prems: "0 < q \ (\b m. exp_equations p q r b m)" + obtain b m where cond: "exp_equations p q r b m" using prems by auto + (* show that the four relevant conditions in exp_alpha hold; first three are easy *) + hence "int b = Exp_Matrices.\ (q + 4) (r + 1) + int (q * q) + 2 \ + m = b * q - q * q - 1 \ p < m" + unfolding exp_equations_def power2_eq_square by auto + (* the last one requires some extra care; but it's mainly mod_diff_cong *) + moreover have "int (p mod m) = (int q * Exp_Matrices.\ b r - + (int b * Exp_Matrices.\ b r - Exp_Matrices.\ b (r + 1))) mod int m" + using cond unfolding exp_equations_def + using mod_diff_cong[of "(p + b * Exp_Matrices.\ b r)" m "(q * Exp_Matrices.\ b r + + Exp_Matrices.\ b (r + 1))" "b * Exp_Matrices.\ b r" "b * Exp_Matrices.\ b r"] + unfolding diff_diff_eq2 by auto + ultimately show "p = q ^ r" using prems exp_alpha by auto + qed +qed + +definition exp ("[_ = _ ^ _]" 1000) + where "[Q = R ^ S] \ (TERNARY (\a b c. a = b ^ c) Q R S)" + +lemma exp_dioph[dioph]: + fixes P Q R :: polynomial + defines "D \ [P = Q ^ R]" + shows "is_dioph_rel D" +proof - + define P' Q' R' where pushed_def: + "P' \ (push_param P 5)" "Q' \ (push_param Q 5)" "R' \ (push_param R 5)" + define b m a0 a1 a2 where params_def: "b = Param 0" "m = Param 1" "a0 = Param 2" + "a1 = Param 3" "a2 = Param 4" + + define S1 where "S1 \ [0=] Q [\] [0=] R [\] P [=] \<^bold>1 [\] + [0=] Q [\] (Const 0) [<] R [\] [0=] P" + define S2 where "S2 \ [a0 = \ (Q' [+] (Const 4)) (R' [+] \<^bold>1)] + [\] b [=] (a0 [+] Q'[^2] [+] Const 2)" + define S3 where "S3 \ (m [+] Q'[^2] [+] Const 1) [=] b [*] Q' + [\] P' [<] m" + define S4 where "S4 \ [a1 = \ b R'] + [\] [a2 = \ b (R' [+] \<^bold>1)] + [\] MOD (P' [+] b [*] a1) m (Q' [*] a1 [+] a2)" + + note S_defs = S1_def S2_def S3_def S4_def + + define S where "S \ S1 [\] (Q [>] Const 0) [\] ([\5] S2 [\] S3 [\] S4)" + + have "is_dioph_rel S" + unfolding S_def S_defs by (auto simp: dioph) + + moreover have "eval S a = eval D a" for a + proof - + define p q r where evaled_defs: "p = peval P a" "q = peval Q a" "r = peval R a" + + show ?thesis + proof (rule) + assume "eval S a" + then show "eval D a" + unfolding S_def S_defs defs apply (simp add: sq_p_eval) + unfolding D_def exp_def defs apply simp_all + unfolding pushed_def params_def apply (auto simp add: push_push[where ?n = 5] push_list_eval) + unfolding exp_repr exp_equations_def apply simp + subgoal for ks + apply (rule exI[of _ "ks!0"], auto) + subgoal by (simp add: power2_eq_square) + subgoal apply (rule exI[of _ "ks!1"], auto) + by (smt int_ops(7) mult_Suc of_nat_Suc of_nat_add power2_eq_square zmod_int) + done + done + next + assume "eval D a" + then obtain b_val m_val where cond: "(q = 0 \ r = 0 \ p = 1) \ + (q = 0 \ 0 < r \ p = 0) \ + (q > 0 \ exp_equations p q r b_val m_val)" + unfolding D_def exp_def exp_repr evaled_defs ternary_eval by auto + + moreover define a0_val a1_val a2_val where + "a0_val \ nat (Exp_Matrices.\ (q + 4) (r + 1))" + "a1_val \ nat (Exp_Matrices.\ b_val r)" + "a2_val \ nat (Exp_Matrices.\ b_val (r + 1))" + ultimately show "eval S a" + unfolding S_def S_defs defs evaled_defs apply (simp add: sq_p_eval) + apply (elim disjE) + subgoal unfolding defs by simp + subgoal unfolding defs by simp + subgoal apply(elim conjE) apply(intro disjI2, intro conjI) + subgoal by simp + subgoal premises prems + proof- + have bg3: "3 < b_val" + proof- + have "b_val = Exp_Matrices.\ (q + 4) (r + 1) + int q * int q + 2" + using cond prems(4) evaled_defs(2) unfolding exp_equations_def by linarith + moreover have "int q * int q > 0" using evaled_defs(2) prems by simp + moreover have "Exp_Matrices.\ (q + 4) (r + 1) > 0" + using Exp_Matrices.alpha_superlinear[of "q+4" "r+1"] by linarith + ultimately show ?thesis by linarith + qed + show ?thesis apply (rule exI[of _ "[b_val, m_val, a0_val, a1_val, a2_val]"], intro conjI) + using prems + unfolding exp_equations_def pushed_def params_def + using push_list_def push_push bg3 Exp_Matrices.alpha_nonnegative apply simp_all + subgoal using push_list_def by (smt Exp_Matrices.alpha_strictly_increasing int_nat_eq + nat_int numeral_Bit0 numeral_One of_nat_1 of_nat_add of_nat_power plus_1_eq_Suc + power2_eq_square) + subgoal using push_list_def apply auto by (smt One_nat_def Suc_1 Suc_less_eq + int_nat_eq less_Suc_eq nat_int numeral_3_eq_3 of_nat_add of_nat_mult zmod_int) + done + qed + done + done + qed + qed + + ultimately show ?thesis + by (auto simp: is_dioph_rel_def) +qed + +declare exp_def[defs] + +end \ No newline at end of file diff --git a/thys/DPRM_Theorem/Diophantine/Exponentiation.thy b/thys/DPRM_Theorem/Diophantine/Exponentiation.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Diophantine/Exponentiation.thy @@ -0,0 +1,1692 @@ +section \Exponentiation is Diophaninte\ + +subsection \Expressing Exponentiation in terms of the alpha function\ + +theory Exponentiation + imports "HOL-Library.Discrete" +begin + +locale Exp_Matrices + begin + +subsubsection \2x2 matrices and operations\ +datatype mat2 = mat (mat_11 : int) (mat_12 : int) (mat_21 : int) (mat_22 : int) +datatype vec2 = vec (vec_1: int) (vec_2: int) + +fun mat_plus:: "mat2 \ mat2 \ mat2" where + "mat_plus A B = mat (mat_11 A + mat_11 B) (mat_12 A + mat_12 B) + (mat_21 A + mat_21 B) (mat_22 A + mat_22 B)" + +fun mat_mul:: "mat2 \ mat2 \ mat2" where + "mat_mul A B = mat (mat_11 A * mat_11 B + mat_12 A * mat_21 B) + (mat_11 A * mat_12 B + mat_12 A * mat_22 B) + (mat_21 A * mat_11 B + mat_22 A * mat_21 B) + (mat_21 A * mat_12 B + mat_22 A * mat_22 B)" + +fun mat_pow:: "nat \ mat2 \ mat2" where + "mat_pow 0 _ = mat 1 0 0 1" | + "mat_pow n A = mat_mul A (mat_pow (n - 1) A)" + +lemma mat_pow_2[simp]: "mat_pow 2 A = mat_mul A A" + by (simp add: numeral_2_eq_2) + +fun mat_det::"mat2 \ int" where + "mat_det M = mat_11 M * mat_22 M - mat_12 M * mat_21 M" + +fun mat_scalar_mult::"int \ mat2 \ mat2" where + "mat_scalar_mult a M = mat (a * mat_11 M) (a * mat_12 M) (a * mat_21 M) (a * mat_22 M)" + +fun mat_minus:: "mat2 \ mat2 \ mat2" where + "mat_minus A B = mat (mat_11 A - mat_11 B) (mat_12 A - mat_12 B) + (mat_21 A - mat_21 B) (mat_22 A - mat_22 B)" + +fun mat_vec_mult:: "mat2 \ vec2 \ vec2" where + "mat_vec_mult M v = vec (mat_11 M * vec_1 v + mat_12 M * vec_2 v) + (mat_21 M * vec_1 v + mat_21 M * vec_2 v)" + +definition ID :: "mat2" where "ID = mat 1 0 0 1" +declare mat_det.simps[simp del] + +subsubsection \Properties of 2x2 matrices\ +lemma mat_neutral_element: "mat_mul ID N = N" by (auto simp: ID_def) + +lemma mat_associativity: "mat_mul (mat_mul D B) C = mat_mul D (mat_mul B C)" + apply auto by algebra+ + +lemma mat_exp_law: "mat_mul (mat_pow n M) (mat_pow m M) = mat_pow (n+m) M" + apply (induction n, auto) by (metis mat2.sel(1,2) mat_associativity mat_mul.simps)+ + +lemma mat_exp_law_mult: "mat_pow (n*m) M = mat_pow n (mat_pow m M)" (is "?P n") + apply (induction n, auto) using mat_exp_law by (metis mat_mul.simps) + +lemma det_mult: "mat_det (mat_mul M1 M2) = (mat_det M1) * (mat_det M2)" + by (auto simp: mat_det.simps algebra_simps) + +subsubsection \Special second-order recurrent sequences\ +text \Equation 3.2\ +fun \:: "nat \ nat \ int" where + "\ b 0 = 0" | + "\ b (Suc 0) = 1" | + alpha_n: "\ b (Suc (Suc n)) = (int b) * (\ b (Suc n)) - (\ b n)" + +text \Equation 3.3\ +lemma alpha_strictly_increasing: + shows "int b \ 2 \ \ b n < \ b (Suc n) \ 0 < \ b (Suc n)" +proof (induct n) + case 0 + show ?case by simp +next + case (Suc n) + have pos: "0 < \ b (Suc n)" by (simp add:Suc) + have "\ b (Suc n) \ (int b) * (\ b (Suc n)) - \ b (Suc n)" using pos Suc by simp + also have "... < \ b (Suc (Suc n))" by (simp add: Suc) + finally show ?case using pos Suc by simp +qed + +lemma alpha_strictly_increasing_general: + fixes b n m::nat + assumes "b > 2 \ m > n" + shows "\ b m > \ b n" +proof - + from alpha_strictly_increasing assms have S2: "\ b n < \ b m" + by (smt less_imp_of_nat_less lift_Suc_mono_less of_nat_0_less_iff pos2) + show ?thesis using S2 by simp +qed + + +text \Equation 3.4\ +lemma alpha_superlinear: "b > 2 \ int n \ \ b n" + apply (induction n, auto) + by (smt Suc_1 alpha_strictly_increasing less_imp_of_nat_less of_nat_1 of_nat_Suc) + +text \A simple consequence that's often useful; could also be generalized to alpha using + alpha linear\ + +lemma alpha_nonnegative: + shows "b > 2 \ \ b n \ 0" + using of_nat_0_le_iff alpha_superlinear order_trans by blast + +text \Equation 3.5\ +lemma alpha_linear: "\ 2 n = n" +proof(induct n rule: nat_less_induct) + case (1 n) + have s0: "n=0 \ \ 2 n = n" by simp + have s1: "n=1 \ \ 2 n = n" by simp + note hyp = \\m < n. \ 2 m = m\ + from hyp have s2: "n>1 \ \ 2 (n-1) = n-1 \ \ 2 (n-2) = n-2" by simp + have s3: "n>1 \ \ 2 (Suc (Suc (n-2))) = 2*\ 2 (Suc (n-2)) - \ 2 (n-2)" by simp + have s4: "n>1 \ Suc (Suc (n-2)) = n" by simp + have s5: "n>1 \ Suc (n-2) = n-1" by simp + from s3 s4 s5 have s6: "n>1 \ \ 2 n = 2*\ 2 (n-1) - \ 2 (n-2)" by simp + from s2 s6 have s7: "n>1 \ \ 2 n = 2*(n-1) - (n-2)" by simp + from s7 have s8: "n>1 \ \ 2 n = n" by simp + from s0 s1 s8 show ?case by linarith +qed + +text \Equation 3.6 (modified)\ +lemma alpha_exponential_1: "b > 0 \ int b ^ n \ \ (b + 1) (n+1)" +proof(induction n) +case 0 + thus ?case by(simp) +next + case (Suc n) + hence "((int b)*(int b)^n) \ (int b)*(\ (b+1) (n+1))" by simp + hence r2: "((int b)^(Suc n)) \ (int (b+1))*(\ (b+1) (n+1)) - (\ (b+1) (n+1))" + by (simp add: algebra_simps) + have "(int b+1) *(\ (b+1) (n+1)) - (\ (b+1) (n+1)) \ (int b+1)*(\ (b+1) (n+1)) - \ (b+1) n" + using alpha_strictly_increasing Suc by (smt Suc_eq_plus1 of_nat_0_less_iff of_nat_Suc) + thus ?case using r2 by auto +qed + +lemma alpha_exponential_2: "int b>2 \ \ b (n+1) \ (int b)^(n)" +proof(induction n) + case 0 + thus ?case by simp +next + case (Suc n) + hence s1: "\ b (n+2) \ (int b)^(n+1) - \ b n" by simp + have "(int b)^(n+1) - (\ b n) \ (int b)^(n+1)" + using alpha_strictly_increasing Suc by (smt \.simps(1) alpha_superlinear of_nat_1 of_nat_add + of_nat_le_0_iff of_nat_less_iff one_add_one) + thus ?case using s1 by simp +qed + +subsubsection \First order relation\ +text \Equation 3.7 - Definition of A\ +fun A :: "nat \ nat \ mat2" where + "A b 0 = mat 1 0 0 1" | + A_n: "A b n = mat (\ b (n + 1)) (-(\ b n)) (\ b n) (-(\ b (n - 1)))" + +text \Equation 3.9 - Definition of B\ +fun B :: "nat \ mat2" where + "B b = mat (int b) (-1) 1 0" + +declare A.simps[simp del] +declare B.simps[simp del] + +text \Equation 3.8\ +lemma A_rec: "b>2 \ A b (Suc n) = mat_mul (A b n) (B b)" + by (induction n, auto simp: A.simps B.simps) + +text \Equation 3.10\ +lemma A_pow: "b>2 \ A b n = mat_pow n (B b)" + apply (induction n, auto simp: A.simps B.simps) + subgoal by (smt A.elims Suc_eq_plus1 \.simps \.simps(2) mat2.sel) + subgoal for n apply (cases "n=0", auto) + using A.simps(2)[of b "n-1"] gr0_conv_Suc mult.commute by auto + subgoal by (metis A.simps(2) Suc_eq_plus1 \.simps(2) mat2.sel(1) mat_pow.elims) + subgoal by (metis A.simps(2) \.simps(1) add.inverse_neutral mat2.sel(2) mat_pow.elims) + done + + +subsubsection \Characteristic equation\ +text \Equation 3.11\ +lemma A_det: "b>2 \ mat_det (A b n) = 1" + apply (auto simp: A_pow, induction n, simp add: mat_det.simps) + using det_mult apply (auto simp del: mat_mul.simps) by (simp add: B.simps mat_det.simps) + +text \Equation 3.12\ +lemma alpha_det1: + assumes "b>2" + shows "(\ b (Suc n))^2 - (int b) * \ b (Suc n) * \ b n + (\ b n)^2 = 1" +proof(cases "n = 0") + case True + thus ?thesis by auto +next + case False + hence "A b n = mat (\ b (n + 1)) (-(\ b n)) (\ b n) (-(\ b (n - 1)))" using A.elims neq0_conv by blast + hence "mat_det (A b n) = (\ b n)^2 - (\ b (Suc n)) * \ b (n-1)" + apply (auto simp: mat_det.simps) by (simp add: power2_eq_square) + moreover hence "... = (\ b (Suc n))^2 - b * (\ b (Suc n)) * \ b n + (\ b n)^2" + using False alpha_n[of b "n-1"] apply(auto simp add: algebra_simps) + by (metis Suc_1 distrib_left mult.commute mult.left_commute power_Suc power_one_right) + ultimately show ?thesis using A_det assms by auto +qed + +text \Equation 3.12\ +lemma alpha_det2: + assumes "b>2" "n>0" + shows "(\ b (n-1))^2 - (int b) * (\ b (n-1) * (\ b n)) + (\ b n)^2 = 1" + using alpha_det1 assms by (smt One_nat_def Suc_diff_Suc diff_zero mult.commute mult.left_commute) + +text \Equations 3.14 to 3.17\ +lemma alpha_char_eq: + fixes x y b:: nat + shows "(y < x \ x * x + y * y = 1 + b * x * y) \ (\m. int y = \ b m \ int x = \ b (Suc m))" +proof (induct y arbitrary: x rule:nat_less_induct) + case (1 n) + + note pre = \n < x \ (x * x + n * n = 1 + b * x * n)\ + + have h0: "int (x * x + n * n) = int (x * x) + int (n * n)" by simp + from pre h0 have pre1: "int x * int x + int(n * n) = int 1 + int(b * x * n)" by simp + + have i0: "int (n * n) = int n * int n" by simp + have i1: "int (b * x * n) = int b * int x * int n" by simp + from pre1 i0 i1 have pre2: "int x * int x + int n * int n = 1 + int b * int x * int n" by simp + + from pre2 have j0: "int n * int n - 1 = int b * int x * int n - int x * int x" by simp + have j1:"\ = int x * (int b * int n - int x)" by (simp add: right_diff_distrib) + from j0 j1 have pre3:"int n * int n - 1 = int x * (int b * int n - int x)" by simp + + have k0: "int n * int n - 1 < int n * int n" by simp + from pre3 k0 have k1:"int n * int n > int x * (int b * int n - int x)" by simp + from pre have k2: "int n \ int x" by simp + from k2 have k3: "int x * int n \ int n * int n" by (simp add: mult_mono) + from k1 k3 have k4: "int x * int n > int x * (int b * int n - int x)" by linarith + from pre k4 have k5: "int n > int b * int n - int x" by simp + + from pre have l0:"n = 0 \ x = 1" by simp + from l0 have l1: "n = 0 \ x = Suc 0" by simp + from l1 have l2: "n = 0 \ int n = \ b 0 \ int x = \ b (Suc 0)" by simp + from l2 have l3: "n = 0 \ \m. int n = \ b m \ int x = \ b (Suc m)" by blast + + have m0: "n > 0 \ int n * int n - 1 \ 0" by simp + from pre3 m0 have m1: "n > 0 \ int x * (int b * int n - int x) \ 0" by simp + from m1 have m2: "n > 0 \ int b * int n - int x \ 0" using zero_le_mult_iff by force + + from j0 have n0: "int x * int x - int b * int x * int n + int n * int n = 1" by simp + have n1: "(int b * int n - int x) * (int b * int n - int x) = int b * int n * (int b * int n - int x) - int x * (int b * int n - int x)" by (simp add: left_diff_distrib) + from n1 have n2: "int n * int n - int b * int n * (int b * int n - int x) + (int b * int n - int x) * (int b * int n - int x) = int n * int n - int x * (int b * int n - int x)" by simp + from n0 n2 j1 have n3: "int n * int n - int b * int n * (int b * int n - int x) + (int b * int n - int x) * (int b * int n - int x) = 1" by linarith + from n3 have n4: "int n * int n + (int b * int n - int x) * (int b * int n - int x) = 1 + int b * int n * (int b * int n - int x)" by simp + have n5: "int b * int n = int (b * n)" by simp + from n5 m2 have n6: "n > 0 \ int b * int n - int x = int (b * n - x)" by linarith + from n4 n6 have n7: "n > 0 \ int (n * n + (b * n - x) * (b * n - x)) = int (1 + b * n * (b * n - x))" by simp + from n7 have n8: "n > 0 \ n * n + (b * n - x) * (b * n - x) = 1 + b * n * (b * n - x)" using of_nat_eq_iff by blast + + note hyp = \\mx. m < x \ x * x + m * m = 1 + b * x * m \ + (\ma. int m = \ b ma \ int x = \ b (Suc ma))\ + + from k5 n6 n8 have o0: "n > 0 \ (b * n - x) < n \ n * n + (b * n - x) * (b * n - x) = 1 + b * n * (b * n - x)" by simp + from o0 hyp have o1: "n > 0 \ (\ma. int (b * n - x) = \ b ma \ int n = \ b (Suc ma))" by simp + + from o1 l3 n6 show ?case by force +qed + +lemma alpha_char_eq2: + assumes "(x*x + y*y = 1 + b * x * y)" "b>2" + shows "(\n. int x = \ b n)" +proof - + have "x \ y" + proof(rule ccontr, auto) + assume "x=y" + hence "2*x*x = 1+b*x*x" using assms by simp + hence "2*x*x \ 1+2*x*x" using assms by (metis add_le_mono le_less mult_le_mono1) + thus False by auto + qed + thus ?thesis + proof(cases "xn. int x = \ b n \ int y = \ b (Suc n)" using alpha_char_eq assms + by (simp add: add.commute power2_eq_square) + thus ?thesis by auto + next + case False + hence "\j. int y = \ b j \ int x = \ b (Suc j)" using alpha_char_eq assms \x \ y\ by auto + thus ?thesis by blast + qed +qed + + +subsubsection \Divisibility properties\ +text \The following lemmas are needed in the proof of equation 3.25\ +lemma representation: + fixes k m :: nat + assumes "k > 0" "n = m mod k" "l = (m-n)div k" + shows "m = n+k*l \ 0\n \ n\k-1" by (metis Suc_pred' assms le_add2 le_add_same_cancel2 + less_Suc_eq_le minus_mod_eq_div_mult minus_mod_eq_mult_div mod_div_mult_eq + mod_less_divisor neq0_conv nonzero_mult_div_cancel_left) + +lemma div_3251: + fixes b k m:: nat + assumes "b>2" and "k>0" + defines "n \ m mod k" + defines "l \ (m-n) div k" + shows "A b m = mat_mul (A b n) (mat_pow l (A b k))" +proof - + from assms(2) l_def n_def representation have m: "m = n+k*l \ 0\n \ n\k-1" by simp + from A_pow assms(1) have Abm2: "A b m = mat_pow m (B b)" by simp + from m have Bm: "mat_pow m (B b) =mat_pow (n+k*l) (B b)" by simp + from mat_exp_law have as1: "mat_pow (n+k*l) (B b) + = mat_mul (mat_pow n (B b)) (mat_pow (k*l) (B b))" by simp + from mat_exp_law_mult have as2: "mat_pow (k*l) (B b) = mat_pow l (mat_pow k (B b))" + by (metis mult.commute) + from A_pow assms have Abn: "mat_pow n (B b) = A b n" by simp + from A_pow assms(1) have Ablk: "mat_pow l (mat_pow k (B b)) = mat_pow l (A b k)" by simp + from Ablk Abm2 Abn Bm as1 as2 show Abm: "A b m = mat_mul (A b n) (mat_pow l (A b k))" by simp +qed + +lemma div_3252: + fixes a b c d m :: int and l :: nat + defines "M \ mat a b c d" + assumes "mat_21 M mod m = 0" + shows "(mat_21 (mat_pow l M)) mod m = 0 " (is "?P l") +proof(induction l) + show "?P 0" by simp +next + fix l assume IH: "?P l" + define Ml where "Ml = mat_pow l M" + have S1: "mat_pow (Suc(l)) M = mat_mul M (mat_pow l M)" by simp + have S2: "mat_21 (mat_mul M Ml) = mat_21 M * mat_11 Ml + mat_22 M * mat_21 Ml" + by (rule_tac mat_mul.induct mat_plus.induct, auto) + have S3: "mat_21 (mat_pow (Suc(l)) M) = mat_21 M * mat_11 Ml + mat_22 M * mat_21 Ml" + using S1 S2 Ml_def by simp + from assms(2) have S4: "(mat_21 M * mat_11 Ml) mod m = 0" by auto + from IH Ml_def have S5: " mat_22 M * mat_21 Ml mod m = 0" by auto + from S4 S5 have S6: "(mat_21 M * mat_11 Ml + mat_22 M * mat_21 Ml) mod m = 0" by auto + from S3 S6 show "?P (Suc(l))" by simp +qed + +lemma div_3253: + fixes a b c d m:: int and l :: nat + defines "M \ mat a b c d" + assumes "mat_21 M mod m = 0" + shows "((mat_11 (mat_pow l M)) - a^l) mod m = 0" (is "?P l") +proof(induction l) + show "?P 0" by simp +next + fix l assume IH: "?P l" + define Ml where "Ml = mat_pow l M" + from Ml_def have S1: "mat_pow (Suc(l)) M = mat_mul M Ml" by simp + have S2: "mat_11 (mat_mul M Ml) = mat_11 M * mat_11 Ml + mat_12 M * mat_21 Ml" + by (rule_tac mat_mul.induct mat_plus.induct, auto) + hence S3: "mat_11 (mat_pow (Suc(l)) M) = mat_11 M * mat_11 Ml + mat_12 M * mat_21 Ml" + using S1 by simp + from M_def Ml_def assms(2) div_3252 have S4: "mat_21 Ml mod m = 0" by auto + from IH Ml_def have S5: "(mat_11 Ml - a ^ l) mod m = 0" by auto + from IH M_def have S6: "(mat_11 M -a) mod m = 0" by simp + from S4 have S7: "(mat_12 M * mat_21 Ml) mod m = 0" by auto + from S5 S6 have S8: "(mat_11 M * mat_11 Ml- a^(Suc(l))) mod m = 0" + by (metis M_def mat2.sel(1) mod_0 mod_mult_right_eq mult_zero_right power_Suc right_diff_distrib) + have S9: "(mat_11 M * mat_11 Ml - a^(Suc(l)) + mat_12 M * mat_21 Ml ) mod m = 0" + using S7 S8 by auto + from S9 have S10: "(mat_11 M * mat_11 Ml + mat_12 M * mat_21 Ml - a^(Suc(l))) mod m = 0" by smt + from S3 S10 show "?P (Suc(l))" by auto +qed + +text \Equation 3.25\ +lemma divisibility_lemma1: + fixes b k m:: nat + assumes "b>2" and "k>0" + defines "n \ m mod k" + defines "l \ (m-n) div k" + shows "\ b m mod \ b k = \ b n * (\ b (k+1)) ^ l mod \ b k" +proof - + from assms(2) l_def n_def representation have m: "m = n+k*l \ 0\n \ n\k-1" by simp + consider (eq0) "n = 0" | (neq0) "n > 0" by auto + thus ?thesis + proof cases + case eq0 + have Abm_gen: "A b m = mat_mul (A b n) (mat_pow l (A b k))" + using assms div_3251 l_def n_def by blast + have Abk: "mat_pow l (A b k) = mat_pow l (mat (\ b (k+1)) (-\ b k) (\ b k) (-\ b (k-1)))" + using assms(2) neq0_conv by (metis A.elims) + from eq0 have Abm: "A b m = mat_pow l (mat (\ b (k+1)) (-\ b k) (\ b k) (-\ b (k-1)))" + using A_pow \b>2\ apply (auto simp: A.simps B.simps) + by (metis Abk Suc_eq_plus1 add.left_neutral m mat_exp_law_mult mult.commute) + have Abm1: "mat_21 (A b m) = \ b m" by (metis A.elims \.simps(1) mat2.sel(3)) + have Abm2: "mat_21 (mat_pow l (mat (\ b (k+1)) (-\ b k) (\ b k) (-\ b (k-1)))) mod (\ b k) = 0" + using Abm div_3252 by simp + from Abm Abm1 Abm2 have MR0: "\ b m mod \ b k = 0" by simp + from MR0 eq0 show ?thesis by simp + next case neq0 + from assms have Abm_gen: "A b m = mat_mul (A b n) (mat_pow l (A b k))" + using div_3251 l_def n_def by blast + from assms(2) neq0_conv have Abk: "mat_pow l (A b k) + = mat_pow l (mat (\ b (k+1)) (-\ b k) (\ b k) (-\ b (k-1)))" by (metis A.elims) + from n_def neq0 have N0: "n>0" by simp + define M where "M = mat (\ b (n + 1)) (-(\ b n)) (\ b n) (-(\ b (n - 1)))" + define N where "N = mat_pow l (mat (\ b (k+1)) (-\ b k) (\ b k) (-\ b (k-1)))" + from Suc_pred' neq0 have Abn: "A b n = mat (\ b (n + 1)) (-(\ b n)) (\ b n) (-(\ b (n - 1)))" + by (metis A.elims neq0_conv) + from Abm_gen Abn Abk M_def N_def have Abm: "A b m = mat_mul M N" by simp + (* substitutions done! next: calculations *) + from Abm have S1: "mat_21 (mat_mul M N) = mat_21 M * mat_11 N + mat_22 M * mat_21 N" + by (rule_tac mat_mul.induct mat_plus.induct, auto) + have S2: "mat_21 (A b m) = \ b m" by (metis A.elims \.simps(1) mat2.sel(3)) + from S1 S2 Abm have S3: "\ b m = mat_21 M * mat_11 N + mat_22 M * mat_21 N" by simp + from S3 have S4: "(\ b m - (mat_21 M * mat_11 N + mat_22 M * mat_21 N)) mod (\ b k) = 0" by simp + from M_def have S5: "mat_21 M = \ b n" by simp + from div_3253 N_def have S6: "(mat_11 N - (\ b (k+1)) ^ l) mod (\ b k) = 0" by simp + from N_def Abm div_3252 have S7: "mat_21 N mod (\ b k) = 0" by simp + from S4 S7 have S8: "(\ b m - mat_21 M * mat_11 N) mod (\ b k) = 0" by algebra + from S5 S6 have S9: "(mat_21 M * mat_11 N - (\ b n) * (\ b (k+1)) ^ l) mod (\ b k) = 0" + by (metis mod_0 mod_mult_left_eq mult.commute mult_zero_left right_diff_distrib') + from S8 S9 show ?thesis + proof - + have "(mat_21 M * mat_11 N - \ b m) mod \ b k = 0" + using S8 by presburger + hence "\i. (\ b m - (mat_21 M * mat_11 N - i)) mod \ b k = i mod \ b k" + by (metis (no_types) add.commute diff_0_right diff_diff_eq2 mod_diff_right_eq) + thus ?thesis + by (metis (no_types) S9 diff_0_right mod_diff_right_eq) + qed + qed + qed + +text \Prerequisite lemma for 3.27\ +lemma div_coprime: + assumes "b>2" "n \ 0" + shows "coprime (\ b k) (\ b (k+1))" (is "?P") +proof(rule ccontr) + assume as: "\ ?P" + define n where "n = gcd (\ b k) (\ b (k+1))" + from n_def have S1: "n > 1" + using alpha_det1 as assms(1) coprime_iff_gcd_eq_1 gcd_pos_int right_diff_distrib' + by (smt add.commute plus_1_eq_Suc) + have S2: "(\ b (Suc k))^2 - (int b) * \ b (Suc k) * (\ b k) + (\ b k)^2 = 1" + using alpha_det1 assms by auto + from n_def have D1: " n dvd (\ b (k+1))^2" by (simp add: numeral_2_eq_2) + from n_def have D2: " n dvd (- (int b) * \ b (k+1) * (\ b k))" by simp + from n_def have D3: "n dvd (\ b k)^2" by (simp add: gcd_dvdI1) + have S3: "n dvd ((\ b (Suc k))^2 - (int b) * \ b (Suc k) * (\ b k) + (\ b k)^2)" + using D1 D2 D3 by simp + from S2 S3 have S4: "n dvd 1" by simp + from S4 n_def as is_unit_gcd show "False" by blast +qed + +text \Equation 3.27\ +lemma divisibility_lemma2: + fixes b k m:: nat + assumes "b>2" and "k>0" + defines "n \ m mod k" + defines "l \ (m-n) div k" + assumes "\ b k dvd \ b m" + shows "\ b k dvd \ b n" +proof - + from assms(2) l_def n_def representation have m: "0\n \ n\k-1" by simp + from divisibility_lemma1 assms(1) assms(2) l_def n_def have S1: + "(\ b m) mod (\ b k) = (\ b n) * (\ b (k+1)) ^ l mod (\ b k)" by blast + from S1 assms(5) have S2: "(\ b k) dvd ((\ b n) * (\ b (k+1)) ^ l)" by auto + show ?thesis + using S1 div_coprime S2 assms(1) apply auto + using coprime_dvd_mult_left_iff coprime_power_right_iff by blast +qed + +text \Equation 3.23 - main result of this section\ +theorem divisibility_alpha: + assumes "b>2" and "k>0" + shows "\ b k dvd \ b m \ k dvd m" (is "?P \ ?Q") +proof + assume Q: "?Q" + define n where "n = m mod k" + have N: "n=0" by (simp add: Q n_def) + from N have Abn: "\ b n = 0" by simp + from Abn divisibility_lemma1 assms(1) assms(2) mult_eq_0_iff n_def show "?P" + by (metis dvd_0_right dvd_imp_mod_0 mod_0_imp_dvd) +next + assume P: "?P" + define n where "n = m mod k" + define l where "l = (m-n) div k" + define B where "B = (mat (int b) (-1) 1 0)" + have S1: "(\ b n) mod (\ b k) = 0" + using divisibility_lemma2 assms(1) assms(2) n_def P by simp + from n_def assms(2) have m: "n < k" using mod_less_divisor by blast + from alpha_strictly_increasing m assms(1) have S2: "\ b n < \ b k" + by (smt less_imp_of_nat_less lift_Suc_mono_less of_nat_0_less_iff pos2) + from S1 S2 have S3: "n=0" + by (smt alpha_superlinear assms(1) mod_pos_pos_trivial neq0_conv of_nat_0_less_iff) + from S3 n_def show "?Q" by auto +qed + +subsubsection \Divisibility properties (continued)\ + +text \Equation 3.28 - main result of this section\ + +lemma divisibility_equations: + assumes 0: "m = k*l" and "b>2" "m>0" + shows "A b m = mat_pow l (mat_minus (mat_scalar_mult (\ b k) (B b)) + (mat_scalar_mult (\ b (k-1)) ID))" + apply (auto simp del: mat_pow.simps mat_mul.simps mat_minus.simps mat_scalar_mult.simps + simp add: A_pow mult.commute[of k l] assms mat_exp_law_mult) + using A_pow[of b k] \m>0\ + apply (auto simp: A.simps \m>0\ ID_def B.simps) + using A.simps(2) alpha_n One_nat_def Suc_eq_plus1 Suc_pred assms \m>0\ assms + mult.commute nat_0_less_mult_iff + by (smt mat_exp_law_mult) + +lemma divisibility_cong: + fixes e f :: int + fixes l :: nat + fixes M :: mat2 + assumes "mat_22 M = 0" "mat_21 M = 1" + shows "(mat_21 (mat_pow l (mat_minus (mat_scalar_mult e M) (mat_scalar_mult f ID)))) mod e^2 = (-1)^(l-1)*l*e*f^(l-1)*(mat_21 M) mod e^2 + \ mat_22 (mat_pow l (mat_minus (mat_scalar_mult e M) (mat_scalar_mult f ID))) mod e^2 = (-1)^l *f^l mod e^2" +(is "?P l \ ?Q l" ) +proof(induction l) + case 0 + then show ?case by simp +next + case (Suc l) + have S2: "mat_pow (Suc(l)) (mat_minus (mat_scalar_mult e M) (mat_scalar_mult f ID)) = + mat_mul (mat_pow l (mat_minus (mat_scalar_mult e M) (mat_scalar_mult f ID))) (mat_minus (mat_scalar_mult e M) (mat_scalar_mult f ID))" + using mat_exp_law[of l _ 1] mat2.sel by (auto, metis)+ + define a1 where "a1 = mat_11 (mat_minus (mat_scalar_mult e M) (mat_scalar_mult f ID))" + define b1 where "b1 = mat_12 (mat_minus (mat_scalar_mult e M) (mat_scalar_mult f ID))" + define c1 where "c1 = mat_21 (mat_minus (mat_scalar_mult e M) (mat_scalar_mult f ID))" + define d1 where "d1 = mat_22 (mat_minus (mat_scalar_mult e M) (mat_scalar_mult f ID))" + define a where "a = mat_11 M" + define b where "b = mat_12 M" + define c where "c = mat_21 M" + define d where "d = mat_22 M" + define g where "g = mat_21 (mat_pow l (mat_minus (mat_scalar_mult e M) (mat_scalar_mult f ID)))" + define h where "h = mat_22 (mat_pow l (mat_minus (mat_scalar_mult e M) (mat_scalar_mult f ID)))" + from S2 g_def a1_def h_def c1_def have S3: "mat_21 (mat_pow (Suc(l)) (mat_minus (mat_scalar_mult e M) (mat_scalar_mult f ID))) = g*a1 + h*c1" + by simp + from S2 g_def b1_def h_def d1_def have S4: "mat_22 (mat_pow (Suc(l)) (mat_minus (mat_scalar_mult e M) (mat_scalar_mult f ID))) + = g*b1+h*d1" by simp + have S5: "mat_11 (mat_scalar_mult e M) = e*a" by (simp add: a_def) + have S6: "mat_12 (mat_scalar_mult e M) = e*b" by (simp add: b_def) + have S7: "mat_21 (mat_scalar_mult e M) = e*c" by (simp add: c_def) + have S8: "mat_22 (mat_scalar_mult e M) = e*d" by (simp add: d_def) + from a1_def S5 have S9: "a1 = e*a-f"by (simp add: Exp_Matrices.ID_def) + from b1_def S6 have S10: "b1 = e*b" by (simp add: Exp_Matrices.ID_def) + from c1_def S7 have S11: "c1 = e*c" by (simp add: Exp_Matrices.ID_def) + from S11 assms(2) c_def have S115: "c1 = e" by simp + from d1_def S8 have S12: "d1 = e*d - f" by (simp add: Exp_Matrices.ID_def) + from S12 assms(1) d_def have S125: "d1 = - f" by simp + from assms(2) c_def Suc g_def c_def have S13: "g mod e^2 = (-1)^(l-1)*l*e*f^(l-1)*c mod e^2" by blast + from assms(2) c_def S13 have S135: "g mod e^2 = (-1)^(l-1)*l*e*f^(l-1) mod e^2" by simp + from Suc h_def have S14: "h mod e^2 = (-1)^l *f^l mod e^2" by simp + from S10 S135 have S27: "g*b1 mod e^2 = (-1)^(l-1)*l*e*f^(l-1)*e*b mod e^2" by (metis mod_mult_left_eq mult.assoc) + from S27 have S28: "g*b1 mod e^2 = 0 mod e^2" by (simp add: power2_eq_square) + from S125 S14 mod_mult_cong have S29: "h*d1 mod e^2 = (-1)^l *f^l*(- f) mod e^2" by blast + from S29 have S30: "h*d1 mod e^2 = (-1)^(l+1) *f^l*f mod e^2" by simp + from S30 have S31: "h*d1 mod e^2 = (-1)^(l+1) *f^(l+1) mod e^2" by (metis mult.assoc power_add power_one_right) + from S31 have F2: "?Q (Suc(l))" by (metis S28 S4 Suc_eq_plus1 add.left_neutral mod_add_cong) + from S9 S13 have S15: "g*a1 mod e^2 = ((-1)^(l-1)*l*e*f^(l-1)*c*(e*a-f))mod e^2" by (metis mod_mult_left_eq) + have S16: "((-1)^(l-1)*l*e*f^(l-1)*c*(e*a-f)) = ((-1)^(l-1)*l*e^2*f^(l-1)*c*a) - f*(-1)^(l-1)*l*e*f^(l-1)*c" by algebra + have S17: "((-1)^(l-1)*l*e^2*f^(l-1)*c*a) mod e^2 = 0 mod e^2" by simp + from S17 have S18: "(((-1)^(l-1)*l*e^2*f^(l-1)*c*a) - f*(-1)^(l-1)*l*e*f^(l-1)*c) mod e^2 = + - f*(-1)^(l-1)*l*e*f^(l-1)*c mod e^2" + proof - + have f1: "\i ia. (ia::int) - (0 - i) = ia + i" + by auto + have "\i ia. ((0::int) - ia) * i = 0 - ia * i" + by auto + then show ?thesis using f1 + proof - + have f1: "\i. (0::int) - i = - i" + by presburger + then have "\i. (i - - ((- 1) ^ (l - 1) * int l * e\<^sup>2 * f ^ (l - 1) * c * a)) mod e\<^sup>2 = i mod e\<^sup>2" + by (metis (no_types) S17 \\i ia. ia - (0 - i) = ia + i\ add.right_neutral mod_add_right_eq) + then have "\i. ((- 1) ^ (l - 1) * int l * e\<^sup>2 * f ^ (l - 1) * c * a - i) mod e\<^sup>2 = - i mod e\<^sup>2" + using f1 by (metis \\i ia. ia - (0 - i) = ia + i\ uminus_add_conv_diff) + then show ?thesis + using f1 \\i ia. (0 - ia) * i = 0 - ia * i\ by presburger + qed + qed + from S15 S16 S18 have S19: "g*a1 mod e^2 = - f*(-1)^(l-1)*l*e*f^(l-1)*c mod e^2" by presburger + from S11 S14 have S20: "h*c1 mod e^2 = (-1)^l *f^l*e*c mod e^2" by (metis mod_mult_left_eq mult.assoc) + from S19 S20 have S21: "(g*a1 + h*c1) mod e^2 = (- f*(-1)^(l-1)*l*e*f^(l-1)*c + (-1)^l *f^l*e*c) mod e^2" using mod_add_cong by blast + from assms(2) c_def have S22: "(- f*(-1)^(l-1)*l*e*f^(l-1)*c + (-1)^l *f^l*e*c) mod e^2=(- f*(-1)^(l-1)*l*e*f^(l-1) + (-1)^l *f^l*e) mod e^2" by simp + have S23: "(- f*(-1)^(l-1)*l*e*f^(l-1) + (-1)^l *f^l*e) mod e^2 = (f*(-1)^(l)*l*e*f^(l-1) + (-1)^l *f^l*e) mod e^2" + by (smt One_nat_def Suc_pred mult.commute mult_cancel_left2 mult_minus_left neq0_conv of_nat_eq_0_iff power.simps(2)) + have S24: "(f*(-1)^(l)*l*e*f^(l-1) + (-1)^l *f^l*e) mod e^2 = ((-1)^(l)*l*e*f^l + (-1)^l *f^l*e) mod e^2" + by (smt One_nat_def Suc_pred mult.assoc mult.commute mult_eq_0_iff neq0_conv of_nat_eq_0_iff power.simps(2)) + have S25: "((-1)^(l)*l*e*f^l + (-1)^l *f^l*e) mod e^2 = ((-1)^(l)*(l+1)*e*f^l) mod e^2" + proof - + have f1: "\i ia. (ia::int) * i = i * ia" + by simp + then have f2: "\i ia. (ia::int) * i - - i = i * (ia - - 1)" + by (metis (no_types) mult.right_neutral mult_minus_left right_diff_distrib') + have "\n. int n - - 1 = int (n + 1)" + by simp + then have "e * (f ^ l * (int l * (- 1) ^ l - - ((- 1) ^ l))) mod e\<^sup>2 = e * (f ^ l * ((- 1) ^ l * int (l + 1))) mod e\<^sup>2" + using f2 by presburger + then have "((- 1) ^ l * int l * e * f ^ l - - ((- 1) ^ l) * f ^ l * e) mod e\<^sup>2 = (- 1) ^ l * int (l + 1) * e * f ^ l mod e\<^sup>2" + using f1 + proof - + have f1: "\i ia ib. (i::int) * (ia * ib) = ia * (i * ib)" + by simp + then have "\i ia ib. (i::int) * (ia * ib) - - (i * ib) = (ia - - 1) * (i * ib)" + by (metis (no_types) \\i ia. ia * i = i * ia\ f2) + then show ?thesis + using f1 by (metis (no_types) \\i ia. ia * i = i * ia\ \e * (f ^ l * (int l * (- 1) ^ l - - ((- 1) ^ l))) mod e\<^sup>2 = e * (f ^ l * ((- 1) ^ l * int (l + 1))) mod e\<^sup>2\ f2 mult_minus_right) + qed + then show ?thesis + by simp + qed + from S21 S22 S23 S24 S25 have S26: "(g*a1 + h*c1) mod e^2 = ((-1)^(l)*(l+1)*e*f^l) mod e^2" by presburger + from S3 S26 have F1: "?P (Suc(l))" by (metis Suc_eq_plus1 assms(2) diff_Suc_1 mult.right_neutral) + from F1 F2 show ?case by simp +qed + +lemma divisibility_congruence: + assumes "m = k*l" and "b>2" "m>0" + shows "\ b m mod (\ b k)^2 = ((-1)^(l-1)*l*(\ b k)*(\ b (k-1))^(l-1)) mod (\ b k)^2" +proof - + have S0: "\ b m = mat_21 (A b m)" by (metis A.elims assms(3) mat2.sel(3) neq0_conv) + from assms S0 divisibility_equations have S1: "\ b m = + mat_21 ( mat_pow l (mat_minus (mat_scalar_mult (\ b k) (B b)) + (mat_scalar_mult (\ b (k-1)) ID)))" by auto + have S2: "mat_21 (B b) = 1" using Binomial.binomial_ring by (simp add: Exp_Matrices.B.simps) + have S3: "mat_22 (B b) = 0" by (simp add: Exp_Matrices.B.simps) + from S1 S2 S3 divisibility_cong show ?thesis by (metis mult.right_neutral) +qed + +text \ Main result section 3.5 \ +theorem divisibility_alpha2: + assumes "b>2" "m>0" + shows "(\ b k)^2 dvd (\ b m) \ k*(\ b k) dvd m" (is "?P \ ?Q") +proof + assume Q: "?Q" + then show "?P" + proof(cases "k dvd m") + case True + then obtain l where mkl: "m = k * l" by blast + from Q assms mkl have S0: "l mod \ b k = 0" by simp + from S0 have S1: "l*(\ b k) mod (\ b k)^2 = 0" by (simp add: power2_eq_square) + from S1 have S2: "((-1)^(l-1)*l*(\ b k)*(\ b (k-1))^(l-1)) mod (\ b k)^2 = 0" + proof - + have "\i. \ b k * (int l * i) mod (\ b k)\<^sup>2 = 0" + by (metis (no_types) S1 mod_0 mod_mult_left_eq mult.assoc mult.left_commute mult_zero_left) + then show ?thesis + by (simp add: mult.assoc mult.left_commute) + qed + from assms divisibility_congruence mkl have S3: + "\ b m mod (\ b k)^2 = ((-1)^(l-1)*l*(\ b k)*(\ b (k-1))^(l-1)) mod (\ b k)^2" by simp + from S2 S3 have S4: "\ b m mod (\ b k)^2 = 0" by linarith + then show ?thesis by auto + next + case False + then show ?thesis using Q dvd_mult_left int_dvd_int_iff by blast + qed +next + assume P: "?P" + show "?Q" + proof(cases "k dvd m") + case True + then obtain l where mkl: "m = k * l" by blast + from assms mkl divisibility_congruence have S0: + "\ b m mod (\ b k)^2 = ((-1)^(l-1)*l*(\ b k)*(\ b (k-1))^(l-1)) mod (\ b k)^2" by simp + from S0 P have S1: "(\ b k)^2 dvd ((-1)^(l-1)*l*(\ b k)*(\ b (k-1))^(l-1))" by auto + from S1 have S2: "(\ b k)^2 dvd l*(\ b k)*(\ b (k-1))^(l-1)" + by (metis (no_types, opaque_lifting) Groups.mult_ac(1) dvd_trans dvd_triv_right left_minus_one_mult_self) + from S2 have S3: "(\ b k) dvd l*(\ b (k-1))^(l-1)" + by (metis (full_types) Exp_Matrices.alpha_superlinear assms(1) assms(2) mkl + mult.assoc mult.commute mult_0 not_less_zero of_nat_le_0_iff power2_eq_square zdvd_mult_cancel) + from div_coprime Suc_eq_plus1 Suc_pred' assms(1) assms(2) mkl less_imp_le_nat nat_0_less_mult_iff + have S4: "coprime (\ b k) (\ b (k-1))" by (metis coprime_commute) + hence "coprime (\ b k) ((\ b (k-1))^(l-1))" using coprime_power_right_iff by blast + hence "(\ b k) dvd l" using S3 using coprime_dvd_mult_left_iff by blast + then show ?thesis by (simp add: mkl) + next + case False + then show ?thesis + apply(cases "0Congruence properties\ +text \In this section we will need the inverse matrices of A and B\ +fun A_inv :: "nat \ nat \ mat2" where + "A_inv b n = mat (-\ b (n-1)) (\ b n) (-\ b n) (\ b (n+1))" + +fun B_inv :: "nat \ mat2" where + "B_inv b = mat 0 1 (-1) b" + +lemma A_inv_aux: "b>2 \ n>0 \ \ b n * \ b n - \ b (Suc n) * \ b (n - Suc 0) = 1" + apply (induction n, auto) subgoal for n using alpha_det1[of b n] apply auto by algebra done + +lemma A_inverse[simp]: "b>2 \ n>0 \ mat_mul (A_inv b n) (A b n) = ID" + using mat2.expand[of "mat_mul (A_inv b n) (A b n)" ID] apply rule + using ID_def A.simps(2)[of _ "n-1"] ID_def apply (auto) + subgoal using mat2.sel(1)[of 1 0 0 1] apply (auto) + using A_inv_aux[of b n] by (auto simp: mult.commute) + subgoal by (metis mat2.sel(2)) + subgoal by (metis mat2.sel(3)) + subgoal using mat2.sel(4)[of 1 0 0 1] apply (auto) + using A_inv_aux[of b n] by (auto simp: mult.commute) + done + +lemma B_inverse[simp]: "mat_mul (B b) (B_inv b) = ID" using B.simps ID_def by auto + +declare A_inv.simps B_inv.simps[simp del] + + +text \Equation 3.33\ +lemma congruence: + assumes "b1 mod q = b2 mod q" + shows "\ b1 n mod q = \ b2 n mod q" +proof (induct n rule:nat_less_induct) + case (1 n) + note hyps = \\m b1 m mod q = \ b2 m mod q\ + have n0:"(\ b1 0) mod q = (\ b2 0) mod q" by simp + have n1:"(\ b1 1) mod q = (\ b2 1) mod q" by simp + from hyps have s1: "n>1\\ b1 (n-1) mod q = \ b2 (n-1) mod q" by auto + from hyps have s2: "n>1\\ b1 (n-2) mod q = \ b2 (n-2) mod q" by auto + have s3: "n>1 \ \ b1 (Suc (Suc n)) = (int b1) * (\ b1 (Suc n)) - (\ b1 n)" by simp + from s3 have s4: "n>1 \ (\ b1 n = (int b1*(\ b1 (n-1)) - \ b1 (n-2)))" + by (smt Suc_1 Suc_diff_Suc diff_Suc_1 alpha_n lessE) + have sw: "n>1 \ \ b2 (Suc (Suc n)) = (int b2) * (\ b2 (Suc n)) - (\ b2 n)" by simp + from sw have sx: "n>1 \ (\ b2 n = (int b2*(\ b2 (n-1)) - \ b2 (n-2)))" + by (smt Suc_1 Suc_diff_Suc diff_Suc_1 alpha_n lessE) + from n0 n1 s1 s2 s3 s4 assms(1) mod_mult_cong have s5: "n>1 + \ b1*(\ b1 (n-1)) mod q = b2*(\ b2 (n-1)) mod q " by (smt mod_mult_eq of_nat_mod) + from hyps have sq: "n>1 \ \ b1 (n-2) mod q = \ b2 (n-2) mod q " by simp + from s5 sq have sd: "n>1 \-( (\ b1 (n-2))) mod q = -((\ b2 (n-2))) mod q " + by (metis mod_minus_eq) + from sd s5 mod_add_cong have s6: "n>1 \ ( b1*(\ b1 (n-1)) - \ b1 (n-2)) mod q + = (b2*(\ b2 (n-1)) - \ b2 (n-2)) mod q" by force + from s4 have sa: "n>1 \( b1*(\ b1 (n-1)) - \ b1 (n-2)) mod q = (\ b1 n) mod q" by simp + from sx have sb: "n>1 \( b2*(\ b2 (n-1)) - \ b2 (n-2)) mod q = (\ b2 n) mod q" by simp + from sb sa s6 sx have s7: "n>1 \ (\ b1 n) mod q = (b2*(\ b2 (n-1)) - \ b2 (n-2)) mod q" by simp + from s7 sx s6 have s9: "\ b1 n mod q = \ b2 n mod q" + by (metis One_nat_def \.simps(1) \.simps(2) less_Suc0 nat_neq_iff) + from s9 n0 n1 show ?case by simp +qed + +text \Equation 3.34\ +lemma congruence2: + fixes b1 :: nat + assumes "b>=2" + shows "(\ b n) mod (b - 2) = n mod (b - 2)" +proof- + from alpha_linear have S1: "\ (nat 2) n = n" by simp + define q where "q = b - (nat 2)" + from q_def assms le_mod_geq have S4: "b mod q = 2 mod q" by auto + from assms S4 congruence have SN: "(\ b n) mod q = (\ 2 n) mod q" by blast + from S1 SN q_def zmod_int show ?thesis by simp +qed + +lemma congruence_jpos: + fixes b m j l :: nat + assumes "b>2" and "2*l*m+j>0" + defines "n \ 2*l*m+j" + shows "A b n = mat_mul (mat_pow l (mat_pow 2 (A b m))) (A b j)" +proof- + from A_pow assms(1) have Abm2: "A b n = mat_pow n (B b)" by simp + from Abm2 n_def have Bn: "mat_pow n (B b) =mat_pow (2*l*m+j) (B b)" by simp + from mat_exp_law have as1: "mat_pow (2*l*m+j) (B b) = mat_mul (mat_pow l (mat_pow m (mat_pow 2 (B b)))) (mat_pow (j) (B b))" + by (metis (no_types, lifting) mat_exp_law_mult mult.commute) + from A_pow assms(1) B.elims mult.commute mat_exp_law_mult have as2: "mat_mul (mat_pow l (mat_pow m (mat_pow 2 (B b)))) (mat_pow (j) (B b)) + = mat_mul (mat_pow l (mat_pow 2 (A b m))) (A b j)" by metis + from as2 as1 Abm2 Bn show ?thesis by auto +qed + + +lemma congruence_inverse: "b>2 \ mat_pow (n+1) (B_inv b) = A_inv b (n+1)" + apply (induction n, simp add: B_inv.simps, auto) by (auto simp add: B_inv.simps) + +lemma congruence_inverse2: + fixes n b :: nat + assumes "b>2" + shows "mat_mul (mat_pow n (B b)) (mat_pow n (B_inv b)) = mat 1 0 0 1" +proof(induct n) + case 0 + thus ?case by simp +next + case (Suc n) + have S1: "mat_pow (Suc(n)) (B b) = mat_mul (B b) (mat_pow n (B b))" by simp + have S2: "mat_pow (Suc(n)) (B_inv b) = mat_mul (mat_pow n (B_inv b)) (B_inv b)" + proof - + have "\i ia ib ic. mat_pow 1 (mat ic ib ia i) = mat ic ib ia i" + by simp + hence "\m ma mb. mat_pow 1 m = m \ mat_mul mb m \ ma" by (metis mat2.exhaust) + thus ?thesis + by (metis (no_types) One_nat_def add_Suc_right diff_Suc_Suc diff_zero mat_exp_law mat_pow.simps(1) mat_pow.simps(2)) + qed + define "C" where "C= (B b)" + define "D" where "D = mat_pow n C" + define "E" where "E = B_inv b" + define "F" where "F = mat_pow n E" + from S1 S2 C_def D_def E_def F_def have S3: "mat_mul (mat_pow (Suc(n)) C) (mat_pow (Suc(n)) E) = mat_mul (mat_mul C D) (mat_mul F E)" by simp + from S3 mat_associativity mat2.exhaust C_def D_def E_def F_def have S4: "mat_mul (mat_pow (Suc(n)) C) (mat_pow (Suc(n)) E) + = mat_mul C (mat_mul (mat_mul D F) E)" by metis + from S4 Suc.hyps mat_neutral_element C_def D_def E_def F_def have S5: "mat_mul (mat_pow (Suc(n)) C) (mat_pow (Suc(n)) E) = mat_mul C E" by simp + from S5 C_def E_def show ?case using B_inverse ID_def by auto +qed + +lemma congruence_mult: + fixes m :: nat + assumes "b>2" + shows "n>m ==> mat_pow (nat(int n- int m)) (B b) = mat_mul (mat_pow n (B b)) (mat_pow m (B_inv b))" +proof(induction n) + case 0 + thus ?case by simp +next + case (Suc n) + consider (eqm) "n == m" | (gm) "n < m" | (lm) "n>m" by linarith + thus ?case + proof cases + case gm + from Suc.prems gm not_less_eq show ?thesis by simp + next case lm + have S1: "mat_pow (nat(int (Suc(n)) - int m)) (B b) = mat_mul (B b) (mat_pow (nat(int n - int m)) (B b))" + by (metis Suc.prems Suc_diff_Suc diff_Suc_1 diff_Suc_Suc mat_pow.simps(2) nat_minus_as_int) + from lm S1 Suc.IH have S2: "mat_pow (nat(int (Suc(n)) - int m)) (B b) = mat_mul (B b) (mat_mul (mat_pow n (B b)) (mat_pow m (B_inv b)))" by simp + from S2 mat_associativity mat2.exhaust have S3: "mat_pow (nat(int (Suc(n)) - int m)) (B b) = mat_mul (mat_mul (B b) (mat_pow n (B b))) (mat_pow m (B_inv b))" by metis + from S3 show ?thesis by simp + next case eqm + from eqm have S1: "nat(int (Suc(n))- int m) = 1" by auto + from S1 have S2: "mat_pow (nat(int (Suc(n))- int m)) (B b) == B b" by simp + from eqm have S3: "(mat_pow (Suc(n)) (B b)) = mat_mul (B b) (mat_pow m (B b))" by simp + from S3 have S35: "mat_mul (mat_pow (Suc(n)) (B b)) (mat_pow m (B_inv b)) = mat_mul (mat_mul (B b) (mat_pow m (B b))) (mat_pow m (B_inv b))" by simp + from mat2.exhaust S35 mat_associativity have S4: "mat_mul (mat_pow (Suc(n)) (B b)) (mat_pow m (B_inv b)) + = mat_mul (B b) (mat_mul (mat_pow m (B b)) (mat_pow m (B_inv b)))" by smt + from congruence_inverse2 assms have S5: "mat_mul (mat_pow m (B b)) (mat_pow m (B_inv b)) = mat 1 0 0 1" by simp + have S6: "mat_mul (B b) (B_inv b) = mat 1 0 0 1" using ID_def B_inverse by auto + from S5 S6 eqm have S7: "mat_mul (mat_pow n (B b)) (mat_pow m (B_inv b)) = mat 1 0 0 1" by metis + from S7 have S8: "mat_mul (B b) (mat_mul (mat_pow n (B b)) (mat_pow m (B_inv b))) == B b" by simp + from eqm S2 S4 S8 show ?thesis by simp + qed +qed + +lemma congruence_jneg: + fixes b m j l :: nat + assumes "b>2" and "2*l*m > j" and "j>=1" + defines "n \ nat(int 2*l*m- int j)" + shows "A b n = mat_mul (mat_pow l (mat_pow 2 (A b m))) (A_inv b j)" +proof- + from A_pow assms(1) have Abm2: "A b n = mat_pow n (B b)" by simp + from Abm2 n_def have Bn: "A b n = mat_pow (nat(int 2*l*m- int j)) (B b)" by simp + from Bn congruence_mult assms(1) assms(2) have Bn2: "A b n = mat_mul (mat_pow (2*l*m) (B b)) (mat_pow j (B_inv b))" by fastforce + from assms(1) assms(3) congruence_inverse Bn2 add.commute le_Suc_ex have Bn3: "A b n = mat_mul (mat_pow (2*l*m) (B b)) (A_inv b j)" by smt + from Bn3 A_pow assms(1) mult.commute B.simps mat_exp_law_mult have as3: "A b n = mat_mul (mat_pow l (mat_pow 2 (A b m))) (A_inv b j)" by metis + from as3 A_pow add.commute assms(1) mat_exp_law mat_exp_law_mult show ?thesis by simp +qed + +lemma matrix_congruence: + fixes Y Z :: mat2 + fixes b m j l :: nat + assumes "b>2" + defines "X \ mat_mul Y Z" + defines "a \ mat_11 Y" and "b0\ mat_12 Y" and "c \ mat_21 Y" and "d \ mat_22 Y" + defines "e \ mat_11 Z" and "f \ mat_12 Z" and "g \ mat_21 Z" and "h \ mat_22 Z" + defines "v \ \ b (m+1) - \ b (m-1)" + assumes "a mod v = a1 mod v" and "b0 mod v = b1 mod v" and "c mod v = c1 mod v" and "d mod v = d1 mod v" + shows "mat_21 X mod v = (c1*e+d1*g) mod v \ mat_22 X mod v = (c1*f+ d1*h) mod v" (is "?P \ ?Q") +proof - + (* proof of ?P *) + from X_def mat2.exhaust_sel c_def e_def d_def g_def have P1: "mat_21 X = (c*e+d*g)" + using mat2.sel by auto + from assms(14) mod_mult_cong have P2: "(c*e) mod v = (c1*e) mod v" by blast + from assms(15) mod_mult_cong have P3: "(d*g) mod v = (d1*g) mod v" by blast + from P2 P3 mod_add_cong have P4: "(c*e+d*g) mod v = (c1*e+d1*g) mod v" by blast + from P1 P4 have F1: ?P by simp + (* proof of ?Q *) + from X_def mat2.exhaust_sel c_def f_def d_def h_def mat2.sel(4) mat_mul.simps + have Q1: "mat_22 X = (c*f+d*h)" by metis + from assms(14) mod_mult_cong have Q2: "(c*f) mod v = (c1*f) mod v" by blast + from assms(15) mod_mult_cong have Q3: "(d*h) mod v = (d1*h) mod v" by blast + from Q1 Q2 Q3 mod_add_cong have F2: ?Q by fastforce + from F1 F2 show ?thesis by auto +qed + +text \3.38\ +lemma congruence_Abm: + fixes b m n :: nat + assumes "b>2" + defines "v \ \ b (m+1) - \ b (m-1)" + shows "(mat_21 (mat_pow n (mat_pow 2 (A b m))) mod v = 0 mod v) + \ (mat_22 (mat_pow n (mat_pow 2 (A b m))) mod v = ((-1)^n) mod v)" (is "?P n \ ?Q n") +proof(induct n) +case 0 + from mat2.exhaust have S1: "mat_pow 0 (mat_pow 2 (A b m)) = mat 1 0 0 1" by simp + thus ?case by simp +next + case (Suc n) + define Z where "Z = mat_pow 2 (A b m)" + define Y where "Y = mat_pow n Z" + define X where "X = mat_mul Y Z" + define c where "c = mat_21 Y" + define d where "d = mat_22 Y" + define e where "e = mat_11 Z" + define f where "f = mat_12 Z" + define g where "g = mat_21 Z" + define h where "h = mat_22 Z" + define d1 where "d1 = (-1)^n mod v" + from d_def d1_def Z_def Y_def Suc.hyps have S1: "d mod v = d1 mod v" by simp + from matrix_congruence assms(1) X_def v_def c_def d_def e_def d1_def g_def S1 + have S2: "mat_21 X mod v = (c*e+d1*g) mod v" by blast + from Z_def Y_def c_def Suc.hyps have S3: "c mod v = 0 mod v" by simp + consider (eq0) "m = 0" | (g0) "m>0" by blast + hence S4: "g mod v = 0" + proof cases + case eq0 + from eq0 have S1: "A b m = mat 1 0 0 1" using A.simps by simp + from S1 Z_def div_3252 g_def show ?thesis by simp + next + case g0 + from g0 A.elims neq0_conv + have S1: "A b m = mat (\ b (m + 1)) (-(\ b m)) (\ b m) (-(\ b (m - 1)))" by metis + from S1 assms(1) mat2.sel(3) mat_mul.simps mat_pow.simps + have S2: "mat_21 (mat_pow 2 (A b m)) = (\ b m)*(\ b (m+1)) + (-\ b (m-1))*(\ b m)" + by (auto) + from S2 g_def Z_def g0 A.elims neq0_conv + have S3: "g = (\ b (m+1))*(\ b m)- (\ b m)*(\ b (m-1))" by simp + from S3 g_def v_def mod_mult_self1_is_0 mult.commute right_diff_distrib show ?thesis by metis + qed + from S2 S3 S4 Z_def div_3252 g_def mat2.exhaust_sel mod_0 have F1: "?P (Suc(n))" by metis + (* Now proof of Q *) + from d_def d1_def Z_def Y_def Suc.hyps have Q1: "d mod v = d1 mod v" by simp + from matrix_congruence assms(1) X_def v_def c_def d_def f_def d1_def h_def S1 + have Q2: "mat_22 X mod v = (c*f+d1*h) mod v" by blast + from Z_def Y_def c_def Suc.hyps have Q3: "c mod v = 0 mod v" by simp + consider (eq0) "m = 0" | (g0) "m>0" by blast + hence Q4: "h mod v = (-1) mod v" + proof cases + case eq0 + from eq0 have S1: "A b m = mat 1 0 0 1" using A.simps by simp + from eq0 v_def have S2: "v = 1" by simp + from S1 S2 show ?thesis by simp + next + case g0 + from g0 A.elims neq0_conv have S1: "A b m = mat (\ b (m + 1)) (-(\ b m)) (\ b m) (-(\ b (m - 1)))" by metis + from S1 A_pow assms(1) mat2.sel(4) mat_exp_law mat_exp_law_mult mat_mul.simps mult_2 + have S2: "mat_22 (mat_pow 2 (A b m)) = (\ b m)*(-(\ b m)) + (-(\ b (m - 1)))*(-(\ b (m - 1)))" + by auto + from S2 Z_def h_def have S3: "h = -(\ b m)*(\ b m) + (\ b (m - 1))*(\ b (m - 1))" by simp + from v_def add.commute diff_add_cancel mod_add_self2 have S4: "(\ b (m - 1)) mod v = \ b (m+1) mod v" by metis + from S3 S4 mod_diff_cong mod_mult_left_eq mult.commute mult_minus_right uminus_add_conv_diff + have S5: "h mod v = (-(\ b m)*(\ b m) + (\ b (m - 1))*(\ b (m + 1))) mod v" by metis + from One_nat_def add.right_neutral add_Suc_right \.elims diff_Suc_1 g0 le_imp_less_Suc le_simps(1) neq0_conv Suc_diff_1 alpha_n + have S6: "\ b (m + 1) = b* (\ b m)- \ b (m-1)" + by (smt Suc_eq_plus1 Suc_pred' \.elims alpha_superlinear assms(1) g0 nat.inject of_nat_0_less_iff of_nat_1 of_nat_add) + from S6 have S7: "(\ b (m - 1))*(\ b (m + 1)) = (int b) * (\ b (m-1) * (\ b m)) - (\ b (m-1))^2" + proof - + have f1: "\i ia. - ((ia::int) * i) = ia * - i" by simp + have "\i ia ib ic. (ic::int) * (ib * ia) + ib * i = ib * (ic * ia + i)" by (simp add: distrib_left) + thus ?thesis using f1 by (metis S6 ab_group_add_class.ab_diff_conv_add_uminus power2_eq_square) + qed + from S7 have S8: "(-(\ b m)*(\ b m) + (\ b (m - 1))*(\ b (m + 1))) + = -1*(\ b (m-1))^2 + (int b) * (\ b (m-1) * (\ b m)) - (\ b m)^2" by (simp add: power2_eq_square) + from alpha_det2 assms(1) g0 have S9: "-1*(\ b (m-1))^2 + (int b) * (\ b (m-1) * (\ b m)) - (\ b m)^2 = -1" by smt + from S5 S8 S9 show ?thesis by simp + qed + from Q2 Q3 Q4 Suc_eq_plus1 add.commute add.right_neutral d1_def mod_add_right_eq mod_mult_left_eq mod_mult_right_eq mult.right_neutral + mult_minus1 mult_minus_right mult_zero_left power_Suc have Q5: "mat_22 X mod v = (-1)^(n+1) mod v" by metis + from Q5 Suc_eq_plus1 X_def Y_def Z_def mat_exp_law mat_exp_law_mult mult.commute mult_2 one_add_one have F2: "?Q (Suc(n))" by metis + from F1 F2 show ?case by blast +qed + +text \ 3.36 requires two lemmas 361 and 362 \ + +lemma 361: + fixes b m j l :: nat + assumes "b>2" + defines "n \ 2*l*m + j" + defines "v \ \ b (m+1) - \ b (m-1)" + shows "(\ b n) mod v = ((-1)^l * \ b j) mod v" +proof - + define Y where "Y = mat_pow l (mat_pow 2 (A b m))" + define Z where "Z = A b j" + define X where "X = mat_mul Y Z" + define c where "c = mat_21 Y" + define d where "d = mat_22 Y" + define e where "e = mat_11 Z" + define g where "g = mat_21 Z" + define d1 where "d1 = (-1)^l mod v" + from congruence_Abm assms(1) d_def v_def Y_def d1_def have S0: "d mod v = d1 mod v" by simp_all + from matrix_congruence assms(1) X_def v_def c_def d_def e_def d1_def g_def S0 have S1: "mat_21 X mod v = (c*e+d1*g) mod v" by blast + from congruence_Abm d1_def v_def mod_mod_trivial have S2: "d1 mod v = (-1)^l mod v" by blast + from congruence_Abm Y_def assms(1) c_def v_def have S3: "c mod v = 0" by simp + from Z_def g_def A.elims \.simps(1) mat2.sel(3) mat2.exhaust have S4: "g = \ b j" by metis + from A_pow assms(1) mat_exp_law mat_exp_law_mult mult_2 mult_2_right n_def X_def Y_def Z_def have S5: "A b n = X" by metis + from S5 A.elims \.simps(1) mat2.sel(3) Z_def Y_def have S6: "mat_21 X = \ b n" by metis + from S2 S3 S4 S6 S1 add.commute mod_0 mod_mult_left_eq mod_mult_self2 mult_zero_left zmod_eq_0_iff show ?thesis by metis +qed + +lemma 362: +fixes b m j l :: nat + assumes "b>2" and "2*l*m > j" and "j>=1" + defines "n \ 2*l*m - j" + defines "v \ \ b (m+1) - \ b (m-1)" + shows "(\ b n) mod v = -((-1)^l * \ b j) mod v" +proof - + define Y where "Y = mat_pow l (mat_pow 2 (A b m))" + define Z where "Z = A_inv b j" + define X where "X = mat_mul Y Z" + define c where "c = mat_21 Y" + define d where "d = mat_22 Y" + define e where "e = mat_11 Z" + define g where "g = mat_21 Z" + define d1 where "d1 = (-1)^l mod v" + from congruence_Abm assms(1) d_def v_def Y_def d1_def have S0: "d mod v = d1 mod v" by simp_all + from matrix_congruence assms(1) X_def v_def c_def d_def e_def d1_def g_def S0 have S1: "mat_21 X mod v = (c*e+d1*g) mod v" by blast + from congruence_Abm d1_def v_def mod_mod_trivial have S2: "d1 mod v = (-1)^l mod v" by blast + from congruence_Abm Y_def assms(1) c_def v_def have S3: "c mod v = 0" by simp + from Z_def g_def have S4: "g = - \ b j" by simp + from congruence_jneg assms(1) assms(2) assms(3) n_def X_def Y_def Z_def have S5: "A b n = X" by (simp add: nat_minus_as_int) + from S5 A.elims \.simps(1) mat2.sel(3) Z_def Y_def have S6: "mat_21 X = \ b n" by metis + from S2 S3 S4 S6 S1 add.commute mod_0 mod_mult_left_eq mod_mult_self2 mult_minus_right mult_zero_left zmod_eq_0_iff show ?thesis by metis +qed + +text \Equation 3.36\ +lemma 36: +fixes b m j l :: nat + assumes "b>2" + assumes "(n = 2 * l * m + j \ (n = 2 * l * m - j \ 2 * l * m > j \ j \ 1)) " + defines "v \ \ b (m+1) - \ b (m-1)" + shows "(\ b n) mod v = \ b j mod v \ (\ b n) mod v = -\ b j mod v" using assms(2) + apply(auto) + subgoal using 361 assms(1) v_def + apply(cases "even l" ) + by simp+ + subgoal using 362 assms(1) v_def + apply(cases "even l" ) + by simp+ + done + +subsubsection \Diophantine definition of a sequence alpha\ +definition alpha_equations :: "nat \ nat \ nat + \ nat \ nat \ nat \ nat \ nat \ nat \ nat \ nat \ bool" where + "alpha_equations a b c r s t u v w x y = ( + \ \3.41\ b > 3 \ + \ \3.42\ u^2 + t ^ 2 = 1 + b * u * t \ + \ \3.43\ s^2 + r ^ 2 = 1 + b * s * r \ + \ \3.44\ r < s \ + \ \3.45\ u ^ 2 dvd s \ + \ \3.46\ v + 2 * r = (b) * s \ + \ \3.47\ w mod v = b mod v \ + \ \3.48\ w mod u = 2 mod u \ + \ \3.49\ 2 < w \ + \ \3.50\ x ^ 2 + y ^ 2 = 1 + w * x * y \ + \ \3.51\ 2 * a < u \ + \ \3.52\ 2 * a < v \ + \ \3.53\ a mod v = x mod v \ + \ \3.54\ 2 * c < u \ + \ \3.55\ c mod u = x mod u)" + +text \The sufficiency\ +lemma alpha_equiv_suff: + fixes a b c::nat + assumes "\r s t u v w x y. alpha_equations a b c r s t u v w x y" + shows "3 < b \ int a = (\ b c)" +proof - + from assms obtain r s t u v w x y where eq: "alpha_equations a b c r s t u v w x y" by auto + have 41: "b > 3" using alpha_equations_def eq by auto + have 42: "u^2 + t ^ 2 = 1 +b * u * t" using alpha_equations_def eq by auto + have 43: "s^2 + r ^ 2 = 1 + b * s * r" using alpha_equations_def eq by auto + have 44: "r < s" using alpha_equations_def eq by auto + have 45: "u ^ 2 dvd s" using alpha_equations_def eq by auto + have 46: "v + 2 * r = b * s" using alpha_equations_def eq by auto + have 47: "w mod v = b mod v" using alpha_equations_def eq by auto + have 48: "w mod u = 2 mod u" using alpha_equations_def eq by auto + have 49: "2 < w" using alpha_equations_def eq by auto + have 50: "x ^ 2 + y ^ 2 = 1 + w * x * y" using alpha_equations_def eq by auto + have 51: "2 * a < u" using alpha_equations_def eq by auto + have 52: "2 * a < v" using alpha_equations_def eq by auto + have 53: "a mod v = x mod v" using alpha_equations_def eq by auto + have 54: "2 * c < u" using alpha_equations_def eq by auto + have 55: "int c mod u = x mod u" using alpha_equations_def eq by auto + + have "b > 2" using \b>3\ by auto + have "u > 0" using 51 by auto + + (* These first three equations are all proved very similarly. *) + text \Equation 3.56\ + have "\k. u = \ b k" using 42 alpha_char_eq2 by (simp add: \2 < b\ power2_eq_square) + then obtain k where 56: "u = \ b k" by auto + + text \Equation 3.57\ + have "\m. s = \ b m \ r = \ b (m-1)" using 43 44 alpha_char_eq[of r s b] diff_Suc_1 + by (metis power2_eq_square) + then obtain m where 57: "s = \ b m \ r = \ b (m-1)" by auto + + (* These are some properties we need multiple times *) + have m_pos: "m \ 0" using "44" "57" not_less_eq by fastforce + have alpha_pos: "\ b m > 0" using "44" "57" by linarith + + + text \Equation 3.58\ + have "\n. x = \ w n" using 50 alpha_char_eq2 by (simp add: "49" power2_eq_square) + then obtain n where 58: "x = \ w n" by auto + + text \Equation 3.59\ + have "\l j. (n = 2 * l * m + j \ n = 2 * l * m - j \ 2 * l * m > j \ j \ 1) \ j \ m" + proof - + define q where "q = n mod m" + obtain p where p_def: "n = p * m + q" using mod_div_decomp q_def by auto + have q1: "q \ m" using 44 57 + by (metis diff_le_self le_0_eq le_simps(1) linorder_not_le mod_less_divisor nat_int q_def) + consider (c1) "even p" | (c2) "odd p" by auto + thus ?thesis + proof(cases) + case c1 + thus ?thesis using p_def q1 by blast + next + case c2 + obtain d where "p=2*d+1" using c2 oddE by blast + define l where "l=d+1" + hence jpt: "l>0" by simp + from \p=2*d+1\ l_def have c21: "p=2*l-1" by auto + have c22: "n=2*l*m-(m-q)" + by (metis Nat.add_diff_assoc2 add.commute c21 diff_diff_cancel diff_le_self jpt mult_eq_if + mult_is_0 neq0_conv p_def q1 zero_neq_numeral) + thus ?thesis using diff_le_self + by (metis add.left_neutral diff_add_inverse2 diff_zero less_imp_diff_less mult.right_neutral + mult_eq_if mult_zero_right not_less zero_less_diff) + qed + qed + + (* here we add some conditions that are implicit in the proof but necessary for Isabelle *) + then obtain l j where 59: " (n = 2 * l * m + j \ n = 2 * l * m - j \ 2 * l * m > j \ j \ 1) \ j \ m" by auto + + text \Equation 3.60\ + have 60: "u dvd m" + using 45 56 57 divisibility_alpha2[of b m k] \b>2\ + by (metis dvd_trans dvd_triv_right int_dvd_int_iff m_pos neq0_conv of_nat_power) + + text \Equation 3.61\ + have 61: "v = \ b (m+1)- \ b (m-1)" + proof- + have "v = b*(\ b m) - 2*(\ b (m-1))" using 46 57 by (metis add_diff_cancel_right' mult_2 of_nat_add of_nat_mult) + thus ?thesis using alpha_n[of b "m-1"] m_pos by auto + qed + + text \Equation 3.62.1\ + have "a mod v = \ b n mod v" using 53 58 47 congruence[of w v b n] by (simp add: zmod_int) + hence "a mod v = \ b j mod v \ a mod v = -\ b j mod v" using 36[of b] 61 59 \2 < b\ by auto + (* more usable version *) + hence 62: "v dvd (a+ \ b j) \ v dvd (a - \ b j)" using mod_eq_dvd_iff zmod_int by auto + + text \Equation 3.63\ + (* This could be rewritten in a single proof if none of the 631-634 are reused *) + + have 631: "2*\ b j \ 2* \ b m" using 59 alpha_strictly_increasing_general[of b j m] \2 < b\ by force + + have "b - 2 \ 2" using 41 by simp + moreover have "\ b m > 0" using "44" "57" by linarith + ultimately have 632: "2 * \ b m \ (b - 2) * \ b m" by auto + + have "(b - 2) * \ b m = b * \ b m - 2* \ b m" using \2 < b\ + by (simp add: int_distrib(4) mult.commute of_nat_diff) + moreover have "b * \ b m - 2* \ b m < b * \ b m - 2 * \ b (m - 1)" using "44" "57" by linarith + ultimately have 633: "(b - 2) * \ b m < b * \ b m - 2 * \ b (m - 1)" by auto + + have 634: " b * \ b m - 2 * \ b (m - 1) = v" using 61 alpha_n[of b "m-1"] m_pos by simp + + have 63: "2*\ b j < v" using 631 632 633 634 by auto + + text \Equation 3.64\ + + hence 64: "a = \ b j" + proof(cases "0 < a + \ b j") + case True + moreover have "a + \ b j < v" using 52 63 by linarith + ultimately show ?thesis using 62 + apply auto + subgoal using zdvd_not_zless by blast + subgoal + by (smt \2 < b\ alpha_superlinear dvd_add_triv_left_iff negative_zle zdvd_not_zless) + done + next + case False + hence "j = 0" using \2 < b\ alpha_strictly_increasing_general by force + thus ?thesis using False by auto + qed + + + text \Equation 3.65\ + have 65: "c mod u = n mod u" + proof - + have "c mod u = \ w n mod u" using 55 58 zmod_int by (simp add: ) + moreover have "... = n mod u" using 48 alpha_linear congruence zmod_int by presburger + ultimately show ?thesis by linarith + qed + + text \Equation 3.66\ + have "2 * j \ 2 * \ b j \ 2 * a < u" + using 51 alpha_superlinear \b>2\ by auto + hence 66: "2*j < u" using "64" by linarith + + text \Equation 3.67\ + have 652: "u dvd (n+j) \ u dvd (n-j)" using 60 59 by auto + + hence "c = j" using 66 54 + proof- + have "c + j < u" using 66 54 by linarith + thus ?thesis using 652 + apply auto + subgoal + by (metis "65" add_cancel_right_right dvd_eq_mod_eq_0 mod_add_left_eq mod_if not_add_less2 not_gr_zero) + subgoal + by (metis "59" "60" "65" "66" Nat.add_diff_assoc2 \\c + j < u; u dvd n + j\ \ c = j\ + add_diff_cancel_right' add_lessD1 dvd_mult le_add2 le_less mod_less mod_nat_eqI mult_2) + done + qed + + show ?thesis using \b>3\ 64 \c=j\ by auto +qed + + +text \ 3.7.2 The necessity \ + +lemma add_mod: + fixes p q :: int + assumes "p mod 2 = 0" "q mod 2 = 0" + shows "(p+q) mod 2 = 0 \ (p-q) mod 2 = 0" + using assms(1) assms(2) by auto + +lemma one_odd: + fixes b n :: nat + assumes "b>2" + shows "(\ b n) mod 2 = 1 \ (\ b (n+1)) mod 2 = 1" +proof(rule ccontr) + assume asm: "\(\ b n mod 2 = 1 \ \ b (n + 1) mod 2 = 1)" + from asm have step1: "(\ b n mod 2 = 0 \ \ b (n + 1) mod 2 = 0)" by simp + from step1 have s1: "(\ b n)^2 mod 2 = 0 \ (\ b (n+1))^2 mod 2 = 0" by auto + from step1 have s2: "(int b)*(\ b n)*(\ b (n+1)) mod 2 = 0" by auto + from s1 have s3: "((\ b (n+1))^2 + (\ b n)^2) mod 2 = 0" by auto + from s2 s3 add_mod have s4: "((\ b (n+1))^2 + (\ b n)^2 - (int b)*((\ b n)*(\ b (n+1)))) mod 2 = 0" + by (simp add: Groups.mult_ac(2) Groups.mult_ac(3)) + have s5: "(\ b (n+1))^2+(\ b n)^2-(int b)*((\ b n)*(\ b (n+1)))=(\ b (n+1))^2-(int b)*(\ b (n+1)*(\ b n))+(\ b n)^2" by simp + from s4 s5 have s6: "((\ b (n+1))^2-(int b)*(\ b (n+1)*(\ b n))+(\ b n)^2) mod 2 = 0" + proof - + have f1: "(\ b (n + 1))\<^sup>2 - int b * (\ b (n + 1) * \ b n) = (\ b (n + 1))\<^sup>2 + - 1 * (int b * (\ b (n + 1) * \ b n))" + by simp + have f2: "(\ b (n + 1))\<^sup>2 + - 1 * (int b * (\ b (n + 1) * \ b n)) + (\ b n)\<^sup>2 = (\ b (n + 1))\<^sup>2 + (\ b n)\<^sup>2 + - 1 * (int b * (\ b n * \ b (n + 1)))" + by simp + have "((\ b (n + 1))\<^sup>2 + (\ b n)\<^sup>2 + - 1 * (int b * (\ b n * \ b (n + 1)))) mod 2 = 0" + using s4 by fastforce + thus ?thesis using f2 f1 by presburger + qed + from s6 alpha_det1 show False by (simp add: assms mult.assoc) +qed + +lemma oneodd: + fixes b n :: nat + assumes "b>2" + shows "odd (\ b n) = True \ odd (\ b (n+1)) = True" + using assms odd_iff_mod_2_eq_one one_odd by auto + +lemma cong_solve_nat: "a \ 0 \ \x. (a*x) mod n = (gcd a n) mod n" + for a n :: nat + apply (cases "n=0") + apply auto + apply (insert bezout_nat [of a n], auto) + by (metis mod_mult_self4) + +lemma cong_solve_coprime_nat: "coprime (a::nat) (n::nat) \ \x. (a*x) mod n = 1 mod n" + using cong_solve_nat[of a n] coprime_iff_gcd_eq_1[of a n] by fastforce + +lemma chinese_remainder_aux_nat: + fixes m1 m2 :: nat + assumes a:"coprime m1 m2" + shows "\b1 b2. b1 mod m1 = 1 mod m1 \ b1 mod m2 = 0 mod m2 \ b2 mod m1 = 0 mod m1 \ b2 mod m2 = 1 mod m2" +proof - + from cong_solve_coprime_nat [OF a] obtain x1 where 1: "(m1*x1) mod m2 = 1 mod m2" by auto + from a have b: "coprime m2 m1" + by (simp add: coprime_commute) + from cong_solve_coprime_nat [OF b] obtain x2 where 2: "(m2*x2) mod m1 = 1 mod m1" by auto + have "(m1*x1) mod m1 = 0" by simp + have "(m2*x2) mod m2 = 0" by simp + show ?thesis using 1 2 + by (metis mod_0 mod_mult_self1_is_0) +qed + +lemma cong_scalar2_nat: "a mod m = b mod m \ (k*a) mod m = (k*b) mod m" + for a b k :: nat + by (rule mod_mult_cong) simp_all + +lemma chinese_remainder_nat: + fixes m1 m2 :: nat + assumes a: "coprime m1 m2" + shows "\x. x mod m1 = u1 mod m1 \ x mod m2 = u2 mod m2" +proof - + from chinese_remainder_aux_nat [OF a] obtain b1 b2 where "b1 mod m1 = 1 mod m1" and "b1 mod m2 = 0 mod m2" and +"b2 mod m1 = 0 mod m1" and "b2 mod m2 = 1 mod m2" by force + let ?x = "u1*b1+u2*b2" + have "?x mod m1 = (u1*1+u2*0) mod m1" + apply (rule mod_add_cong) + apply(rule cong_scalar2_nat) + apply (rule \b1 mod m1 = 1 mod m1\) + apply(rule cong_scalar2_nat) + apply (rule \b2 mod m1 = 0 mod m1\) + done + hence 1:"?x mod m1 = u1 mod m1" by simp + have "?x mod m2 = (u1*0+u2*1) mod m2" + apply (rule mod_add_cong) + apply(rule cong_scalar2_nat) + apply (rule \b1 mod m2 = 0 mod m2\) + apply(rule cong_scalar2_nat) + apply (rule \b2 mod m2 = 1 mod m2\) + done + hence "?x mod m2 = u2 mod m2" by simp + with 1 show ?thesis by blast +qed + +lemma nat_int1: "\(w::nat) (u::int).u>0 \ (w mod nat u = 2 mod nat u \ int w mod u = 2 mod u)" + by blast + +lemma nat_int2: "\(w::nat) (b::nat) (v::int).u>0 \ (w mod nat v = b mod nat v \ int w mod v = int b mod v)" + by (metis mod_by_0 nat_eq_iff zmod_int) + +lemma lem: + fixes u t::int and b::nat + assumes "u^2-int b*u*t+t^2=1" "u\0" "t\0" + shows "(nat u)^2+(nat t)^2=1+b*(nat u)*(nat t)" +proof - + define U where "U=nat u" + define T where "T=nat t" + from U_def T_def assms have UT: "int U=u \ int T = t" using int_eq_iff by blast + from UT have UT1: "int (b*U*T) = b*u*t" by simp + from UT have UT2: "int (U^2+T^2)=u^2+t^2" by simp + from UT2 assms have sth: "int (U^2+T^2)\b*u*t" by auto + from sth assms have sth1: "U^2+T^2\b*U*T" using UT1 by linarith + from sth1 of_nat_diff have sth2: "int (U^2+T^2-b*U*T) = int (U^2+T^2) - int (b*U*T)" by blast + from UT1 UT2 have UT3: "int (U^2+T^2)-int (b*U*T)=u^2+t^2-b*u*t" by simp + from sth2 UT3 assms have sth4: "int (U^2+T^2-b*U*T) = 1" by auto + from sth4 have sth5: "U^2+T^2-b*U*T=1" by simp + from sth5 have sth6: "U^2+T^2=1+b*U*T" by simp + show ?thesis using sth6 U_def T_def by simp +qed + +text \The necessity\ + +lemma alpha_equiv_nec: + "b > 3 \ a = \ b c \ \r s t u v w x y. alpha_equations a b c r s t u v w x y" +proof - + assume assms: "b > 3 \ a = \ b c" + have s1: "\(k::nat) (u::int) (t::int).u=\ b k \ odd u = True \ 2*int a u u^2-(int b)*u*t+t^2=1 \ k>0 \ t = \ b (k+1)" + proof - + define j::nat where "j=2*(a)+1" + have rd: "j>0" by (simp add: j_def) + consider (c1) "odd (\ b j) = True" | (c2) "odd (\ b (j+1)) = True" + using assms oneodd by fastforce + thus ?thesis + proof cases + case c1 + define k::nat where "k=j" + define u::int where "u=\ b k" + define t::int where "t=\ b (k+1)" + have stp: "k>0" by (simp add: k_def j_def) + from alpha_strictly_increasing assms have abc: "u2*a" by (simp add: k_def j_def) + from alpha_superlinear c12 have c13: "2*a b k" + define t::int where "t=\ b (k+1)" + have stc: "k>0" by (simp add: k_def j_def) + from alpha_strictly_increasing assms have abc: "u2*a" by (simp add: k_def j_def) + from alpha_superlinear c22 have c23: "2*a b k \ odd u = True \ 2*int a u u^2-(int b)*u*t+t^2=1 \ k>0 \ t= \ b (k+1)" by force + define m where "m=(nat u)*k" + define s where "s=\ b m" + define r where "r=\ b (m-1)" + note udef = \u = \ b k \ odd u = True \ 2 * int a < u \ u < t \ u\<^sup>2 - int b * u * t + t\<^sup>2 = 1 \ 0 < k \ t = \ b (k+1)\ + from assms have s211: "int b > 3" by simp + from assms alpha_superlinear have a354: "c\a" + by (simp add: nat_int_comparison(3)) + from a354 udef have 354: "2* int c b k \ int k" by simp + from alpha_strictly_increasing s211 s1 m_def s_def udef r_def have s212: "\ b (m-1) < \ b m" + by (smt One_nat_def Suc_pred nat_0_less_mult_iff zero_less_nat_eq) + from s212 r_def s_def have 344: "r b k) dvd (int m) \ k dvd m" by simp + from xyz divisibility_alpha2 have wxyz: "(\ b k)*(\ b k) dvd (\ b m)" by (smt assms dvd_mult_div_cancel int_nat_eq less_imp_le_nat m_def mult_pos_pos neq0_conv not_less not_less_eq numeral_2_eq_2 numeral_3_eq_3 of_nat_0_less_iff power2_eq_square udef) + from wxyz udef s_def have 345: "u^2 dvd s" by (simp add: power2_eq_square) + define v where "v = b*s-2*r" + from v_def s_def r_def alpha_n have 370: "v = \ b (m+1) - \ b (m-1)" + by (smt Suc_eq_plus1 add_diff_inverse_nat diff_Suc_1 neq0_conv not_less_eq s212 zero_less_diff) + have 371: "v = b*\ b m - 2*\ b (m-1)" using v_def s_def r_def by simp + from alpha_strictly_increasing assms m_def udef have asd: "\ b m > 0" + by (smt Suc_pred nat_0_less_mult_iff s211 zero_less_nat_eq) + from assms asd 371 have 372: "v\4*\ b m -2*\ b (m-1)" by simp + from 372 assms have 373: "v>2*\ b m \ 4*\ b m -2*\ b (m-1) > 2*\ b m" using s212 by linarith + from 373 assms alpha_superlinear have 374: "2*\ b m \ 2*m \ v>2*m" + by (smt One_nat_def Suc_eq_plus1 add_lessD1 distrib_right mult.left_neutral numeral_3_eq_3 of_nat_add one_add_one) + from udef have pre1: "k\1 \ u\1" using rd by linarith + from pre1 374 m_def have pre2: "m\u" by simp + from pre2 374 have 375: "2*m\2*u \ v>2*u" by simp + from 375 udef have 376: "2*u>2*a \ v>2*a" using pre1 by linarith + have u_v_coprime: "coprime u v" + proof - + obtain d::nat where d: "d = gcd u v" + by (metis gcd_int_def) + from \d = gcd u v\ have ddef: "d dvd u \ d dvd v" by simp + from 345 ddef have stp1: "d dvd s" using dvd_mult_left dvd_trans s_def udef wxyz by blast + from v_def stp1 ddef have stp2: "d dvd 2*r" by algebra + from ddef udef have d_odd: "odd d" using dvd_trans by auto + have r2even: "even (2*r)" by simp + from stp2 d_odd r2even have stp3: "(2*d) dvd (2*r)" by fastforce + from stp3 have stp4: "d dvd r" by simp + from stp1 stp4 have stp5: "d dvd s^2 \ d dvd (-int b*s*r) \ d dvd r^2" by (simp add: power2_eq_square) + from stp5 have stp6: "d dvd (s^2-int b*s*r+r^2)" by simp + from 343 stp6 have stp7: "d dvd 1" by simp + show ?thesis using stp7 d by auto + qed + have wdef: "\w::nat. int w mod u = 2 mod u \ int w mod v = int b mod v \ w>2" + proof - + from pre1 m_def have mg: "m\1" by auto + from s_def r_def 344 have srg: "s-r\1" by simp + from assms have bg: "b\4" by simp + from bg have bsr: "((int b)*s-2*r)\(4*s-2*r)" using 372 r_def s_def v_def by blast + have t1: "v\2+2*s" using bsr srg v_def by simp + from s_def have sg: "s\1" using asd by linarith + from sg t1 have t2: "v\4" by simp + from u_v_coprime have u_v_coprime1:"coprime (nat u) (nat v)" using pre1 t2 + using coprime_int_iff by fastforce + obtain z::nat where "z mod nat u = 2 mod nat u \ z mod nat v = b mod nat v" using chinese_remainder_nat u_v_coprime1 by force + note zdef = \z mod nat u = 2 mod nat u \ z mod nat v = b mod nat v\ + from t2 pre1 have t31: "nat v\4 \ nat u\1" by auto + from t31 have t3: "nat v*nat u\4" using mult_le_mono by fastforce + define w::nat where "w=z+ nat u*nat v" + from w_def t3 have t4: "w\4" by (simp add: mult.commute) + have t51: "(nat u*nat v) mod nat u = 0 \ (nat u*nat v) mod nat v = 0" using algebra by simp + from t51 w_def have t5: "w mod nat u = z mod nat u" by presburger + from t51 w_def have t6: "w mod nat v = z mod nat v" by presburger + from t5 t6 zdef have t7: "w mod nat u = 2 mod nat u \ w mod nat v = b mod nat v" by simp + from t7 t31 have t8: "int w mod u = 2 mod u \ int w mod v = int b mod v" + using nat_int1 nat_int2 by force + from t4 t8 show ?thesis by force + qed + obtain w::nat where "int w mod u = 2 mod u \ int w mod v = int b mod v \ w>2" using wdef by force + note wd = \int w mod u = 2 mod u \ int w mod v = int b mod v \ w>2\ + define x where "x = \ w c" + define y where "y = \ w (c+1)" + from alpha_det1 wd x_def y_def have 350: "x^2-int w*x*y+y^2 =1" by (metis add_gr_0 alpha_det2 diff_add_inverse2 less_one mult.assoc) + from x_def wd congruence have 353: "a mod v = x mod v" + by (smt "374" assms int_nat_eq nat_int nat_mod_distrib) + from congruence2 wd x_def have 379: "x mod int (w-2) = int c mod (int w-2)" + using int_ops(6) zmod_int by auto + from wd have wc: "u dvd (int w-2)" using mod_diff_cong mod_eq_0_iff_dvd by fastforce + from wc have wb: "\k1. u*k1=int w-2" by (metis dvd_def) + obtain k1 where "u*k1=int w-2" using wb by force + note k1def = \u*k1=int w-2\ + define r1 where "r1=int c mod (int w-2)" + from r1_def 379 have wa: "r1 = x mod (int w-2)" + using int_ops(6) wd by auto + obtain k2 where "int c = (int w-2)*k2+r1" by (metis mult_div_mod_eq r1_def) + note k2def = \int c = (int w-2)*k2+r1\ + from k2def k1def have a355: "int c = u*k1*k2+r1" by simp + from udef k1def k2def have bh: "u*k1*k2 mod u = 0" by (metis mod_mod_trivial mod_mult_left_eq mod_mult_self1_is_0 mult_eq_0_iff) + from a355 bh have b355: "(u*k1*k2+r1) mod u = r1 mod u" by (simp add: mod_eq_dvd_iff) + from a355 b355 have c355: "int c mod u = r1 mod u" by simp + from wa have waa: "\k3. x = k3*(int w-2)+r1" by (metis div_mult_mod_eq) + obtain k3 where "x=k3*(int w-2)+r1" using waa by force + from k1def \x = k3*(int w-2)+r1\ have d355: "x=u*k1*k3+r1" by simp + from udef k1def have ch: "u*k1*k3 mod u = 0" by (metis mod_mod_trivial mod_mult_left_eq mod_mult_self1_is_0 mult_eq_0_iff) + from d355 ch have e355: "(u*k1*k3+r1) mod u = r1 mod u" by (simp add: mod_eq_dvd_iff) + from d355 e355 have f355: "x mod u = r1 mod u" by simp + from c355 f355 have 355: "int c mod u = x mod u" by simp + from assms s1 wdef udef 343 344 345 v_def wd 350 376 353 354 355 have prefinal: "u^2-b*u*t+t^2=1 \ s^2-b*s*r+r^2=1 \ r u^2 dvd s \ b*s=v+2*r \ w mod v = b mod v \ w mod u = 2 mod u \ w>2 \ x^2-w*x*y+y^2=1 \ + 2*a 2*a a mod v = x mod v \ 2*c c mod u = x mod u" by fastforce + from alpha_strictly_increasing have s_pos: "s\0" using asd s_def by linarith + define S where "S=nat s" + from alpha_strictly_increasing have r_pos: "r\0" using asd r_def by (smt One_nat_def Suc_1 alpha_superlinear assms(1) lessI less_trans numeral_3_eq_3 of_nat_0_le_iff) + define R where "R=nat r" + from udef alpha_strictly_increasing have ut_pos:"u\0 \ t\0" using pre1 by linarith + from assms have a_pos: "a\0" using a354 by linarith + from a_pos have v_pos: "v\0" using "376" by linarith + from x_def y_def have xy_pos: "x\0 \ y\0" by (smt alpha_superlinear of_nat_0_le_iff wd) + define U where "U=nat u" + define T where "T=nat t" + define V where "V=nat v" + define X where "X=nat x" + define Y where "Y=nat y" + from lem U_def T_def S_def R_def X_def Y_def prefinal have lem1: "U^2+T^2=1+b*U*T \ S^2+R^2=1+b*S*R \ X^2+Y^2=1+w*X*Y" using s_pos ut_pos r_pos xy_pos by blast + from R_def S_def have lem2: "R2*r" using v_def v_pos by simp + from aq have aq1: "nat (int b*s) \ nat (2*r)" by simp + from s_pos r_pos assms have aq2: "nat (int b*s) = b*(nat s) \ nat (2*r) = 2*(nat r)" by (simp add: nat_mult_distrib) + from aq1 aq2 have aq3: "b*S\2*R" using S_def R_def by simp + from aq3 have aq4: "int (b*S-2*R) = int (b*S)-int (2*R)" using of_nat_diff by blast + have aq5: "int (b*S) = int b*int S \ int (2*R) = 2*int R" by simp + from aq4 aq5 have aq6: "int (b*S-2*R) = int b*s-2*r" using R_def S_def r_pos s_pos by simp + from aq6 v_def v_pos V_def have lem4: "b*S-2*R = V" by simp + from prefinal v_pos V_def ut_pos U_def xy_pos X_def a_pos have lem5: "w mod V = b mod V \ w mod U = 2 mod U \ a mod V=X mod V \ c mod U = X mod U" + by (metis int_nat_eq nat_int of_nat_numeral zmod_int) + from a_pos ut_pos v_pos U_def V_def prefinal have lem6: "2*nat a 2*nat a 2*c2" by simp + have third_last: "\b s v r::nat. b * s = v + 2 * r \ int (b * s) = int (v + 2 * r)" using of_nat_eq_iff by blast + have onemore: "\u t b. u ^ 2 + t ^ 2 = 1 + b * u * t \ int u ^ 2 + int t ^ 2 = 1 + int b * int u * int t" + by (metis (no_types) nat_int of_nat_1 of_nat_add of_nat_mult of_nat_power) + from lem1 lem2 lem3 lem4 lem5 lem6 lem7 third_last onemore show ?thesis + unfolding Exp_Matrices.alpha_equations_def[of a b c] apply auto + using assms apply blast + apply (rule exI[of _ R], rule exI[of _ S], rule exI[of _ T], rule exI[of _ U], simp) + apply (rule exI[of _ V], simp) + apply (rule exI[of _ w], simp) + apply (rule exI[of _ X], simp) + using aq4 aq5 lem5 by auto +qed + +subsubsection \Exponentiation is Diophantine\ + +text \Equations 3.80-3.83\ + +lemma 86: + fixes b r and q::int + defines "m \ b * q - q * q - 1" + shows "(q * \ b (r + 1) - \ b r) mod m = (q ^ (r + 1)) mod m" +proof(induction r) + case 0 + show ?case by simp +next + case (Suc n) + from m_def have a0: "(q * q - b * q + 1) mod m = ((-(q * q - b * q + 1)) mod m + (q * q - b * q + 1) mod m) mod m" by simp + have a1: "\ = 0" by (simp add:mod_add_eq) + from a0 a1 have a2: "(q * q - b * q + 1) mod m = 0" by simp + + from a2 have b0: "(b * q - 1) mod m = ((q * q - b * q + 1) mod m + (b * q - 1) mod m) mod m" by simp + have b1: "\ = (q * q) mod m" by (simp add: mod_add_eq) + from b0 b1 have b2: "(b * q - 1) mod m = (q * q) mod m" by simp + + have "(q * (\ b (Suc n + 1)) -\ b (Suc n)) mod m = (q * (int b *\ b (Suc n) -\ b n) -\ b (Suc n)) mod m" by simp + also have "\ = ((b * q - 1) *\ b (Suc n) - q *\ b n) mod m" by algebra + also have "\ = (((b * q - 1) *\ b (Suc n)) mod m - (q *\ b n) mod m) mod m" by (simp add: mod_diff_eq) + also have "\ = ((((b * q - 1) mod m) * ((\ b (Suc n)) mod m)) mod m - (q *\ b n) mod m) mod m" by (simp add: mod_mult_eq) + also have "\ = ((((q * q) mod m) * ((\ b (Suc n)) mod m)) mod m - (q *\ b n) mod m) mod m" by (simp add: b2) + also have "\ = (((q * q) * (\ b (Suc n))) mod m - (q *\ b n) mod m) mod m" by (simp add: mod_mult_eq) + also have "\ = ((q * q) * (\ b (Suc n)) - q *\ b n) mod m" by (simp add: mod_diff_eq) + also have "\ = (q * (q * (\ b (Suc n)) -\ b n)) mod m" by algebra + also have "\ = ((q mod m) * ((q * (\ b (Suc n)) -\ b n) mod m)) mod m" by (simp add: mod_mult_eq) + finally have c0: "(q * (\ b (Suc n + 1)) -\ b (Suc n)) mod m = ((q mod m) * ((q * (\ b (Suc n)) -\ b n) mod m)) mod m" by simp + from Suc.IH have c1: "\ = ((q ^ (n + 2))) mod m" by (simp add: mod_mult_eq) + + from c0 c1 show ?case by simp +qed + +text \This is a more convenient version of (86)\ + +lemma 860: + fixes b r and q::int + defines "m \ b * q - q * q - 1" + shows "(q * \ b r - (int b * \ b r - \ b (Suc r))) mod m = (q ^ r) mod m" +proof(cases "r=0") + case True + then show ?thesis by simp +next + case False + thus ?thesis using alpha_n[of b "r-1"] 86[of q b "r-1"] m_def by auto +qed + +text \We modify the equivalence (88) in a similar manner\ + +lemma 88: + fixes b r p q:: nat + defines "m \ int b * int q - int q * int q - 1" + assumes "int q ^ r < m" and "q > 0" + shows "int p = int q ^ r \ int p < m \ (q * \ b r - (int b * \ b r - \ b (Suc r))) mod m = int p mod m" + using "Exp_Matrices.860" assms(2) m_def by auto + +lemma 89: + fixes r p q :: nat + assumes "q > 0" + defines "b \ nat (\ (q + 4) (r + 1)) + q * q + 2" + defines "m \ int b * int q - int q * int q - 1" + shows "int q ^ r < m" +proof - + have a0: "int q * int q - 2 * int q + 1 = (int q - 1) * (int q - 1)" by algebra + from assms have a1: "int q * int q * int q \ int q * int q" by simp + from assms a0 a1 have a2: "\ > (int q - 1) * (int q - 1)" by linarith + + from alpha_strictly_increasing have c0: " \ (q + 4) (r + 1) > 0" by simp + from c0 have c1: " \ (q + 4) (r + 1) = int (nat (\ (q + 4) (r + 1)))" by simp + + then have b1: "(q+3) ^ r \ \ (q + 4) (r + 1)" using alpha_exponential_1[of "q+3"] + by(auto simp add: add.commute) + have b3: "int q ^ r \ (q + 3) ^ r" by (simp add: power_mono) + also have b4: "... \ (q + 3) ^ r * int q" using assms by simp + also from assms b1 have b5: "\ \ \ (q + 4) (r + 1) * int q" by simp + also from a2 have b6: "\ < \ (q + 4) (r + 1) * int q + int q * int q * int q - (int q - 1) * (int q - 1)" by simp + also have b7: "\ = ( \ (q + 4) (r + 1) + int q * int q + 2) * q - int q * int q - 1" by algebra + also from assms m_def have b8: "\ = m" using c1 by auto + finally show ?thesis by linarith +qed +end + +text \The final equivalence\ +theorem exp_alpha: + fixes p q r :: nat + shows "p = q ^ r \ ((q = 0 \ r = 0 \ p = 1) \ + (q = 0 \ 0 < r \ p = 0) \ + (q > 0 \ (\b m. + b = Exp_Matrices.\ (q + 4) (r + 1) + q * q + 2 \ + m = b * q - q * q - 1 \ + p < m \ + p mod m = ((q * Exp_Matrices.\ b r) - (int b * Exp_Matrices.\ b r - Exp_Matrices.\ b (r + 1))) mod m)))" +proof(cases "q>0") + case True + show ?thesis (is "?P = ?Q") + proof (rule) + assume ?P + define b where "b = nat (Exp_Matrices.\ (q + 4) (r + 1)) + q * q + 2" + define m where "m = int b * int q - int q * int q - 1" + have sg1: "int b = Exp_Matrices.\ (q + 4) (Suc r) + int q * int q + 2" using b_def + proof- + have "0 \ (Exp_Matrices.\ (q + 4) (r + 1))" using Exp_Matrices.alpha_exponential_1[of "q+3" r] + apply (simp add: add.commute) using zero_le_power[of "int q+3" r] by linarith + then show ?thesis using b_def int_nat_eq[of "(Exp_Matrices.\ (q + 4) (r + 1))"] by simp + qed + have sg2: "q ^ r < b * q - Suc (q * q)" using True "Exp_Matrices.89"[of q r] + of_nat_less_of_nat_power_cancel_iff[of q r " b * q - Suc (q * q)"] + b_def int_ops(6)[of "b * q" "Suc (q * q)"] of_nat_1 of_nat_add of_nat_mult plus_1_eq_Suc by smt + have sg3: "int (q ^ r mod (b * q - Suc (q * q))) + = (int q * Exp_Matrices.\ b r - (int b * Exp_Matrices.\ b r - Exp_Matrices.\ b (Suc r))) + mod int (b * q - Suc (q * q))" + proof- + have "int b * int q - int q * int q - 1 = b * q - Suc (q * q)" + using \q ^ r < b * q - Suc (q * q)\ int_ops(6) by auto + then show ?thesis using "Exp_Matrices.860"[of q b r] by (simp add: zmod_int) + qed + from sg1 sg2 sg3 True show ?Q + by (smt (verit) Suc_eq_plus1_left \p = q ^ r\ add.commute diff_diff_eq of_nat_mult) + next + assume Q: ?Q (is "?A \ ?B \ ?C") + thus ?P + proof (elim disjE) + show "?A \ ?P" by auto + show "?B \ ?P" by auto + show "?C \ ?P" + proof- + obtain b where b_def: "int b = Exp_Matrices.\ (q + 4) (Suc r) + int q * int q + 2" using Q True by auto + have prems3: "p < b * q - Suc (q * q)" using Q True b_def apply (simp add: add.commute) by (metis of_nat_eq_iff) + have prems4: " int p = (int q * Exp_Matrices.\ b r - ((Exp_Matrices.\ (q + 4) (Suc r) + + int q * int q + 2) * Exp_Matrices.\ b r - Exp_Matrices.\ b (Suc r))) mod int (b * q - Suc (q * q))" + using Q True b_def apply (simp add: add.commute) by (metis mod_less of_nat_eq_iff) + define m where "m = int b * int q - int q * int q - 1" + have "int q ^ r < int b * int q - int q * int q - 1" using "Exp_Matrices.89"[of q r] b_def True + by (smt Exp_Matrices.alpha_strictly_increasing One_nat_def Suc_eq_plus1 int_nat_eq nat_2 + numeral_Bit0 of_nat_0_less_iff of_nat_add of_nat_mult one_add_one) + moreover have "int p < m" by (smt gr_implies_not0 int_ops(6) int_ops(7) less_imp_of_nat_less + m_def of_nat_Suc of_nat_eq_0_iff prems3) + moreover have "(int q * Exp_Matrices.\ b r - (int b * Exp_Matrices.\ b r - Exp_Matrices.\ b (Suc r))) mod m = int p mod m" + using prems4 by (smt calculation(2) int_ops(6) m_def mod_pos_pos_trivial of_nat_0_le_iff + of_nat_1 of_nat_add of_nat_mult plus_1_eq_Suc b_def) + ultimately show ?thesis using True "Exp_Matrices.88"[of q "r" b p] m_def by simp + qed + qed + qed +next + case False + then show ?thesis by auto +qed + +lemma alpha_equivalence: + fixes a b c + shows "3 < b \ int a = Exp_Matrices.\ b c \ (\r s t u v w x y. Exp_Matrices.alpha_equations a b c r s t u v w x y)" + using Exp_Matrices.alpha_equiv_suff Exp_Matrices.alpha_equiv_nec + by meson+ + +end \ No newline at end of file diff --git a/thys/DPRM_Theorem/Diophantine/Modulo_Divisibility.thy b/thys/DPRM_Theorem/Diophantine/Modulo_Divisibility.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Diophantine/Modulo_Divisibility.thy @@ -0,0 +1,74 @@ +subsection \Mod is Diophantine\ + +theory Modulo_Divisibility + imports Existential_Quantifier +begin + +text \ Divisibility is diophantine \ +definition dvd ("DVD _ _" 1000) where "DVD Q R \ (BINARY (dvd) Q R)" + +lemma dvd_repr: + fixes a b :: nat + shows "a dvd b \ (\x. x * a = b)" + using dvd_class.dvd_def by auto + +lemma dvd_dioph[dioph]: "is_dioph_rel (DVD Q R)" +proof - + define Q' R' where pushed_defs: "Q' \ push_param Q 1" "R' \ push_param R 1" + define D where "D \ [\] (Param 0 [*] Q' [=] R')" + + have "eval (DVD Q R) a = eval D a" for a + unfolding D_def pushed_defs defs using push_push1 apply (auto simp: push0) + unfolding dvd_def by (auto simp: dvd_repr binary_eval) + + moreover have "is_dioph_rel D" + unfolding D_def by (auto simp: dioph) + + ultimately show ?thesis + by (auto simp: is_dioph_rel_def) +qed + +declare dvd_def[defs] + +(* Congruence is diophantine *) +definition mod ("MOD _ _ _" 1000) + where "MOD A B C \ (TERNARY (\a b c. a mod b = c mod b) A B C)" +declare mod_def[defs] + +lemma mod_repr: + fixes a b c :: nat + shows "a mod b = c mod b \ (\x y. c + x*b = a + y*b)" + by (metis mult.commute nat_mod_eq_iff) + +lemma mod_dioph[dioph]: + fixes A B C + defines "D \ (MOD A B C)" + shows "is_dioph_rel D" +proof - + define A' B' C' where pushed_defs: "A' \ push_param A 2" "B' \ push_param B 2" "C' \ push_param C 2" + define DS where "DS \ [\2] (Param 0 [*] B' [+] C' [=] Param 1 [*] B' [+] A')" + + have "eval DS a = eval D a" for a + proof + show "eval DS a \ eval D a" + unfolding DS_def defs D_def mod_def + by auto (metis mod_mult_self3 push_push_simp pushed_defs(1) pushed_defs(2) pushed_defs(3)) + show "eval D a \ eval DS a" + unfolding DS_def defs D_def mod_def + apply (auto simp add: mod_repr) + subgoal for x y + apply (rule exI[of _ "[x, y]"]) + unfolding pushed_defs by (simp add: push_push[where ?n = 2] push_list_eval) + done + qed + + moreover have "is_dioph_rel DS" + unfolding DS_def by (simp add: dioph) + + ultimately show ?thesis + by (auto simp: is_dioph_rel_def) +qed + +declare mod_def[defs] + +end \ No newline at end of file diff --git a/thys/DPRM_Theorem/Diophantine/Parametric_Polynomials.thy b/thys/DPRM_Theorem/Diophantine/Parametric_Polynomials.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Diophantine/Parametric_Polynomials.thy @@ -0,0 +1,82 @@ +section \Diophantine Equations\ + +theory Parametric_Polynomials + imports Main + abbrevs ++ = "\<^bold>+" and + -- = "\<^bold>-" and + ** = "\<^bold>*" and + 00 = "\<^bold>0" and + 11 = "\<^bold>1" +begin + +subsection \Parametric Polynomials\ + +text \This section defines parametric polynomials and builds up the infrastructure to later prove + that a given predicate or relation is Diophantine. The formalization follows + @{cite "h10lecturenotes"}.\ + +type_synonym assignment = "nat \ nat" + +text \Definition of parametric polynomials with natural number coefficients and their evaluation + function\ + +datatype ppolynomial = + Const nat | + Param nat | + Var nat | + Sum ppolynomial ppolynomial (infixl "\<^bold>+" 65) | + NatDiff ppolynomial ppolynomial (infixl "\<^bold>-" 65) | + Prod ppolynomial ppolynomial (infixl "\<^bold>*" 70) + +fun ppeval :: "ppolynomial \ assignment \ assignment \ nat" where + "ppeval (Const c) p v = c" | + "ppeval (Param x) p v = p x" | + "ppeval (Var x) p v = v x" | + "ppeval (D1 \<^bold>+ D2) p v = (ppeval D1 p v) + (ppeval D2 p v)" | + (* The next line lifts subtraction of type "nat \ nat \ nat" *) + "ppeval (D1 \<^bold>- D2) p v = (ppeval D1 p v) - (ppeval D2 p v)" | + "ppeval (D1 \<^bold>* D2) p v = (ppeval D1 p v) * (ppeval D2 p v)" + +definition Sq_pp ("_ \<^bold>^\<^bold>2" [99] 75) where "Sq_pp P = P \<^bold>* P" + +definition is_dioph_set :: "nat set \ bool" where + "is_dioph_set A = (\P1 P2::ppolynomial. \a. (a \ A) + \ (\v. ppeval P1 (\x. a) v = ppeval P2 (\x. a) v))" + +datatype polynomial = + Const nat | + Param nat | + Sum polynomial polynomial (infixl "[+]" 65) | + NatDiff polynomial polynomial (infixl "[-]" 65) | + Prod polynomial polynomial (infixl "[*]" 70) + +fun peval :: "polynomial \ assignment \ nat" where + "peval (Const c) p = c" | + "peval (Param x) p = p x" | + "peval (Sum D1 D2) p = (peval D1 p) + (peval D2 p)" | + (* The next line lifts subtraction of type "nat \ nat \ nat" *) + "peval (NatDiff D1 D2) p = (peval D1 p) - (peval D2 p)" | + "peval (Prod D1 D2) p = (peval D1 p) * (peval D2 p)" + +definition sq_p :: "polynomial \ polynomial" ("_ [^2]" [99] 75) where "sq_p P = P [*] P" + +definition zero_p :: "polynomial" ("\<^bold>0") where "zero_p = Const 0" +definition one_p :: "polynomial" ("\<^bold>1") where "one_p = Const 1" + +lemma sq_p_eval: "peval (P[^2]) p = (peval P p)^2" + unfolding sq_p_def by (simp add: power2_eq_square) + +fun convert :: "polynomial \ ppolynomial" where + "convert (Const c) = (ppolynomial.Const c)" | + "convert (Param x) = (ppolynomial.Param x)" | + "convert (D1 [+] D2) = (convert D1) \<^bold>+ (convert D2)" | + "convert (D1 [-] D2) = (convert D1) \<^bold>- (convert D2)" | + "convert (D1 [*] D2) = (convert D1) \<^bold>* (convert D2)" + +lemma convert_eval: "peval P a = ppeval (convert P) a v" (* implicit for all v *) + by (induction P, auto) + +definition list_eval :: "polynomial list \ assignment \ (nat \ nat)" where + "list_eval PL a = nth (map (\x. peval x a) PL)" + +end \ No newline at end of file diff --git a/thys/DPRM_Theorem/Diophantine/Register_Machine_Sums.thy b/thys/DPRM_Theorem/Diophantine/Register_Machine_Sums.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Diophantine/Register_Machine_Sums.thy @@ -0,0 +1,123 @@ +subsubsection \Preliminary: Register machine sums are Diophantine\ + +theory Register_Machine_Sums imports Diophantine_Relations + "../Register_Machine/RegisterMachineSimulation" + +begin + +fun sum_polynomial :: "(nat \ polynomial) \ nat list \ polynomial" where + "sum_polynomial f [] = Const 0" | + "sum_polynomial f (i#idxs) = f i [+] sum_polynomial f idxs" + +lemma sum_polynomial_eval: + "peval (sum_polynomial f idxs) a = (\k=0.. (xs ! xa) = (xs @ [x]) ! xa" for xa + by (simp add: nth_append) + ultimately show ?case by auto +qed + + +definition sum_program :: "program \ (nat \ polynomial) \ polynomial" + ("[\_] _" [100, 100] 100) where + "[\p] f \ sum_polynomial f [0.. length l = length p \ + peval ([\p] (\k. if g k then map (\x. push_param x m) l ! k else h k)) (push_list a ns) + = peval ([\p] (\k. if g k then l ! k else h k)) a" + unfolding sum_program_def apply (induction p, auto) + oops + +definition sum_radd_polynomial :: "program \ register \ (nat \ polynomial) \ polynomial" + ("[\R+] _ _ _") where + "[\R+] p l f \ [\p] (\k. if isadd (p!k) \ l = modifies (p!k) then f k else Const 0)" + +lemma sum_radd_polynomial_eval[defs]: + assumes "length p > 0" + shows "peval ([\R+] p l f) a = (\R+ p l (\x. peval (f x) a))" +proof - + have 1: "x \ length p - Suc 0 \ x < length p" for x using assms by linarith + have 2: "x \ length p - Suc 0 \ peval (f ([0.. register \ (nat \ polynomial) \ polynomial" + ("[\R-] _ _ _") where + "[\R-] p l f \ [\p] (\k. if issub (p!k) \ l = modifies (p!k) then f k else Const 0)" + +lemma sum_rsub_polynomial_eval[defs]: + assumes "length p > 0" + shows "peval ([\R-] p l f) a = (\R- p l (\x. peval (f x) a))" +proof - + have 1: "x \ length p - Suc 0 \ x < length p" for x using assms by linarith + have 2: "x \ length p - Suc 0 \ peval (f ([0.. state \ (nat \ polynomial) \ polynomial" + ("[\S+] _ _ _") where + "[\S+] p d f \ [\p] (\k. if isadd (p!k) \ d = goes_to (p!k) then f k else Const 0)" + +lemma sum_sadd_polynomial_eval[defs]: + assumes "length p > 0" + shows "peval ([\S+] p d f) a = (\S+ p d (\x. peval (f x) a))" +proof - + have 1: "x \ length p - Suc 0 \ x < length p" for x using assms by linarith + have 2: "x \ length p - Suc 0 \ peval (f ([0.. state \ (nat \ polynomial) \ polynomial" + ("[\S-] _ _ _") where + "[\S-] p d f \ [\p] (\k. if issub (p!k) \ d = goes_to (p!k) then f k else Const 0)" + +lemma sum_ssub_nzero_polynomial_eval[defs]: + assumes "length p > 0" + shows "peval ([\S-] p d f) a = (\S- p d (\x. peval (f x) a))" +proof - + have 1: "x \ length p - Suc 0 \ x < length p" for x using assms by linarith + have 2: "x \ length p - Suc 0 \ peval (f ([0.. state \ (nat \ polynomial) \ polynomial" + ("[\S0] _ _ _") where + "[\S0] p d f \ [\p] (\k. if issub (p!k) \ d = goes_to_alt (p!k) then f k else Const 0)" + +lemma sum_ssub_zero_polynomial_eval[defs]: + assumes "length p > 0" + shows "peval ([\S0] p d f) a = (\S0 p d (\x. peval (f x) a))" +proof - + have 1: "x \ length p - Suc 0 \ x < length p" for x using assms by linarith + have 2: "x \ length p - Suc 0 \ peval (f ([0..Wrap-Up: Combining all equations\ + +theory All_Equations + imports All_Equations_Invariance + +begin + +context register_machine +begin + +definition all_equations_relation :: "polynomial \ polynomial \ polynomial \ polynomial + \ polynomial \ polynomial \ polynomial \ polynomial list \ polynomial list \ polynomial list + \ relation" ("[ALLEQ] _ _ _ _ _ _ _ _ _ _") where + "[ALLEQ] a q b c d e f r z s + \ LARY (\ll. all_equations (ll!0!0) (ll!0!1) (ll!0!2) (ll!0!3) (ll!0!4) (ll!0!5) (ll!0!6) + (nth (ll!1)) (nth (ll!2)) (nth (ll!3))) + [[a, q, b, c, d, e, f], r, z, s]" + +lemma all_equations_dioph: + fixes A f e d c b q :: polynomial + fixes r z s :: "polynomial list" + assumes "length r = n" "length z = n" "length s = Suc m" + defines "DR \ [ALLEQ] A q b c d e f r z s" + shows "is_dioph_rel DR" +proof - + define DS where "DS \ ([REG] A b q r z s) + [\] ([STATE] b e q z s) + [\] ([MASK] c d e f r z) + [\] ([CONST] b c d e f q) + [\] [MISC] A c q" + + have "eval DS a = eval DR a" for a + unfolding DR_def DS_def all_equations_relation_def all_equations_def + unfolding register_equations_relation_def state_equations_relation_def + mask_equations_relation_def rm_constant_equations_def rm_miscellaneous_equations_def + by (simp add: defs) + + moreover have "is_dioph_rel DS" + unfolding DS_def apply (rule and_dioph)+ + apply (simp_all add: rm_constant_equations_dioph rm_miscellaneous_equations_dioph) + using assms reg_dioph[of r z s A b q] state_equations_dioph[of s z b e q] + mask_equations_relation_dioph[of r z c d e f] by metis+ + + ultimately show ?thesis using is_dioph_rel_def by auto +qed + +definition rm_equations :: "nat \ bool" where + "rm_equations a \ \ q :: nat. + \ b c d e f :: nat. + \ r z :: register \ nat. + \ s :: state \ nat. + all_equations a q b c d e f r z s" + +definition rm_equations_relation :: "polynomial \ relation" ("[RM] _") where + "[RM] A \ UNARY (rm_equations) A" + +(* Correct assumptions: Need validity of p *) + +lemma rm_dioph: + fixes A + fixes ic :: configuration + defines "DR \ [RM] A" + shows "is_dioph_rel DR" +proof - + define q b c d e f where "q \ Param 0" and + "b \ Param 1" and + "c \ Param 2" and + "d \ Param 3" and + "e \ Param 4" and + "f \ Param 5" + + define r where "r \ map Param [6.. map Param [n+6..<2*n + 6]" + define s where "s \ map Param [2*n + 6..<2*n + 6 + m + 1]" + + define number_of_ex_vars where "number_of_ex_vars \ 2*n + 6 + m + 1" + + define A' where "A' \ push_param A number_of_ex_vars" + + define DS where "DS \ [\number_of_ex_vars] [ALLEQ] A' q b c d e f r z s" + + have "length r = n" and "length z = n" and "length s = Suc m" + unfolding r_def z_def s_def by auto + + have "eval DS a = eval DR a" for a + proof (rule) + assume "eval DS a" + then obtain ks where + ks_length: "number_of_ex_vars = length ks" and + ALLEQ: "eval ([ALLEQ] A' q b c d e f r z s) (push_list a ks)" + unfolding DS_def by (auto simp add: defs) + + define q' b' c' d' e' f' where "q' \ ks!0" and + "b' \ ks!1" and + "c' \ ks!2" and + "d' \ ks!3" and + "e' \ ks!4" and + "f' \ ks!5" + + define r_list where "r_list \ (take n (drop 6 ks))" + define z_list where "z_list \ (take n (drop (6+n) ks))" + define s_list where "s_list \ (drop (6 + 2*n) ks)" + + define r' where "r' \ (!) r_list" + define z' where "z' \ (!) z_list" + define s' where "s' \ (!) s_list" + + have A: "peval A' (push_list a ks) = peval A a" for a + using ks_length push_push_simp unfolding A'_def by auto + + have q: "peval q (push_list a ks) = q'" + unfolding q_def q'_def push_list_def using ks_length unfolding number_of_ex_vars_def by auto + have b: "peval b (push_list a ks) = b'" + unfolding b_def b'_def push_list_def using ks_length unfolding number_of_ex_vars_def by auto + have c: "peval c (push_list a ks) = c'" + unfolding c_def c'_def push_list_def using ks_length unfolding number_of_ex_vars_def by auto + have d: "peval d (push_list a ks) = d'" + unfolding d_def d'_def push_list_def using ks_length unfolding number_of_ex_vars_def by auto + have e: "peval e (push_list a ks) = e'" + unfolding e_def e'_def push_list_def using ks_length unfolding number_of_ex_vars_def by auto + have f: "peval f (push_list a ks) = f'" + unfolding f_def f'_def push_list_def using ks_length unfolding number_of_ex_vars_def by auto + + have r: "(!) (map (\P. peval P (push_list a ks)) r) x = (!) r_list x" if "x < n" for x + unfolding r_def r_list_def using that unfolding push_list_def + using ks_length unfolding number_of_ex_vars_def by auto + + have z: "(map (\P. peval P (push_list a ks)) z) ! x = z_list ! x" if "x < n" for x + unfolding z_def z_list_def using that unfolding push_list_def + using ks_length unfolding number_of_ex_vars_def by (auto simp add: add.commute) + + have s: "(map (\P. peval P (push_list a ks)) s) ! x = s_list ! x" if "x < Suc m" for x + unfolding s_def s_list_def using that unfolding push_list_def + using ks_length unfolding number_of_ex_vars_def by (auto simp add: add.commute) + + have "all_equations (peval A a) q' b' c' d' e' f' r' z' s'" + using ALLEQ unfolding all_equations_relation_def apply (simp add: defs) + unfolding A q b c d e f + using all_equations_invariance[of + "(!) (map (\P. peval P (push_list a ks)) r)" r' + "(!) (map (\P. peval P (push_list a ks)) z)" z' + "(!) (map (\P. peval P (push_list a ks)) s)" s' + "peval A a" q' b' c' d' e' f'] r z s + using r'_def s'_def z'_def by fastforce + + thus "eval DR a" + unfolding DR_def rm_equations_def rm_equations_relation_def by (auto simp: defs) (blast) + next + assume "eval DR a" + then obtain q' b' c' d' e' f' r' z' s' where + all_eq: "all_equations (peval A a) q' b' c' d' e' f' r' z' s'" + unfolding DR_def rm_equations_def rm_equations_relation_def by (auto simp: defs) + + define r_list where "r_list \ map r' [0.. map z' [0.. map s' [0.. [q', b', c', d', e', f'] @ r_list @ z_list @ s_list" + + have "number_of_ex_vars = length ks" + unfolding number_of_ex_vars_def ks_def r_list_def z_list_def s_list_def by auto + + have A: "peval A' (push_list a ks) = peval A a" for a + unfolding A'_def + using push_push_simp[of A ks a] unfolding \number_of_ex_vars = length ks\ by auto + + + have q: "peval q (push_list a ks) = q'" + unfolding q_def ks_def push_list_def by auto + have b: "peval b (push_list a ks) = b'" + unfolding b_def ks_def push_list_def by auto + have c: "peval c (push_list a ks) = c'" + unfolding c_def ks_def push_list_def by auto + have d: "peval d (push_list a ks) = d'" + unfolding d_def ks_def push_list_def by auto + have e: "peval e (push_list a ks) = e'" + unfolding e_def ks_def push_list_def by auto + have f: "peval f (push_list a ks) = f'" + unfolding f_def ks_def push_list_def by auto + + have r: "(map (\P. peval P (push_list a ks)) r) ! x = r' x" if "x < n" for x + using that unfolding ks_def r_list_def r_def push_list_def + using nth_append[of "map r' [0..P. peval P (push_list a ks)) z) ! x = z' x" if "x < n" for x + using that unfolding ks_def z_list_def r_list_def z_def push_list_def apply simp + using nth_append[of "map r' [0..P. peval P (push_list a ks)) s) ! x = s' x" if "x < Suc m" for x + using that unfolding ks_def r_list_def z_list_def s_list_def s_def push_list_def apply simp + using nth_append[of "map r' [0..P. peval P (push_list a ks)) r)" r' + "(!) (map (\P. peval P (push_list a ks)) z)" z' + "(!) (map (\P. peval P (push_list a ks)) s)" s' + "peval A a" q' b' c' d' e' f'] r z s + using r_list_def s_list_def z_list_def by auto + + thus "eval DS a" + unfolding DS_def using \number_of_ex_vars = length ks\ by (auto) + qed + + moreover have "is_dioph_rel DS" + unfolding DS_def + using all_equations_dioph \length r = n\ \length z = n\ \length s = Suc m\ assms + by (auto simp: dioph) + + ultimately show ?thesis + using is_dioph_rel_def by auto + +qed + +end + +end \ No newline at end of file diff --git a/thys/DPRM_Theorem/Machine_Equations/All_Equations_Invariance.thy b/thys/DPRM_Theorem/Machine_Equations/All_Equations_Invariance.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Machine_Equations/All_Equations_Invariance.thy @@ -0,0 +1,140 @@ +subsubsection \Invariance of equations\ + +theory All_Equations_Invariance + imports Register_Equations All_State_Equations Mask_Equations Constants_Equations + +begin + +context register_machine +begin + +definition all_equations where + "all_equations a q b c d e f r z s + \ rm_eq_fixes.register_equations p n a b q r z s + \ rm_eq_fixes.state_equations p b e q z s + \ rm_eq_fixes.mask_equations n c d e f r z + \ rm_eq_fixes.constants_equations b c d e f q + \ rm_eq_fixes.miscellaneous_equations a c q" + +lemma all_equations_invariance: + fixes r z s :: "nat \ nat" + and r' z' s' :: "nat \ nat" + assumes "\iii r i = r' i" for i + using assms by auto + have z: "i z i = z' i" for i + using assms by auto + have s: "i s i = s' i" for i + using assms by auto + + have "length p > 0" using p_nonempty by auto + have "n > 0" using n_gt_0 by auto + + have z_at_modifies: "z (modifies (p ! k)) = z' (modifies (p ! k))" if "k < length p" for k + using z[of "modifies (p!k)"] m_def modifies_yields_valid_register that by auto + + have "rm_eq_fixes.register_equations p n a b q r z s + = rm_eq_fixes.register_equations p n a b q r' z' s'" + proof - + + have sum_radd: "\R+ p d s = \R+ p d s'" for d + by (rule sum_radd_cong, auto simp: s m_def) + + have sum_rsub: "\R- p d (\k. s k && z d) = \R- p d (\k. s' k && z' d)" for d + apply (rule sum_rsub_cong) using s z m_def z_at_modifies \length p > 0\ + by (auto, metis Suc_pred \0 < length p\ le_imp_less_Suc) + + have "rm_eq_fixes.register_0 p a b r z s = rm_eq_fixes.register_0 p a b r' z' s'" + using rm_eq_fixes_def local.register_machine_axioms rm_eq_fixes.register_0_def sum_radd[of 0] + sum_rsub[of 0] using r \n > 0\ by auto + + moreover have "rm_eq_fixes.register_l p n b r z s = rm_eq_fixes.register_l p n b r' z' s'" + using rm_eq_fixes.register_l_def sum_radd sum_rsub rm_eq_fixes_def + local.register_machine_axioms using r \n > 0\ by auto + + moreover have "rm_eq_fixes.register_bound n b q r = rm_eq_fixes.register_bound n b q r'" + using rm_eq_fixes_def local.register_machine_axioms rm_eq_fixes.register_bound_def + using r by auto + + ultimately show ?thesis + using rm_eq_fixes_def local.register_machine_axioms rm_eq_fixes.register_equations_def + by auto + qed + + moreover have "rm_eq_fixes.state_equations p b e q z s + = rm_eq_fixes.state_equations p b e q z' s'" + proof - + have "rm_eq_fixes.state_relations_from_recursion p b e z s + = rm_eq_fixes.state_relations_from_recursion p b e z' s'" + proof - + + have sum_sadd: "\S+ p d s = \S+ p d s'" for d + by (rule sum_sadd_cong, auto simp: s m_def) + + have sum_ssub_nzero: "\S- p d (\k. s k && z (modifies (p ! k))) + = \S- p d (\k. s' k && z' (modifies (p ! k)))" for d + apply (rule sum_ssub_nzero_cong) using z_at_modifies z s + by (metis One_nat_def Suc_pred \0 < length p\ le_imp_less_Suc m_def) + + have sum_ssub_zero: "\S0 p d (\k. s k && e - z (modifies (p ! k))) + = \S0 p d (\k. s' k && e - z' (modifies (p ! k)))" for d + apply (rule sum_ssub_zero_cong) using z_at_modifies z s + by (metis One_nat_def Suc_pred \0 < length p\ le_imp_less_Suc m_def) + + have "rm_eq_fixes.state_0 p b e z s = rm_eq_fixes.state_0 p b e z' s'" + using rm_eq_fixes.state_0_def sum_sadd sum_ssub_nzero sum_ssub_zero + rm_eq_fixes_def local.register_machine_axioms + using s by auto + + moreover have "rm_eq_fixes.state_d p b e z s = rm_eq_fixes.state_d p b e z' s'" + using rm_eq_fixes.state_d_def sum_sadd sum_ssub_nzero sum_ssub_zero + rm_eq_fixes_def local.register_machine_axioms + using s by auto + + ultimately show ?thesis + using rm_eq_fixes_def local.register_machine_axioms + rm_eq_fixes.state_relations_from_recursion_def by auto + qed + + moreover have "rm_eq_fixes.state_unique_equations p b e q s + = rm_eq_fixes.state_unique_equations p b e q s'" + using rm_eq_fixes.state_unique_equations_def + rm_eq_fixes_def local.register_machine_axioms rm_eq_fixes.state_mask_def + rm_eq_fixes.state_bound_def + using s by force + + ultimately show ?thesis + using rm_eq_fixes_def local.register_machine_axioms rm_eq_fixes.state_equations_def + rm_eq_fixes.state_mask_def rm_eq_fixes.state_bound_def rm_eq_fixes.state_m_def + rm_eq_fixes.state_partial_sum_mask_def using s z by auto + qed + + moreover have "rm_eq_fixes.mask_equations n c d e f r z = + rm_eq_fixes.mask_equations n c d e f r' z'" + proof - + have "rm_eq_fixes.register_mask n d r = rm_eq_fixes.register_mask n d r'" + using rm_eq_fixes_def local.register_machine_axioms rm_eq_fixes.register_mask_def r by auto + + moreover have "rm_eq_fixes.zero_indicator_mask n e z = rm_eq_fixes.zero_indicator_mask n e z'" + using rm_eq_fixes.zero_indicator_mask_def rm_eq_fixes_def local.register_machine_axioms z + by auto + + moreover have "rm_eq_fixes.zero_indicator_0_or_1 n c d f r z + = rm_eq_fixes.zero_indicator_0_or_1 n c d f r' z'" + using rm_eq_fixes_def local.register_machine_axioms rm_eq_fixes.zero_indicator_0_or_1_def + using r z by auto + + ultimately show ?thesis + using rm_eq_fixes_def local.register_machine_axioms rm_eq_fixes.mask_equations_def by auto + qed + + ultimately show ?thesis + unfolding all_equations_def by auto +qed + + +end + +end \ No newline at end of file diff --git a/thys/DPRM_Theorem/Machine_Equations/All_State_Equations.thy b/thys/DPRM_Theorem/Machine_Equations/All_State_Equations.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Machine_Equations/All_State_Equations.thy @@ -0,0 +1,142 @@ +subsubsection \Wrap-up: Combining all state equations\ + +theory All_State_Equations imports State_Unique_Equations State_d_Equation + +begin + +text \The remaining equations:\ + +context rm_eq_fixes +begin + + text \Equation 4.27\ + definition state_m :: "bool" where + "state_m \ s m = b ^ q" + + text \Equation not in the book\ + definition state_partial_sum_mask :: "bool" where + "state_partial_sum_mask \ \M\m. (\k\M. s k) \ e" + + text \Wrapping it all up\ + + definition state_equations :: "bool" where + "state_equations \ state_relations_from_recursion \ state_unique_equations + \ state_partial_sum_mask \ state_m" + +end + +context register_machine +begin + +lemma state_m_dioph: + fixes b q :: polynomial + fixes s :: "polynomial list" + assumes "length s = Suc m" + defines "DR \ LARY (\ll. rm_eq_fixes.state_m p (ll!0!0) (ll!0!1) (nth (ll!1))) [[b, q], s]" + shows "is_dioph_rel DR" +proof - + define DS where "DS \ [(s!m) = b ^ q]" + + have "eval DS a = eval DR a" for a + using DS_def DR_def rm_eq_fixes.state_m_def rm_eq_fixes_def local.register_machine_axioms + using assms by (simp add: defs) + + moreover have "is_dioph_rel DS" + unfolding DS_def by (auto simp: dioph) + + ultimately show ?thesis using is_dioph_rel_def by auto +qed + +lemma state_partial_sum_mask_dioph: + fixes e :: polynomial + fixes s :: "polynomial list" + assumes "length s = Suc m" + defines "DR \ LARY (\ll. rm_eq_fixes.state_partial_sum_mask p (ll!0!0) (nth (ll!1))) [[e], s]" + shows "is_dioph_rel DR" +proof - + + define partial_sum_mask where "partial_sum_mask \ (\m. (sum_polynomial (nth s) [0..] e))" + define DS where "DS \ [\j = 0.. peval e a) + = ((\j = 0.. peval e a)" for k + proof - + have "[0..j = 0..j = 0..j = 0..P. peval P a) s)) {..k})" if "k\m" for k + apply (rule sum.cong, auto) + by (metis assms(1) dual_order.strict_trans le_imp_less_Suc nth_map + order.not_eq_order_implies_strict that) + + have "eval DS a = (\kj = 0.. peval e a)" + unfolding DS_def partial_sum_mask_def using aux + by (simp add: defs \length s = Suc m\ sum_polynomial_eval) + + also have "... = (\k\m. + (\j = 0.. peval e a)" + by (simp add: less_Suc_eq_le) + + finally show ?thesis using rm_eq_fixes_def local.register_machine_axioms DR_def + rm_eq_fixes.state_partial_sum_mask_def aux2 by (simp add: defs) + qed + + moreover have "is_dioph_rel DS" + unfolding DS_def partial_sum_mask_def by (auto simp: dioph) + + ultimately show ?thesis using is_dioph_rel_def by auto +qed + +definition state_equations_relation :: "polynomial \ polynomial \ polynomial \ polynomial list + \ polynomial list \ relation" ("[STATE] _ _ _ _ _")where + "[STATE] b e q z s \ LARY (\ll. rm_eq_fixes.state_equations p (ll!0!0) (ll!0!1) (ll!0!2) + (nth (ll!1)) (nth (ll!2))) + [[b, e, q], z, s]" + +lemma state_equations_dioph: + fixes b e q :: polynomial + fixes s z :: "polynomial list" + assumes "length s = Suc m" "length z = n" + defines "DR \ [STATE] b e q z s" + shows "is_dioph_rel DR" +proof - + + define DS where + "DS \ (LARY (\ll. rm_eq_fixes.state_relations_from_recursion p (ll!0!0) (ll!0!1) + (nth (ll!1)) (nth (ll!2))) [[b, e], z, s]) + [\] (LARY (\ll. rm_eq_fixes.state_unique_equations p (ll!0!0) (ll!0!1) (ll!0!2) (nth (ll!1))) + [[b, e, q], s]) + [\] (LARY (\ll. rm_eq_fixes.state_partial_sum_mask p (ll!0!0) (nth (ll!1))) [[e], s]) + [\] (LARY (\ll. rm_eq_fixes.state_m p (ll!0!0) (ll!0!1) (nth (ll!1))) [[b, q], s])" + + have "eval DS a = eval DR a" for a + using DS_def DR_def rm_eq_fixes.state_equations_def + state_equations_relation_def rm_eq_fixes_def local.register_machine_axioms by (auto simp: defs) + + moreover have "is_dioph_rel DS" + unfolding DS_def using assms state_relations_from_recursion_dioph[of z s] state_m_dioph[of s] + state_partial_sum_mask_dioph state_unique_equations_dioph and_dioph + by (auto simp: dioph) + + ultimately show ?thesis using is_dioph_rel_def by auto +qed + +end + +end + diff --git a/thys/DPRM_Theorem/Machine_Equations/Constants_Equations.thy b/thys/DPRM_Theorem/Machine_Equations/Constants_Equations.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Machine_Equations/Constants_Equations.thy @@ -0,0 +1,137 @@ +subsubsection \Equations for arithmetization constants\ + +theory Constants_Equations imports Equation_Setup "../Register_Machine/MachineMasking" + "../Diophantine/Binary_And" + +begin + +context rm_eq_fixes +begin + + text \Equation 4.14\ + definition constant_b :: "bool" where + "constant_b \ b = B c" + + text \Equation 4.16\ + definition constant_d :: "bool" where + "constant_d \ d = D q c b" + + text \Equation 4.18\ + definition constant_e :: "bool" where + "constant_e \ e = E q b" + + text \Equation 4.21\ + definition constant_f :: "bool" where + "constant_f \ f = F q c b" + + text \Equation not in the book\ + definition c_gt_0 :: "bool" where + "c_gt_0 \ c > 0" + + text \Equation 4.26\ + definition a_bound :: "bool" where + "a_bound \ a < 2 ^ c" + + text \Equation not in the book\ + definition q_gt_0 :: "bool" where + "q_gt_0 \ q > 0" + + + definition constants_equations :: "bool" where + "constants_equations \ constant_b \ constant_d \ constant_e \ constant_f" + + definition miscellaneous_equations :: "bool" where + "miscellaneous_equations \ c_gt_0 \ a_bound \ q_gt_0 " + +end + +context register_machine +begin + +definition rm_constant_equations :: + "polynomial \ polynomial \ polynomial \ polynomial \ polynomial \ polynomial \ relation" + ("[CONST] _ _ _ _ _ _") where + "[CONST] b c d e f q \ NARY (\l. rm_eq_fixes.constants_equations + (l!0) (l!1) (l!2) (l!3) (l!4) (l!5)) [b, c, d, e, f, q]" + +definition rm_miscellaneous_equations :: + "polynomial \ polynomial \ polynomial \ relation" + ("[MISC] _ _ _") where + "[MISC] c a q \ NARY (\l. rm_eq_fixes.miscellaneous_equations + (l!0) (l!1) (l!2)) [c, a, q]" + +lemma rm_constant_equations_dioph: + fixes b c d e f q + defines "DR \ [CONST] b c d e f q" + shows "is_dioph_rel DR" +proof- + have fx: "rm_eq_fixes p n" + using rm_eq_fixes_def local.register_machine_axioms by auto + + define b' c' d' e' f' q' where pushed_defs: + "b' = (push_param b 2)" "c' = (push_param c 2)" "d' = (push_param d 2)" + "e' = (push_param e 2)" "f' = (push_param f 2)" "q' = (push_param q 2)" + + define s t where params_def: "s = Param 0" "t = Param 1" + + (* using equations (4.31) - (4.33) *) + define DS1 where "DS1 \ [b' = Const 2 ^ (c' [+] \<^bold>1)] [\] + [s = Const 2 ^ c'] [\] [t = b' ^ (q' [+] \<^bold>1)] [\] + (b' [-] \<^bold>1) [*] d' [=] (s [-] \<^bold>1) [*] (t [-] \<^bold>1)" + + define DS2 where "DS2 \ (b' [-] \<^bold>1) [*] e' [=] t [-] \<^bold>1 [\] + (b' [-] \<^bold>1) [*] f' [=] s [*] (t [-] \<^bold>1)" + + define DS where "DS \ [\2] DS1 [\] DS2" + + have "eval DS a = eval DR a" for a + unfolding DR_def DS_def DS1_def DS2_def rm_constant_equations_def defs + apply (auto simp add: fx rm_eq_fixes.constants_equations_def[of p n]) + unfolding pushed_defs params_def push_push apply (auto simp add: push_list_eval) + apply (auto simp add: fx rm_eq_fixes.constant_b_def[of p n] B_def + rm_eq_fixes.constant_d_def[of p n] rm_eq_fixes.constant_e_def[of p n] + rm_eq_fixes.constant_f_def[of p n]) + using d_geom_series[of "2 * 2 ^ peval c a" "peval c a" "(peval q a)" " peval d a"] + using e_geom_series[of "(2 * 2 ^ peval c a)" "peval q a" "peval e a"] + using f_geom_series[of "2 * 2 ^ peval c a" "peval c a" "(peval q a)" " peval f a"] + apply (auto) + apply (rule exI[of _ "[2 ^ peval c a, peval b a * peval b a ^ peval q a]"]) + using push_list_def push_push by auto + + + moreover have "is_dioph_rel DS" unfolding DS_def DS1_def DS2_def by (simp add: dioph) + + ultimately show ?thesis + by (simp add: is_dioph_rel_def) +qed + +lemma rm_miscellaneous_equations_dioph: + fixes c a q + defines "DR \ [MISC] a c q" + shows "is_dioph_rel DR" +proof- + define c' a' q' where pushed_defs: + "c' == (push_param c 1)" "a' == (push_param a 1)" "q' = (push_param q 1)" + + define DS where "DS \ [\] c' [>] \<^bold>0 + [\] [(Param 0) = (Const 2) ^ c'] [\] a'[<] Param 0 + [\] q' [>] \<^bold>0" + + have "eval DS a = eval DR a" for a unfolding DS_def defs DR_def + using rm_miscellaneous_equations_def + rm_eq_fixes.miscellaneous_equations_def rm_eq_fixes.c_gt_0_def rm_eq_fixes.a_bound_def + rm_eq_fixes.q_gt_0_def rm_eq_fixes_def local.register_machine_axioms apply auto + unfolding pushed_defs push_push1 + apply (auto, rule exI[of _ "2 ^ peval c a"]) unfolding push0 by auto + + + moreover have "is_dioph_rel DS" unfolding DS_def by (simp add: dioph) + + ultimately show ?thesis + by (simp add: is_dioph_rel_def) +qed + + +end + +end \ No newline at end of file diff --git a/thys/DPRM_Theorem/Machine_Equations/Equation_Setup.thy b/thys/DPRM_Theorem/Machine_Equations/Equation_Setup.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Machine_Equations/Equation_Setup.thy @@ -0,0 +1,38 @@ +subsection \Arithmetizing equations are Diophantine\ + +theory Equation_Setup imports "../Register_Machine/RegisterMachineSpecification" + "../Diophantine/Diophantine_Relations" + +begin + +locale register_machine = + fixes p :: program + and n :: nat + assumes p_nonempty: "length p > 0" + and valid_program: "program_register_check p n" + assumes n_gt_0: "n > 0" + +begin + + definition m :: "nat" where + "m \ length p - 1" + + lemma modifies_yields_valid_register: + assumes "k < length p" + shows "modifies (p!k) < n" + proof - + have "instruction_register_check n (p!k)" + using valid_program assms list_all_length program_register_check.simps by auto + + thus ?thesis by (cases "p!k", auto simp: n_gt_0) + qed + +end + +locale rm_eq_fixes = register_machine + + fixes a b c d e f :: "nat" + and q :: nat + and r z :: "register \ nat" + and s :: "state \ nat" + +end \ No newline at end of file diff --git a/thys/DPRM_Theorem/Machine_Equations/Machine_Equation_Equivalence.thy b/thys/DPRM_Theorem/Machine_Equations/Machine_Equation_Equivalence.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Machine_Equations/Machine_Equation_Equivalence.thy @@ -0,0 +1,257 @@ +subsection \Equivalence of register machine and arithmetizing equations\ + +theory Machine_Equation_Equivalence imports All_Equations + "../Register_Machine/MachineEquations" + "../Register_Machine/MultipleToSingleSteps" + +begin + +context register_machine +begin + +lemma conclusion_4_5: + assumes is_val: "is_valid_initial ic p a" + and n_def: "n \ length (snd ic)" + shows "(\q. terminates ic p q) = rm_equations a" +proof (rule) + assume "\q. terminates ic p q" + then obtain q::nat where terminates: "terminates ic p q" by auto + hence "q>0" using terminates_def by auto + + have "\c>1. cells_bounded ic p c" + using terminate_c_exists terminates is_val is_valid_initial_def by blast + then obtain c where c: "cells_bounded ic p c \ c > 1" by auto + + define b where "b \ B c" + define d where "d \ D q c b" + define e where "e \ E q b" + define f where "f \ F q c b" + + have "c>1" using c by auto + + have "b>1" using c b_def B_def + using nat_neq_iff by fastforce + + define r where "r \ RLe ic p b q" + define s where "s \ SKe ic p b q" + define z where "z \ ZLe ic p b q" + + interpret equations: rm_eq_fixes p n a b c d e f q r z s by unfold_locales + + have "equations.mask_equations" + proof - + have "\l d" + using lm04_15_register_masking[of "ic" "p" "c" _ "q"] r_def n_def d_def b_def c by auto + moreover have "\l e" + using lm04_15_zero_masking z_def n_def e_def b_def c by auto + moreover have "\lR+ p 0 s - b * \R- p 0 (\k. s k && z 0)" + using lm04_23_multiple_register1[of "ic" "p" "a" "c" "0" "q"] is_val c terminates `q>0` r_def + s_def z_def b_def bitAND_commutes by auto + moreover have "\l>0. l < n \ r l = b * r l + b * \R+ p l s - b * \R- p l (\k. s k && z l)" + using lm04_22_multiple_register[of "ic" "p" "a" "c" _ "q"] + b_def c terminates r_def s_def z_def is_val bitAND_commutes n_def `q>0` by auto + moreover have "l r l < b^q" for l + proof - + assume "l 2 ^ Suc c - Suc 0" using `c>1` by auto + have "\t. R ic p l t < 2 ^ c" using c `lt. R ic p l t < 2 ^ Suc c - Suc 0" using c_ineq + by (metis dual_order.strict_trans linorder_neqE_nat not_less) + have "(\t = 0..q. b ^ t * R ic p l t) = (\t = 0..(Suc (q-1)). b ^ t * R ic p l t)" + using `q>0` by auto + also have "... = (\t = 0..q-1. b ^ t * R ic p l t) + b^q * R ic p l q" + using Set_Interval.comm_monoid_add_class.sum.atLeast0_atMost_Suc[of _ "q-1"] `q>0` by auto + also have "... = (\t = 0..q-1. b ^ t * R ic p l t)" using Rlq by auto + also have "... < b ^ q" using b_def R_bound + base_summation_bound[of "R ic p l" "c" "q-1"] `q>0` by (auto simp: mult.commute) + finally show ?thesis using r_def RLe_def by auto + qed + ultimately show ?thesis unfolding equations.register_equations_def equations.register_0_def + equations.register_l_def equations.register_bound_def by auto + qed + + moreover have "equations.state_equations" + proof - + have "equations.state_relations_from_recursion" + proof - + have "\d>0. d\m \ s d = b*\S+ p d (\k. s k) + b*\S- p d (\k. s k && z (modifies (p!k))) + + b*\S0 p d (\k. s k && (e - z (modifies (p!k))))" + apply (auto simp: s_def z_def) + using lm04_24_multiple_step_states[of "ic" "p" "a" "c" _ "q"] + b_def c terminates s_def z_def is_val bitAND_commutes m_def `q>0` e_def E_def by auto + moreover have "s 0 = 1 + b*\S+ p 0 (\k. s k) + b*\S- p 0 (\k. s k && z (modifies (p!k))) + + b*\S0 p 0 (\k. s k && (e - z (modifies (p!k))))" + using lm04_25_multiple_step_state1[of "ic" "p" "a" "c" _ "q"] + b_def c terminates s_def z_def is_val bitAND_commutes m_def `q>0` e_def E_def by auto + ultimately show ?thesis unfolding equations.state_relations_from_recursion_def + equations.state_0_def equations.state_d_def equations.state_m_def by auto + qed + + moreover have "equations.state_unique_equations" + proof - + have "k s k < b ^ q" for k + using state_q_bound is_val terminates \q>0\ b_def s_def m_def c by auto + moreover have "k\m \ s k \ e" for k + using state_mask is_val terminates \q>0\ b_def e_def s_def c by auto + ultimately show ?thesis unfolding equations.state_unique_equations_def + equations.state_mask_def equations.state_bound_def by auto + qed + + moreover have "\M\m. sum s {..M} \ e" + using state_sum_mask is_val terminates \q>0\ b_def e_def s_def c `b>1` m_def by auto + + moreover have "s m = b^q" + using halting_condition_04_27[of "ic" "p" "a" "q" "c"] m_def b_def is_val `q>0` terminates + s_def by auto + + ultimately show ?thesis unfolding equations.state_equations_def + equations.state_partial_sum_mask_def equations.state_m_def by auto + qed + + moreover have "equations.constants_equations" + unfolding equations.constants_equations_def equations.constant_b_def + equations.constant_d_def equations.constant_e_def equations.constant_f_def + using b_def d_def e_def f_def by auto + + moreover have "equations.miscellaneous_equations" + proof - + have tapelength: "length (snd ic) > 0" + using is_val is_valid_initial_def[of "ic" "p" "a"] by auto + have "R ic p 0 0 = a" using is_val is_valid_initial_def[of "ic" "p" "a"] + R_def List.hd_conv_nth[of "snd ic"] by auto + moreover have "R ic p 0 0 < 2^c" using c tapelength by auto + ultimately have "a < 2^c" by auto + thus ?thesis unfolding equations.miscellaneous_equations_def equations.c_gt_0_def + equations.a_bound_def equations.q_gt_0_def + using \q > 0\ \c > 1\ by auto + qed + + ultimately show "rm_equations a" unfolding rm_equations_def all_equations_def by blast +next + assume "rm_equations a" + + then obtain q b c d e f r z s where + reg: "rm_eq_fixes.register_equations p n a b q r z s" and + state: "rm_eq_fixes.state_equations p b e q z s" and + mask: "rm_eq_fixes.mask_equations n c d e f r z" and + const: "rm_eq_fixes.constants_equations b c d e f q" and + misc: "rm_eq_fixes.miscellaneous_equations a c q" + unfolding rm_equations_def all_equations_def by auto + + have fx: "rm_eq_fixes p n" + unfolding rm_eq_fixes_def using local.register_machine_axioms by auto + + have "q>0" using misc fx rm_eq_fixes.miscellaneous_equations_def + rm_eq_fixes.q_gt_0_def by auto + have "b>1" using B_def const rm_eq_fixes.constants_equations_def + rm_eq_fixes.constant_b_def fx + by (metis One_nat_def Zero_not_Suc less_one n_not_Suc_n nat_neq_iff nat_power_eq_Suc_0_iff + numeral_2_eq_2 of_nat_0 of_nat_power_eq_of_nat_cancel_iff of_nat_zero_less_power_iff pos2) + have "n>0" using is_val is_valid_initial_def[of "ic" "p" "a"] n_def by auto + have "m>0" using m_def is_val is_valid_initial_def[of "ic" "p"] is_valid_def[of "ic" "p"] by auto + + define Seq where "Seq \ (\k t. nth_digit (s k) t b)" + define Req where "Req \ (\l t. nth_digit (r l) t b)" + define Zeq where "Zeq \ (\l t. nth_digit (z l) t b)" + + (* Quick and dirty: :\| *) + have mask_old: "mask_equations n r z c d e f" and + reg_old: "reg_equations p r z s b a (length (snd ic)) q" and + state_old: "state_equations p s z b e q (length p - 1)" and + const_old: "rm_constants q c b d e f a" + subgoal + using mask rm_eq_fixes.mask_equations_def rm_eq_fixes.register_mask_def fx + mask_equations_def rm_eq_fixes.zero_indicator_0_or_1_def rm_eq_fixes.zero_indicator_mask_def + by simp + subgoal + using reg state mask const misc using rm_eq_fixes.register_equations_def + rm_eq_fixes.register_0_def rm_eq_fixes.register_l_def rm_eq_fixes.register_bound_def + reg_equations_def n_def fx by simp + subgoal + using state fx state_equations_def rm_eq_fixes.state_equations_def + rm_eq_fixes.state_relations_from_recursion_def rm_eq_fixes.state_0_def rm_eq_fixes.state_m_def + rm_eq_fixes.state_d_def rm_eq_fixes.state_unique_equations_def rm_eq_fixes.state_mask_def + rm_eq_fixes.state_bound_def rm_eq_fixes.state_partial_sum_mask_def m_def by simp + subgoal unfolding rm_constants_def + using const misc fx rm_eq_fixes.constants_equations_def + rm_eq_fixes.miscellaneous_equations_def rm_eq_fixes.constant_b_def rm_eq_fixes.constant_d_def + rm_eq_fixes.constant_e_def rm_eq_fixes.constant_f_def rm_eq_fixes.c_gt_0_def + rm_eq_fixes.q_gt_0_def rm_eq_fixes.a_bound_def by simp + done + + hence RZS_eq: "l j\m \ t\q \ + R ic p l t = Req l t \ Z ic p l t = Zeq l t \ S ic p j t = Seq j t" for l j t + using rzs_eq[of "m" "p" "n" "ic" "a" "r" "z"] mask_old reg_old state_old const_old + m_def n_def is_val `q>0` Seq_def Req_def Zeq_def by auto + + have R_eq: "l t\q \ R ic p l t = Req l t" for l t using RZS_eq by auto + have Z_eq: "l t\q \ Z ic p l t = Zeq l t" for l t using RZS_eq by auto + have S_eq: "j\m \ t\q \ S ic p j t = Seq j t" for j t using RZS_eq[of "0"] `n>0` by auto + + have "ishalt (p!m)" using m_def is_val + is_valid_initial_def[of "ic" "p" "a"] is_valid_def[of "ic" "p"] by auto + have "Seq m q = 1" using state nth_digit_def Seq_def `b>1` + using fx rm_eq_fixes.state_equations_def + rm_eq_fixes.state_relations_from_recursion_def + rm_eq_fixes.state_m_def by auto + hence "S ic p m q = 1" using S_eq by auto + hence "fst (steps ic p q) = m" using S_def by(cases "fst (steps ic p q) = m"; auto) + hence qhalt: "ishalt (p ! (fst (steps ic p q)))" using S_def `ishalt (p!m)` by auto + + hence rempty: "snd (steps ic p q) ! l = 0" if "l < n" for l + unfolding R_def[symmetric] + using R_eq[of l q] \l < n\ apply auto + using reg Req_def nth_digit_def + using rm_eq_fixes.register_equations_def + rm_eq_fixes.register_l_def + rm_eq_fixes.register_0_def + rm_eq_fixes.register_bound_def + by (smt Euclidean_Division.div_eq_0_iff mod_if mult_0_right mult_mod_right nat_mult_1_right + zero_less_one fx) + + have state_m_0: "t S ic p m t = 0" for t + proof - + assume "t1 < b\ \t < q\ less_imp_le not_one_le_zero power_diff) + also have "... mod b = 0" using \1 < b\ \t < q\ by simp + finally have mod: "b^q div b^t mod b = 0" by auto + have "s m = b^q" using state fx rm_eq_fixes.state_equations_def + rm_eq_fixes.state_m_def + rm_eq_fixes.state_relations_from_recursion_def by auto + hence "Seq m t = 0" using Seq_def nth_digit_def mod by auto + with S_eq `t < q` show ?thesis by auto + qed + have "\k ishalt (p!k)" + using is_val is_valid_initial_def[of "ic" "p" "a"] is_valid_def[of "ic" "p"] m_def by auto + moreover have "t fst (steps ic p t) < length p - 1" for t + proof (rule ccontr) + assume asm: "\ (t < q \ fst (steps ic p t) < length p - 1)" + hence "t length p - 1" by auto + moreover have "fst (steps ic p t) \ length p - 1" + using p_contains[of "ic" "p" "a" "t"] is_val by auto + ultimately have "fst (steps ic p t) = m" using m_def by auto + hence "S ic p m t = 1" using S_def by auto + thus "False" using state_m_0[of "t"] `t \ ishalt (p ! (fst (steps ic p t)))" for t using m_def by auto + hence no_early_halt: "t \ ishalt (p ! (fst (steps ic p t)))" for t using state_m_0 by auto + + have "correct_halt ic p q" using qhalt rempty correct_halt_def n_def by auto + thus "(\q. terminates ic p q)" using no_early_halt terminates_def `q>0` by auto +qed + +end + +end \ No newline at end of file diff --git a/thys/DPRM_Theorem/Machine_Equations/Mask_Equations.thy b/thys/DPRM_Theorem/Machine_Equations/Mask_Equations.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Machine_Equations/Mask_Equations.thy @@ -0,0 +1,159 @@ +subsubsection \Equations for masking relations\ + +theory Mask_Equations + imports "../Register_Machine/MachineMasking" Equation_Setup "../Diophantine/Binary_And" + +abbrevs mb = "\" + +begin + +context rm_eq_fixes +begin + + text \Equation 4.15\ + definition register_mask :: "bool" where + "register_mask \ \l < n. r l \ d" + + text \Equation 4.17\ + definition zero_indicator_mask :: "bool" where + "zero_indicator_mask \ \l < n. z l \ e" + + text \Equation 4.20\ + definition zero_indicator_0_or_1 :: "bool" where + "zero_indicator_0_or_1 \ \l register_mask \ zero_indicator_mask \ zero_indicator_0_or_1" + +end + +context register_machine +begin + +lemma register_mask_dioph: + fixes d r + assumes "n = length r" + defines "DR \ (NARY (\l. rm_eq_fixes.register_mask n (l!0) (shift l 1)) ([d] @ r))" + shows "is_dioph_rel DR" +proof - + define DS where "DS \ [\i. ((r!i) [\] d))" + + have "eval DS a = eval DR a" for a + proof - + have "eval DR a = rm_eq_fixes.register_mask n (peval d a) (list_eval r a)" + unfolding DR_def by (auto simp add: shift_def list_eval_def) + also have "... = (\l < n. (peval (r!l) a) \ peval d a)" + using rm_eq_fixes.register_mask_def \n = length r\ rm_eq_fixes_def + local.register_machine_axioms by (auto simp: list_eval_def) + finally show ?thesis + unfolding DS_def defs by simp + qed + + moreover have "is_dioph_rel DS" + unfolding DS_def by (auto simp add: dioph) + + ultimately show ?thesis + by (simp add: is_dioph_rel_def) +qed + +lemma zero_indicator_mask_dioph: + fixes e z + assumes "n = length z" + defines "DR \ (NARY (\l. rm_eq_fixes.zero_indicator_mask n (l!0) (shift l 1)) ([e] @ z))" + shows "is_dioph_rel DR" +proof - + define DS where "DS \ [\i. ((z!i) [\] e))" + + have "eval DS a = eval DR a" for a + proof - + have "eval DR a = rm_eq_fixes.zero_indicator_mask n (peval e a) (list_eval z a)" + unfolding DR_def by (auto simp add: shift_def list_eval_def) + also have "... = (\l < n. (peval (z!l) a) \ peval e a)" + using rm_eq_fixes.zero_indicator_mask_def \n = length z\ + rm_eq_fixes_def local.register_machine_axioms by (auto simp: list_eval_def) + finally show ?thesis + unfolding DS_def defs by simp + qed + + moreover have "is_dioph_rel DS" + unfolding DS_def by (auto simp add: dioph) + + ultimately show ?thesis + by (simp add: is_dioph_rel_def) +qed + +lemma zero_indicator_0_or_1_dioph: + fixes c d f r z + assumes "n = length r" and "n = length z" + defines "DR \ LARY (\ll. rm_eq_fixes.zero_indicator_0_or_1 n (ll!0!0) (ll!0!1) (ll!0!2) + (nth (ll!1)) (nth (ll!2))) [[c, d, f], r, z]" + shows "is_dioph_rel DR" +proof - + let ?N = 2 + define c' d' f' r' z' where pushed_def: "c' = push_param c ?N" "d' = push_param d ?N" + "f' = push_param f ?N" "r' = map (\x. push_param x ?N) r" + "z' = map (\x. push_param x ?N) z" + define DS where "DS \ [\i. ([\2] [Param 0 = (Const 2) ^ c'] + [\] [Param 1 = (r'!i) [+] d' && f'] + [\] Param 0 [*] (z'!i) [=] Param 1))" + + have "eval DS a = eval DR a" for a + proof - + have "eval DR a = rm_eq_fixes.zero_indicator_0_or_1 n (peval c a) (peval d a) (peval f a) + (list_eval r a) (list_eval z a)" + unfolding DR_def defs by (auto simp add: assms shift_def list_eval_def) + also have "... = (\l < n. 2^(peval c a) * (peval (z!l) a) + = (peval (r!l) a + peval d a) && peval f a)" + using rm_eq_fixes.zero_indicator_0_or_1_def \n = length r\ using assms + rm_eq_fixes_def local.register_machine_axioms by (auto simp: list_eval_def) + finally show ?thesis + unfolding DS_def defs pushed_def using push_push apply (auto) + subgoal for k + apply (rule exI[of _ "[2^peval c a, peval (r!k) a + peval d a && peval f a]"]) + apply (auto simp: push_list_def assms(1-2)) + by (metis assms(1) assms(2) length_Cons list.size(3) nth_map numeral_2_eq_2) + subgoal + using assms by auto + done + qed + + moreover have "is_dioph_rel DS" + unfolding DS_def by (auto simp add: dioph) + + ultimately show ?thesis + by (simp add: is_dioph_rel_def) +qed + + +definition mask_equations_relation ("[MASK] _ _ _ _ _ _") where + "[MASK] c d e f r z \ LARY (\ll. rm_eq_fixes.mask_equations n + (ll!0!0) (ll!0!1) (ll!0!2) (ll!0!3) (nth (ll!1)) (nth (ll!2))) + [[c, d, e, f], r, z]" + +lemma mask_equations_relation_dioph: + fixes c d e f r z + assumes "n = length r" and "n = length z" + defines "DR \ [MASK] c d e f r z" + shows "is_dioph_rel DR" +proof - + define DS where "DS \ NARY (\l. rm_eq_fixes.register_mask n (l!0) (shift l 1)) ([d] @ r) + [\] NARY (\l. rm_eq_fixes.zero_indicator_mask n (l!0) (shift l 1)) ([e] @ z) + [\] LARY (\ll. rm_eq_fixes.zero_indicator_0_or_1 n (ll!0!0) (ll!0!1) (ll!0!2) + (nth (ll!1)) (nth (ll!2))) [[c, d, f], r, z]" + + have "eval DS a = eval DR a" for a + using DS_def DR_def mask_equations_relation_def rm_eq_fixes.mask_equations_def + rm_eq_fixes_def local.register_machine_axioms by (simp add: defs shift_def) + + moreover have "is_dioph_rel DS" + unfolding DS_def using assms dioph + using register_mask_dioph zero_indicator_mask_dioph zero_indicator_0_or_1_dioph + by (metis (no_types, lifting)) + + ultimately show ?thesis + by (simp add: is_dioph_rel_def) +qed + +end + +end \ No newline at end of file diff --git a/thys/DPRM_Theorem/Machine_Equations/RM_Sums_Diophantine.thy b/thys/DPRM_Theorem/Machine_Equations/RM_Sums_Diophantine.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Machine_Equations/RM_Sums_Diophantine.thy @@ -0,0 +1,110 @@ +theory RM_Sums_Diophantine imports Equation_Setup "../Diophantine/Register_Machine_Sums" + "../Diophantine/Binary_And" + +begin + +context register_machine +begin + +definition sum_ssub_nzero_of_bit_and :: "polynomial \ nat \ polynomial list \ polynomial list + \ relation" + ("[_ = \S- _ '(_ && _')]") where + "[x = \S- d (s && z)] \ let x' = push_param x (length p); + s' = push_param_list s (length p); + z' = push_param_list z (length p) + in [\length p] [\i. [Param i = s'!i && z'!i]) + [\] x' [=] ([\S-] p d Param)" + +lemma sum_ssub_nzero_of_bit_and_dioph[dioph]: + fixes s z :: "polynomial list" and d :: nat and x + shows "is_dioph_rel [x = \S- d (s && z)]" + unfolding sum_ssub_nzero_of_bit_and_def by (auto simp add: dioph) + +lemma sum_rsub_nzero_of_bit_and_eval: + fixes z s :: "polynomial list" and d :: nat and x :: polynomial + assumes "length s = Suc m" "length z = Suc m" "length p > 0" + shows "eval [x = \S- d (s && z)] a + \ peval x a = \S- p d (\k. peval (s!k) a && peval (z!k) a)" (is "?P \ ?Q") +proof - + have invariance: "\k \S- p d y1 = \S- p d y2" for y1 y2 + unfolding sum_ssub_nzero.simps apply (intro sum.cong, simp) + using \length p > 0\ by auto (metis Suc_pred le_imp_less_Suc length_greater_0_conv) + + have len_ps: "length s = length p" + using m_def \length s = Suc m\ \length p > 0\ by auto + have len_pz: "length z = length p" + using m_def \length z = Suc m\ \length p > 0\ by auto + + show ?thesis + proof (rule) + assume ?P + thus ?Q + using sum_ssub_nzero_of_bit_and_def \length p > 0\ apply (auto simp add: defs push_push) + using push_push_map_i apply (simp add: push_param_list_def len_ps len_pz) + unfolding list_eval_def apply (auto simp: assms len_ps len_pz invariance) + apply (rule sum_ssub_nzero_cong) apply auto + by (metis (no_types, lifting) One_nat_def assms(1) assms(2) + le_imp_less_Suc len_ps m_def nth_map) + + next + assume ?Q + thus ?P + using sum_ssub_nzero_of_bit_and_def \length p > 0\ apply (auto simp add: defs push_push) + apply (rule exI[of _ "map (\k. peval (s ! k) a && peval (z ! k) a) [0.. nat \ polynomial list \ polynomial list + \ relation" + ("[_ = \S0 _ '(_ && _')]") where + "[x = \S0 d (s && z)] \ let x' = push_param x (length p); + s' = push_param_list s (length p); + z' = push_param_list z (length p) + in [\length p] [\i. [Param i = s'!i && z'!i]) + [\] x' [=] [\S0] p d Param" + +lemma sum_ssub_zero_of_bit_and_dioph[dioph]: + fixes s z :: "polynomial list" and d :: nat and x + shows "is_dioph_rel [x = \S0 d (s && z)]" + unfolding sum_ssub_zero_of_bit_and_def by (auto simp add: dioph) + +lemma sum_rsub_zero_of_bit_and_eval: + fixes z s :: "polynomial list" and d :: nat and x :: polynomial + assumes "length s = Suc m" "length z = Suc m" "length p > 0" + shows "eval [x = \S0 d (s && z)] a + \ peval x a = \S0 p d (\k. peval (s!k) a && peval (z!k) a)" (is "?P \ ?Q") +proof - + have invariance: "\k \S0 p d y1 = \S0 p d y2" for y1 y2 + unfolding sum_ssub_zero.simps apply (intro sum.cong, simp) + using \length p > 0\ by auto (metis Suc_pred le_imp_less_Suc length_greater_0_conv) + + have len_ps: "length s = length p" + using m_def \length s = Suc m\ \length p > 0\ by auto + have len_pz: "length z = length p" + using m_def \length z = Suc m\ \length p > 0\ by auto + + show ?thesis + proof (rule) + assume ?P + thus ?Q + using sum_ssub_zero_of_bit_and_def \length p > 0\ apply (auto simp add: defs push_push) + using push_push_map_i apply (simp add: push_param_list_def len_ps len_pz) + unfolding list_eval_def apply (auto simp: assms len_ps len_pz invariance) + apply (rule sum_ssub_zero_cong) apply auto + by (metis (no_types, lifting) One_nat_def assms(1) assms(2) + le_imp_less_Suc len_ps m_def nth_map) + next + assume ?Q + thus ?P + using sum_ssub_zero_of_bit_and_def \length p > 0\ apply (auto simp add: defs push_push) + apply (rule exI[of _ "map (\k. peval (s ! k) a && peval (z!k) a) [0..Register Equations\ + +theory Register_Equations imports "../Register_Machine/MultipleStepRegister" + Equation_Setup "../Diophantine/Register_Machine_Sums" + "../Diophantine/Binary_And" "HOL-Library.Rewrite" + +begin + +context rm_eq_fixes +begin + + text \Equation 4.22\ + definition register_0 :: "bool" where + "register_0 \ r 0 = a + b*r 0 + b*\R+ p 0 s - b*\R- p 0 (\k. s k && z 0)" + + text \Equation 4.23\ + definition register_l :: "bool" where + "register_l \ \l>0. l < n \ r l = b*r l + b*\R+ p l s - b*\R- p l (\k. s k && z l)" + + text \Extra equation not in Matiyasevich's book\ + definition register_bound :: "bool" where + "register_bound \ \l < n. r l < b ^ q" + + definition register_equations :: "bool" where + "register_equations \ register_0 \ register_l \ register_bound" + +end + +context register_machine +begin + +definition sum_rsub_of_bit_and :: "polynomial \ nat \ polynomial list \ polynomial + \ relation" + ("[_ = \R- _ '(_ && _')]") where + "[x = \R- d (s && zl)] \ let x' = push_param x (length p); + s' = push_param_list s (length p); + zl' = push_param zl (length p) + in [\length p] [\i. [Param i = s'!i && zl']) + [\] x' [=] [\R-] p d Param" + +lemma sum_rsub_of_bit_and_dioph[dioph]: + fixes s :: "polynomial list" and d :: nat and x zl :: polynomial + shows "is_dioph_rel [x = \R- d (s && zl)]" + unfolding sum_rsub_of_bit_and_def by (auto simp add: dioph) + +lemma sum_rsub_of_bit_and_eval: + fixes z s :: "polynomial list" and d :: nat and x :: polynomial + assumes "length s = Suc m" "length p > 0" + shows "eval [x = \R- d (s && zl)] a + \ peval x a = \R- p d (\k. peval (s!k) a && peval zl a)" (is "?P \ ?Q") +proof - + have invariance: "\k \R- p d y1 = \R- p d y2" for y1 y2 + unfolding sum_rsub.simps apply (intro sum.cong, simp) + using \length p > 0\ by auto (metis Suc_pred le_imp_less_Suc length_greater_0_conv) + + have len_ps: "length s = length p" + using m_def \length s = Suc m\ \length p > 0\ by auto + + have aux1: "peval ([\R-] p l f) a = \R- p l (\x. peval (f x) a) " for l f + using defs \length p > 0\ by auto + + show ?thesis + proof (rule) + assume ?P + thus ?Q + unfolding sum_rsub_of_bit_and_def + using aux1 apply simp + apply (auto simp add: aux1 push_push defs) + using push_push_map_i apply (simp add: push_param_list_def len_ps) + unfolding list_eval_def apply (simp add: assms len_ps invariance) + using assms(2) invariance len_ps sum_rsub_polynomial_eval by force + next + assume ?Q + thus ?P + unfolding sum_rsub_of_bit_and_def apply (auto simp add: aux1 defs push_push) + apply (rule exI[of _ "map (\k. peval (s ! k) a && peval zl a) [0..length p > 0\ defs by simp + qed +qed + + +lemma register_0_dioph[dioph]: + fixes A b :: polynomial + fixes r z s :: "polynomial list" + assumes "length r = n" "length z = n" "length s = Suc m" + defines "DR \ LARY (\ll. rm_eq_fixes.register_0 p (ll!0!0) (ll!0!1) + (nth (ll!1)) (nth (ll!2)) (nth (ll!3))) [[A, b], r, z, s]" + shows "is_dioph_rel DR" +proof - + let ?N = "1" + define A' b' r' z' s' where pushed_def: "A' = push_param A ?N" "b' = push_param b ?N" + "r' = map (\x. push_param x ?N) r" "z' = map (\x. push_param x ?N) z" + "s' = map (\x. push_param x ?N) s" + + define DS where "DS \ [\] ([Param 0 = \R- 0 (s' && (z'!0))] [\] + r'!0 [=] A' [+] b' [*] r'!0 [+] b' [*] ([\R+] p 0 (nth s')) + [-] b' [*] (Param 0))" + + have "length p > 0" using p_nonempty by auto + have "n > 0" using n_gt_0 by auto + + have "length p = length s" + using \length s = Suc m\ m_def \length p > 0\ by auto + have "length s' = length s" + unfolding pushed_def by auto + have "length z > 0" + using \length z = n\ \n > 0\ by simp + have "length r > 0" + using \length r = n\ \n > 0\ by simp + + have "eval DS a = eval DR a" for a + proof - + (* the key to this proof is to write these intermediate steps with list_eval on the RHS + because that is needed in the final proof; otherwise showing the equivalence again + re-requires unfolding the sum definitions *) + have sum_radd_push: "\R+ p 0 (\x. peval (s' ! x) (push a k)) = \R+ p 0 (list_eval s a)" for k + unfolding sum_radd.simps pushed_def apply (intro sum.cong, simp) + using push_push_map1 \length p = length s\ \length s = Suc m\ by simp + + have sum_rsub_push: "\R- p 0 (\x. peval (s' ! x) (push a k) && peval (z' ! 0) (push a k)) + = \R- p 0 (\x. list_eval s a x && peval (z ! 0) a)" for k + unfolding sum_rsub.simps pushed_def apply (intro sum.cong, simp) + using push_push_map1 \length p = length s\ \length s = Suc m\ \length z > 0\ + by (simp add: list_eval_def) + + have 1: "peval ([\R-] p l f) a = \R- p l (\x. peval (f x) a) " for f l + using defs \length p > 0\ by auto + + show ?thesis + unfolding DS_def rm_eq_fixes.register_0_def + register_machine_axioms rm_eq_fixes_def apply (simp add: defs) + using \length p > 0\ apply (simp add: sum_rsub_of_bit_and_eval \length s' = length s\ + \length s = Suc m\) + apply (simp add: sum_radd_push sum_rsub_push) + unfolding pushed_def using push_push1 push_push_map1 \length r > 0\ apply simp + unfolding DR_def assms defs \length p > 0\ + using rm_eq_fixes_def rm_eq_fixes.register_0_def register_machine_axioms apply (simp) + using \length z > 0\ push_def list_eval_def 1 apply (simp add: 1 defs \length p > 0\) + using One_nat_def sum_radd_push unfolding pushed_def(5) list_eval_def by presburger + qed + + moreover have "is_dioph_rel DS" + unfolding DS_def by (simp add: dioph) + + ultimately show ?thesis + by (auto simp: is_dioph_rel_def) +qed + +lemma register_l_dioph[dioph]: + fixes b :: polynomial + fixes r z s :: "polynomial list" + assumes "length r = n" "length z = n" "length s = Suc m" + defines "DR \ LARY (\ll. rm_eq_fixes.register_l p n (ll!0!0) + (nth (ll!1)) (nth (ll!2)) (nth (ll!3))) [[b], r, z, s]" + shows "is_dioph_rel DR" +proof - + define indices where "indices \ [Suc 0..x. push_param x ?N) r" + "z' = map (\x. push_param x ?N) z" + "s' = map (\x. push_param x ?N) s" + + define param_l_is_sum_rsub_of_bitand where + "param_l_is_sum_rsub_of_bitand \ \l. [Param l = \R- l (s' && (z'!l))]" + define params_are_sum_rsub_of_bitand where + "params_are_sum_rsub_of_bitand \ [\ in indices] param_l_is_sum_rsub_of_bitand" + define single_register where + "single_register \ \l. r'!l [=] b' [*] r'!l [+] b' [*] ([\R+] p l (nth s')) [-] b' [*] (Param l)" + + define DS where "DS \ [\n] params_are_sum_rsub_of_bitand [\] [\ in indices] single_register" + + have "length p > 0" using p_nonempty by auto + have "n > 0" using n_gt_0 by auto + have "length p = length s" + using \length s = Suc m\ m_def \length p > 0\ by auto + have "length s' = length s" + unfolding pushed_def by auto + have "length z > 0" + using \length z = n\ \n > 0\ by simp + have "length r > 0" + using \length r = n\ \n > 0\ by simp + have "length indices + 1 = n" + unfolding indices_def \n>0\ + using Suc_pred' \n > 0\ length_upt by presburger + have "length s' = Suc m" + using \length s' = length s\ \length s = Suc m\ by auto + + have "eval DS a = eval DR a" for a + proof - + + have eval_to_peval: + "eval [polynomial.Param (indices ! k) + = \R- indices ! k (s' && z' ! (indices ! k))] y + \(peval (polynomial.Param (indices ! k)) y + = \R- p (indices ! k) (\ka. peval (s' ! ka) y && peval (z' ! (indices ! k)) y) )" for k y + using sum_rsub_of_bit_and_eval \length p > 0\ \length s' = Suc m\ by auto + + have b'_unfold: "peval b' (push_list a ks) = peval b a" if "length ks = n" for ks + unfolding pushed_def using indices_def push_push that \length indices + 1 = n\ by auto + + have r'_unfold: "peval (r' ! (indices ! k)) (push_list a ks) = peval (r!(indices!k)) a" + if "k < length indices" and "length ks = n" for k ks + using indices_def push_push pushed_def that(1) that(2) \length r = n\ by auto + + have Param_unfold: "peval (Param (indices ! k)) (push_list a ks) = ks!(indices!k)" + if "k < length indices" and "length ks = n" for k ks + using One_nat_def Suc_pred indices_def length_upt nat_add_left_cancel_less + nth_upt peval.simps(2) plus_1_eq_Suc push_list_eval that(1) that(2) by (metis \0 < n\) + + have unfold_4: "push_list a ks (indices ! k) = ks!(indices!k)" + if "k < length indices" and "length ks = n" for k ks + using Param_unfold that(1) that(2) by force + + have unfold_sum_radd: "\R+ p (indices ! k) (\x. peval (s' ! x) (push_list a ks)) + = \R+ p (indices ! k) (list_eval s a)" + if "length ks = n" for k ks + apply (rule sum_radd_cong) unfolding pushed_def + using push_push_map_i[of ks n _ s a] \length indices + 1 = n\ that + using \length p = length s\ + by (metis \0 < length p\ add.left_neutral add_lessD1 le_neq_implies_less less_add_one + less_diff_conv less_diff_conv2 nat_le_linear not_add_less1) + + have unfold_sum_rsub: "\R- p (indices ! k) (\ka. peval (s' ! ka) (push_list a ks) + && peval (z' ! (indices ! k)) (push_list a ks)) + = \R- p (indices ! k) (\ka. list_eval s a ka + && peval (z ! (indices ! k)) a)" + if "length ks = n" for k ks + apply (rule sum_rsub_cong) unfolding pushed_def + using push_push_map_i[of ks n _ s a] unfolding \length indices + 1 = n\ + using \length p = length s\ assms apply simp + using nth_map[of _ z "\x. push_param x (Suc (length indices))"] + using modifies_yields_valid_register \length z = n\ + by (smt assms le_imp_less_Suc nth_map push_push_simp that) + + have indices_unfold: "(\k < length indices. P (indices!k)) \ (\l>0. l P l)" for P + unfolding indices_def apply auto + using \n>0\ by (metis Suc_diff_Suc diff_zero not_less_eq) + + have alternative_sum_rsub: + "(\R- p l (\ka. list_eval s a ka && peval (z ! l) a)) + =(\R- p l (\k. map (\P. peval P a) s ! k && map (\P. peval P a) z ! l))" for l + apply (rule sum_rsub_cong) unfolding list_eval_def apply simp + using modifies_yields_valid_register + One_nat_def assms(3) nth_map \length z = n\ \length s = Suc m\ + by (metis \length p = length s\ le_imp_less_Suc m_def) + + (* Start of chain of equalities *) + have "(eval DS a) = (\ks. n = length ks \ + (\kR- (indices ! k) (s' && z' ! (indices ! k))] (push_list a ks)) \ + (\kks. n = length ks \ + (\kR- p (indices ! k) (\ka. peval (s' ! ka) (push_list a ks) + && peval (z' ! (indices ! k)) (push_list a ks)) \ + peval (r' ! (indices ! k)) (push_list a ks) + = peval b' (push_list a ks) * peval (r' ! (indices ! k)) (push_list a ks) + + peval b' (push_list a ks) * \R+ p (indices ! k) + (\x. peval (s' ! x) (push_list a ks)) + - peval b' (push_list a ks) * (push_list a ks (indices ! k))))" + using eval_to_peval unfolding single_register_def + using sum_radd_polynomial_eval \length p > 0\ by (simp add: defs) (blast) + + also have "... = (\ks. n = length ks \ + (\kR- p (indices ! k) (\ka. peval (s' ! ka) (push_list a ks) + && peval (z' ! (indices ! k)) (push_list a ks)) \ + peval (r!(indices!k)) a + = peval b a * peval (r!(indices!k)) a + + peval b a * \R+ p (indices ! k) (\x. peval (s' ! x) (push_list a ks)) + - peval b a * (ks!(indices!k))))" + using b'_unfold r'_unfold Param_unfold unfold_4 by (smt (z3)) + + also have "... = (\ks. n = length ks \ + (\kR- p (indices ! k) (\ka. peval (s' ! ka) (push_list a ks) + && peval (z' ! (indices ! k)) (push_list a ks))) \ + peval (r!(indices!k)) a + = peval b a * peval (r!(indices!k)) a + + peval b a * (\R+ p (indices ! k) (list_eval s a)) + - peval b a * (ks!(indices!k))))" + using unfold_sum_radd by (smt (z3)) + + also have "... = (\ks. n = length ks \ + (\kR- p (indices ! k) (\ka. list_eval s a ka && peval (z ! (indices ! k)) a) + \ peval (r!(indices!k)) a + = peval b a * peval (r!(indices!k)) a + + peval b a * (\R+ p (indices ! k) (list_eval s a)) + - peval b a * (ks!(indices!k))))" + using unfold_sum_rsub by auto + + also have "... = (\ks. n = length ks \ + (\kR- p (indices ! k) (\ka. list_eval s a ka && peval (z ! (indices ! k)) a) + \ peval (r!(indices!k)) a + = peval b a * peval (r!(indices!k)) a + + peval b a * (\R+ p (indices ! k) (list_eval s a)) + - peval b a * + (\R- p (indices ! k) (\ka. list_eval s a ka && peval (z ! (indices ! k)) a))))" + by smt + + also have "... = (\kR+ p (indices ! k) (list_eval s a)) + - peval b a * + (\R- p (indices ! k) (\ka. list_eval s a ka && peval (z ! (indices ! k)) a)))" + unfolding indices_def apply auto + apply (rule exI[of _ + "map (\k. \R- p k (\ka. list_eval s a ka && peval (z ! k) a)) [0..l>0. l < n \ + peval (r!l) a + = peval b a * peval (r!l) a + + peval b a * (\R+ p l (list_eval s a)) + - peval b a * + (\R- p l (\ka. list_eval s a ka && peval (z ! l) a)))" + using indices_unfold[of "\x. peval (r ! x) a = + peval b a * peval (r ! x) a + peval b a * (\R+ p x (list_eval s a)) - + peval b a * (\R- p x (\ka. (list_eval s a ka) && peval (z ! x) a))"] + by auto + + also have "... = (\l>0. l < n \ + peval (r!l) a = + peval b a * map (\P. peval P a) r ! l + + peval b a * (\R+ p l ((!) (map (\P. peval P a) s))) + - peval b a * (\R- p l (\k. map (\P. peval P a) s ! k && map (\P. peval P a) z ! l)))" + using nth_map[of _ r "(\P. peval P a)"] unfolding \length r = n\ + using alternative_sum_rsub list_eval_def by auto + + also have "... = (eval DR a)" + apply (simp add: DR_def defs) using rm_eq_fixes_def rm_eq_fixes.register_l_def + local.register_machine_axioms + using nth_map[of _ r "\P. peval P a"] unfolding \length r = n\ by auto + + finally show "eval DS a = eval DR a" by auto + qed + + moreover have "is_dioph_rel DS" + proof - + have "list_all (is_dioph_rel \ param_l_is_sum_rsub_of_bitand) indices" + unfolding param_l_is_sum_rsub_of_bitand_def indices_def list_all_def by (auto simp:dioph) + hence "is_dioph_rel params_are_sum_rsub_of_bitand" + unfolding params_are_sum_rsub_of_bitand_def by (auto simp: dioph) + + have "list_all (is_dioph_rel \ single_register) indices" + unfolding single_register_def list_all_def indices_def by (auto simp: dioph) + thus ?thesis + unfolding DS_def using \is_dioph_rel params_are_sum_rsub_of_bitand\ by (auto simp: dioph) + qed + + ultimately show ?thesis + by (auto simp: is_dioph_rel_def) +qed + +lemma register_bound_dioph: + fixes b q :: polynomial + fixes r :: "polynomial list" + assumes "length r = n" + defines "DR \ LARY (\ll. rm_eq_fixes.register_bound n (ll!0!0) (ll!0!1) (nth (ll!1))) + [[b, q], r]" + shows "is_dioph_rel DR" +proof - + + define indices where "indices \ [0..x. push_param x ?N) r" + + define bound where + "bound \ \l. (r'!l [<] (Param l) [\] [Param l = b' ^ q'])" + + define DS where "DS \ [\n] [\ in indices] bound" + + have "eval DS a = eval DR a" for a + proof - + + have r'_unfold: "peval (r' ! k) (push_list a ks) = peval (r ! k) a" + if "length ks = n" and "k < length ks" for k ks + unfolding pushed_def \length indices = n\ + using push_push_map_i[of ks n k r] that \length r = n\ list_eval_def by auto + + have b'_unfold: "peval b' (push_list a ks) = peval b a" + and q'_unfold: "peval q' (push_list a ks) = peval q a" + if "length ks = n" and "k < length ks" for k ks + unfolding pushed_def \length indices = n\ + using push_push_simp that \length r = n\ list_eval_def by auto + + have "eval DS a = (\ks. n = length ks \ + (\k + push_list a ks k = peval b' (push_list a ks) ^ peval q' (push_list a ks)))" + unfolding DS_def indices_def bound_def by (simp add: defs) + + also have "... = (\ks. n = length ks \ + (\k + push_list a ks k = peval b a ^ peval q a))" + using r'_unfold b'_unfold q'_unfold by (metis (full_types)) + + also have "... = (\kk. peval b a ^ peval q a) [0..lP. peval P a) r ! l < peval b a ^ peval q a)" + using nth_map[of _ r "\P. peval P a"] \length r = n\ by force + + finally show ?thesis unfolding DR_def + using rm_eq_fixes.register_bound_def rm_eq_fixes_def register_machine_def + p_nonempty n_gt_0 valid_program by (auto simp add: defs) + + qed + + moreover have "is_dioph_rel DS" + proof - + have "list_all (is_dioph_rel \ bound) indices" + unfolding bound_def indices_def list_all_def by (auto simp:dioph) + thus ?thesis unfolding DS_def indices_def bound_def by (auto simp: dioph) + qed + + ultimately show ?thesis + by (auto simp: is_dioph_rel_def) +qed + + + +definition register_equations_relation :: "polynomial \ polynomial \ polynomial + \ polynomial list \ polynomial list \ polynomial list \ relation" ("[REG] _ _ _ _ _ _") where + "[REG] a b q r z s \ LARY (\ll. rm_eq_fixes.register_equations p n (ll!0!0) (ll!0!1) (ll!0!2) + (nth (ll!1)) (nth (ll!2)) (nth (ll!3))) [[a, b, q], r, z, s]" + +lemma reg_dioph: + fixes A b q r z s + assumes "length r = n" "length z = n" "length s = Suc m" + defines "DR \ [REG] A b q r z s" + shows "is_dioph_rel DR" +proof - + + define DS where "DS \ (LARY (\ll. rm_eq_fixes.register_0 p (ll!0!0) (ll!0!1) + (nth (ll!1)) (nth (ll!2)) (nth (ll!3))) [[A, b], r, z, s]) + [\] (LARY (\ll. rm_eq_fixes.register_l p n (ll!0!0) + (nth (ll!1)) (nth (ll!2)) (nth (ll!3))) [[b], r, z, s]) + [\] (LARY (\ll. rm_eq_fixes.register_bound n (ll!0!0) (ll!0!1) (nth (ll!1))) + [[b, q], r])" + + have "eval DS a = eval DR a" for a + unfolding DS_def DR_def register_equations_relation_def rm_eq_fixes.register_equations_def + apply (simp add: defs) + by (simp add: register_machine_axioms rm_eq_fixes.intro rm_eq_fixes.register_equations_def) + + moreover have "is_dioph_rel DS" + unfolding DS_def using assms register_0_dioph[of r z s] register_l_dioph[of r z s] + register_bound_dioph by (auto simp: dioph) + + ultimately show ?thesis by (auto simp: is_dioph_rel_def) +qed + +end + +end \ No newline at end of file diff --git a/thys/DPRM_Theorem/Machine_Equations/State_0_Equation.thy b/thys/DPRM_Theorem/Machine_Equations/State_0_Equation.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Machine_Equations/State_0_Equation.thy @@ -0,0 +1,313 @@ +subsubsection \State 0 equation\ + +theory State_0_Equation imports "../Register_Machine/MultipleStepState" + RM_Sums_Diophantine "../Diophantine/Binary_And" + +begin + +context rm_eq_fixes +begin + + text \Equation 4.24\ + definition state_0 :: "bool" where + "state_0 \ s 0 = 1 + b*\S+ p 0 s + b*\S- p 0 (\k. s k && z (modifies (p!k))) + + b*\S0 p 0 (\k. s k && (e - z (modifies (p!k))))" + +end + +context register_machine +begin + +no_notation ppolynomial.Sum (infixl "\<^bold>+" 65) +no_notation ppolynomial.NatDiff (infixl "\<^bold>-" 65) +no_notation ppolynomial.Prod (infixl "\<^bold>*" 70) + + +lemma state_0_dioph: + fixes b e :: polynomial + fixes z s :: "polynomial list" + assumes "length z = n" "length s = Suc m" + defines "DR \ LARY (\ll. rm_eq_fixes.state_0 p (ll!0!0) (ll!0!1) + (nth (ll!1)) (nth (ll!2))) [[b, e], z, s]" + shows "is_dioph_rel DR" +proof - + let ?N = "2" + define b' e' z' s' where pushed_def: "b' = push_param b ?N" + "e' = push_param e ?N" + "z' = map (\x. push_param x ?N) z" + "s' = map (\x. push_param x ?N) s" + + define z0 z1 where z_def: "z0 \ map (\i. z' ! modifies (p!i)) [0.. map (\i. e' [-] z' ! modifies (p!i)) [0.. [Param 0 = \S- 0 (s' && z0)]" + + define param_1_is_sum_sub_zero_term where + "param_1_is_sum_sub_zero_term \ [Param 1 = \S0 0 (s' && z1)]" + + define step_relation where + "step_relation \ (s'!0 [=] \<^bold>1 [+] b' [*] ([\S+] p 0 (nth s')) + [+] b' [*] Param 0 [+] b' [*] Param 1)" + + define DS where "DS \ [\?N] step_relation + [\] param_0_is_sum_sub_nzero_term + [\] param_1_is_sum_sub_zero_term" + + have "p \ []" using p_nonempty by auto + have ps_lengths: "length p = length s" + using \length s = Suc m\ m_def \p \ []\ by auto + have s_len: "length s > 0" + using ps_lengths \p \ []\ by auto + have p_len: "length p > 0" + using ps_lengths s_len by auto + have p_len2: "length p = Suc m" + using ps_lengths \length s = Suc m\ by auto + have len_s': "length s' = Suc m" + unfolding pushed_def using \length s = Suc m\ by auto + have "length z0 = Suc m" + unfolding z_def ps_lengths \length s = Suc m\ by simp + have "length z1 = Suc m" + unfolding z_def ps_lengths \length s = Suc m\ by simp + + have modifies_le_n: "k < length p \ modifies (p!k) < n" for k + using modifies_yields_valid_register \length z = n\ by auto + + have "eval DS a = eval DR a" for a + proof - + + have b'_unfold: "peval b' (push_list a ks) = peval b a" if "length ks = 2" for ks + using push_push_simp unfolding pushed_def using that by metis + + have s'_0_unfold: "peval (s' ! 0) (push_list a ks) = peval (s ! 0) a" if "length ks = 2" for ks + unfolding pushed_def using push_push_map_i[of ks 2 0 s a] that unfolding list_eval_def + \length s > 0\ using s_len by auto + + have sum_nzero_unfold: + "eval [polynomial.Param 0 = \S- 0 (s' && z0)] (push_list a ks) + = (peval (polynomial.Param 0) (push_list a ks) + = \S- p 0 (\k. peval (s' ! k) (push_list a ks) && peval (z0 ! k) (push_list a ks)))" for ks + using sum_rsub_nzero_of_bit_and_eval[of s' z0 "Param 0" 0 "push_list a ks"] + \length p > 0\ \length s' = Suc m\ \length z0 = Suc m\ by auto + + have sum_zero_unfold: + "eval [polynomial.Param 1 = \S0 0 (s' && z1)] (push_list a ks) + = (peval (polynomial.Param 1) (push_list a ks) + = \S0 p 0 (\k. peval (s' ! k) (push_list a ks) && peval (z1 ! k) (push_list a ks)))" for ks + using sum_rsub_zero_of_bit_and_eval[of s' z1 "Param 1" 0 "push_list a ks"] + \length p > 0\ \length s' = Suc m\ \length z1 = Suc m\ by auto + + have param_0_unfold: "peval (Param 0) (push_list a ks) = ks ! 0" if "length ks = 2" for ks + unfolding push_list_def using that by auto + + have param_1_unfold: "peval (Param 1) (push_list a ks) = ks ! 1" if "length ks = 2" for ks + unfolding push_list_def using that by auto + + have sum_sadd_unfold: + "peval ([\S+] p 0 ((!) s')) (push_list a ks) = \S+ p 0 (\x. peval (s ! x) a)" + if "length ks = 2" for ks + using sum_sadd_polynomial_eval \length p > 0\ apply auto + apply (rule sum_sadd_cong, auto) + unfolding pushed_def using push_push_map_i[of ks 2 _ s a] that + unfolding \length p = length s\ list_eval_def + by (smt One_nat_def assms le_imp_less_Suc m_def nth_map p_len2) + + have z0_unfold: + "peval (s' ! k) (push_list a ks) && peval (z0 ! k) (push_list a ks) + = peval (s ! k) a && peval (z ! modifies (p ! k)) a" + if "length ks = 2" and "k < length p" for k ks + proof - + have map: "map (\i. z' ! modifies (p ! i)) [0..i. z' ! modifies (p ! i)"] + using \k < length p\ by auto + + have s: "peval (map (\x. push_param x 2) s ! k) (push_list a ks) = peval (s ! k) a" + using push_push_map_i[of ks 2 k s] that nth_map[of k s] + unfolding \length s = Suc m\ \length p = Suc m\ list_eval_def by auto + + have z: "peval (map (\x. push_param x 2) z ! modifies (p ! k)) (push_list a ks) + = peval (z ! modifies (p ! k)) a" + using push_push_map_i[of ks 2 "modifies (p!k)" z a] modifies_le_n[of k] that nth_map[of _ z] + unfolding \length z = n\ list_eval_def by auto + + show ?thesis + unfolding z_def map unfolding pushed_def s z by auto + qed + + have z1_unfold: + "peval (s' ! k) (push_list a ks) && peval (z1 ! k) (push_list a ks) + = peval (s ! k) a && (peval e a - peval (z ! modifies (p ! k)) a)" + if "length ks = 2" and "k < length p" for k ks + proof - + have map: + "map (\i. e' [-] (z' ! (modifies (p ! i)))) [0..i. z' ! modifies (p ! i)"] + using \k < length p\ by auto + + have s: "peval (map (\x. push_param x 2) s ! k) (push_list a ks) = peval (s ! k) a" + using push_push_map_i[of ks 2 k s] that nth_map[of k s] + unfolding \length s = Suc m\ \length p = Suc m\ list_eval_def by auto + + have z: "peval (push_param e 2) (push_list a ks) + - peval (map (\x. push_param x 2) z ! modifies (p ! k)) (push_list a ks) + = peval e a - peval (z ! (modifies (p!k))) a" + using push_push_simp[of e ks a] unfolding \length ks = 2\ apply simp + using push_push_map_i[of ks 2 "modifies (p!k)" z a] modifies_le_n[of k] that + nth_map[of "modifies (p!k)" z "(\x. peval x a)"] + unfolding \length z = n\ list_eval_def by auto + + show ?thesis + unfolding z_def map unfolding pushed_def s using z by auto + qed + + have z0sum_unfold: + "(\S- p 0 (\k. peval (s' ! k) (push_list a ks) && peval (z0 ! k) (push_list a ks))) + =(\S- p 0 (\k. peval (s ! k) a && peval (z ! modifies (p ! k)) a))" + if "length ks = 2" for ks + apply (rule sum_ssub_nzero_cong) using z0_unfold[of ks] that + by (metis \length s = Suc m\ le_imp_less_Suc m_def ps_lengths) + + have z1sum_unfold: + "(\S0 p 0 (\k. peval (s' ! k) (push_list a ks) && peval (z1 ! k) (push_list a ks))) + =(\S0 p 0 (\k. peval (s ! k) a && peval e a - peval (z ! modifies (p ! k)) a))" + if "length ks = 2" for ks + apply (rule sum_ssub_zero_cong) using z1_unfold[of ks] that + by (metis \length s = Suc m\ le_imp_less_Suc m_def ps_lengths) + + have sum_sadd_map: "\S+ p 0 ((!) (map (\P. peval P a) s)) = \S+ p 0 (\x. peval (s ! x) a)" + apply (rule sum_sadd_cong, auto) + using nth_map[of _ s "(\P. peval P a)"] m_def \length s = Suc m\ by auto + + have sum_ssub_nzero_map: + "(\S- p 0 (\k. peval (s ! k) a && peval (z ! modifies (p ! k)) a)) + = (\S- p 0 (\k. map (\P. peval P a) s ! k && map (\P. peval P a) z ! modifies (p ! k)))" + proof - + have 1: "peval (s ! k) a && peval (z ! modifies (p ! k)) a = + map (\P. peval P a) s ! k && map (\P. peval P a) z ! modifies (p ! k) " + if "k < length p" for k + proof - + have "peval (s ! k) a = map (\P. peval P a) s ! k" + using nth_map that ps_lengths by auto + moreover have "peval (z ! modifies (p ! k)) a + = map (\P. peval P a) z ! modifies (p ! k) " + using nth_map[of "modifies (p!k)" z "(\P. peval P a)"] modifies_le_n[of k] that + using \length z = n\ by auto + ultimately show ?thesis by auto + qed + show ?thesis apply (rule sum_ssub_nzero_cong, auto) + using 1 by (metis Suc_le_mono Suc_pred less_eq_Suc_le p_len) + qed + + have sum_ssub_zero_map: + "(\S0 p 0 (\k. peval (s ! k) a && peval e a - peval (z ! modifies (p ! k)) a)) + = (\S0 p 0 (\k. map (\P. peval P a) s ! k && peval e a + - map (\P. peval P a) z ! modifies (p ! k)))" + proof - + have 1: "peval (s ! k) a && peval e a - peval (z ! modifies (p ! k)) a = + map (\P. peval P a) s ! k && peval e a - map (\P. peval P a) z ! modifies (p ! k) " + if "k < length p" for k + proof - + have "peval (s ! k) a = map (\P. peval P a) s ! k" + using nth_map that ps_lengths by auto + moreover have "peval (z ! modifies (p ! k)) a + = map (\P. peval P a) z ! modifies (p ! k) " + using nth_map[of "modifies (p!k)" z "(\P. peval P a)"] modifies_le_n[of k] that + using \length z = n\ by auto + ultimately show ?thesis by auto + qed + show ?thesis apply (rule sum_ssub_zero_cong, auto) + using 1 by (metis Suc_le_mono Suc_pred less_eq_Suc_le p_len) + qed + + + have "eval DS a = + (\ks. length ks = 2 \ + eval (s' ! 0 [=] \<^bold>1 [+] b' [*] ([\S+] p 0 (!) s') [+] b' [*] Param 0 + [+] b' [*] Param (Suc 0)) (push_list a ks) + \ eval [polynomial.Param 0 = \S- 0 (s' && z0)] (push_list a ks) + \ eval [polynomial.Param 1 = \S0 0 (s' && z1)] (push_list a ks))" + unfolding DS_def step_relation_def param_0_is_sum_sub_nzero_term_def + param_1_is_sum_sub_zero_term_def by (simp add: defs) + + also have "... = (\ks. length ks = 2 \ + peval (s' ! 0) (push_list a ks) = + Suc (peval b' (push_list a ks) * peval ([\S+] p 0 ((!) s')) (push_list a ks) + + peval b' (push_list a ks) * push_list a ks 0 + + peval b' (push_list a ks) * push_list a ks (Suc 0)) + \ (peval (Param 0) (push_list a ks) + = \S- p 0 (\k. peval (s' ! k) (push_list a ks) && peval (z0 ! k) (push_list a ks))) + \ (peval (Param 1) (push_list a ks) + = \S0 p 0 (\k. peval (s' ! k) (push_list a ks) && peval (z1 ! k) (push_list a ks))))" + unfolding sum_nzero_unfold sum_zero_unfold by (simp add: defs ) + + also have "... = (\ks. length ks = 2 \ + peval (s ! 0) a = + Suc (peval b a * peval ([\S+] p 0 ((!) s')) (push_list a ks) + + peval b a * push_list a ks 0 + + peval b a * push_list a ks (Suc 0)) + \ (ks!0 + = \S- p 0 (\k. peval (s' ! k) (push_list a ks) && peval (z0 ! k) (push_list a ks))) + \ (ks!1 + = \S0 p 0 (\k. peval (s' ! k) (push_list a ks) && peval (z1 ! k) (push_list a ks))))" + using b'_unfold s'_0_unfold param_0_unfold param_1_unfold by auto + + also have "... = (\ks. length ks = 2 \ + peval (s ! 0) a = + Suc (peval b a * \S+ p 0 (\x. peval (s ! x) a) + + peval b a * (ks!0) + peval b a * (ks!1)) + \ (ks!0 = \S- p 0 (\k. peval (s' ! k) (push_list a ks) && peval (z0 ! k) (push_list a ks))) + \ (ks!1 = \S0 p 0 (\k. peval (s' ! k) (push_list a ks) && peval (z1 ! k) (push_list a ks))))" + using push_list_def sum_sadd_unfold by auto + + also have "... = (\ks. length ks = 2 \ + peval (s ! 0) a = Suc (peval b a * \S+ p 0 (\x. peval (s ! x) a) + + peval b a * (ks!0) + peval b a * (ks!1)) + \ (ks!0 = \S- p 0 (\k. peval (s ! k) a && peval (z ! modifies (p ! k)) a)) + \ (ks!1 = \S0 p 0 (\k. peval (s ! k) a && peval e a - peval (z ! modifies (p ! k)) a)))" + using z0sum_unfold z1sum_unfold by auto + + also have "... = (\ks. length ks = 2 \ + peval (s ! 0) a + = Suc (peval b a * \S+ p 0 (\x. peval (s ! x) a) + + peval b a * \S- p 0 (\k. peval (s ! k) a && peval (z ! modifies (p ! k)) a) + + peval b a * \S0 p 0 (\k. peval (s ! k) a && peval e a - peval (z ! modifies (p ! k)) a)) + \ (ks!0 = \S- p 0 (\k. peval (s ! k) a && peval (z ! modifies (p ! k)) a)) + \ (ks!1 = \S0 p 0 (\k. peval (s ! k) a && peval e a - peval (z ! modifies (p ! k)) a)))" + by auto + + also have "... = (peval (s ! 0) a + = Suc (peval b a * \S+ p 0 (\x. peval (s ! x) a) + + peval b a * \S- p 0 (\k. peval (s ! k) a && peval (z ! modifies (p ! k)) a) + + peval b a * \S0 p 0 (\k. peval (s ! k) a && peval e a - peval (z ! modifies (p ! k)) a)))" + apply auto + apply (rule exI[of _ "[(\S- p 0 (\k. peval (s ! k) a && peval (z ! modifies (p ! k)) a)), + \S0 p 0 (\k. peval (s ! k) a && peval e a - peval (z ! modifies (p ! k)) a) ]"]) + by auto + + also have "... = (map (\P. peval P a) s ! 0 = + Suc (peval b a * \S+ p 0 ((!) (map (\P. peval P a) s)) + + peval b a * \S- p 0 (\k. map (\P. peval P a) s ! k + && map (\P. peval P a) z ! modifies (p ! k)) + + peval b a * + \S0 p 0 (\k. map (\P. peval P a) s ! k && peval e a + - map (\P. peval P a) z ! modifies (p ! k)) ))" + using nth_map[of _ _ "(\P. peval P a)"] \length s > 0\ + using sum_ssub_zero_map sum_sadd_map sum_ssub_nzero_map by auto + + finally show ?thesis unfolding DR_def using rm_eq_fixes_def local.register_machine_axioms + rm_eq_fixes.state_0_def by (simp add: defs) + qed + + moreover have "is_dioph_rel DS" + unfolding DS_def param_1_is_sum_sub_zero_term_def param_0_is_sum_sub_nzero_term_def + step_relation_def by (auto simp add: dioph) + + ultimately show ?thesis + by (simp add: is_dioph_rel_def) +qed + +end + +end \ No newline at end of file diff --git a/thys/DPRM_Theorem/Machine_Equations/State_Unique_Equations.thy b/thys/DPRM_Theorem/Machine_Equations/State_Unique_Equations.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Machine_Equations/State_Unique_Equations.thy @@ -0,0 +1,156 @@ +subsubsection \State unique equations\ + +theory State_Unique_Equations imports "../Register_Machine/MultipleStepState" + Equation_Setup "../Diophantine/Register_Machine_Sums" + "../Diophantine/Binary_And" + +begin + +context rm_eq_fixes +begin + + text \Equations not in the book: \ + definition state_mask :: "bool" where + "state_mask \ \k\m. s k \ e" + + definition state_bound :: "bool" where + "state_bound \ \k— s m,t = 1 at t = q and nowhere else*) + + definition state_unique_equations :: "bool" where + "state_unique_equations \ state_mask \ state_bound" + +end + +context register_machine +begin + +lemma state_mask_dioph: + fixes e :: polynomial + fixes s :: "polynomial list" + assumes "length s = Suc m" + defines "DR \ LARY (\ll. rm_eq_fixes.state_mask p (ll!0!0) (nth (ll!1))) [[e], s]" + shows "is_dioph_rel DR" +proof - + define mask where "mask \ (\l. (s!l [\] e))" + define DS where "DS \ [\k\m. peval (s ! k) a \ peval e a)" + unfolding DS_def mask_def by (simp add: less_Suc_eq_le defs) + + also have "... = (\k\m. map (\P. peval P a) s ! k \ peval e a)" + using nth_map[of _ s "(\P. peval P a)"] \length s = Suc m\ by auto + + finally show ?thesis + unfolding DR_def using rm_eq_fixes_def local.register_machine_axioms + rm_eq_fixes.state_mask_def by (simp add: defs) + qed + + moreover have "is_dioph_rel DS" + proof - + have "list_all (is_dioph_rel \ mask) [0.. LARY (\ll. rm_eq_fixes.state_bound p (ll!0!0) (ll!0!1) (nth (ll!1))) [[b, q], s]" + shows "is_dioph_rel DR" +proof - + let ?N = "m" + define b' q' s' where pushed_def: "b' = push_param b ?N" + "q' = push_param q ?N" + "s' = map (\x. push_param x ?N) s" + + define bound where + "bound \ \l. s'!l [<] (Param l) [\] [Param l = b' ^ q']" + + define DS where "DS \ [\m] [\length s = Suc m\ list_eval_def + by (metis less_SucI nth_map push_push) + + have b'_unfold: "peval b' (push_list a ks) = peval b a" + and q'_unfold: "peval q' (push_list a ks) = peval q a" + if "length ks = m" and "k < length ks" for k ks + unfolding pushed_def + using push_push_simp that \length s = Suc m\ list_eval_def by metis+ + + have "eval DS a = (\ks. m = length ks \ + (\k + push_list a ks k = peval b' (push_list a ks) ^ peval q' (push_list a ks)))" + unfolding DS_def bound_def by (simp add: defs) + + also have "... = (\ks. m = length ks \ + (\k + push_list a ks k = peval b a ^ peval q a))" + using s'_unfold b'_unfold q'_unfold by metis + + also have "... = (\kk. peval b a ^ peval q a) [0..lP. peval P a) s ! l < peval b a ^ peval q a)" + using nth_map[of _ s "\P. peval P a"] \length s = Suc m\ by force + + finally show ?thesis unfolding DR_def + using rm_eq_fixes_def local.register_machine_axioms rm_eq_fixes.state_bound_def + by (simp add: defs) + + qed + + moreover have "is_dioph_rel DS" + proof - + have "list_all (is_dioph_rel \ bound) [0.. LARY + (\ll. rm_eq_fixes.state_unique_equations p (ll!0!0) (ll!0!1) (ll!0!2) (nth (ll!1))) + [[b, e, q], s]" + shows "is_dioph_rel DR" +proof - + + define DS where "DS \ LARY (\ll. rm_eq_fixes.state_mask p (ll!0!0) (nth (ll!1))) [[e], s] + [\] LARY (\ll. rm_eq_fixes.state_bound p (ll!0!0) (ll!0!1) (nth (ll!1))) + [[b, q], s]" + + have "eval DS a = eval DR a" for a + unfolding DR_def DS_def using rm_eq_fixes.state_unique_equations_def rm_eq_fixes_def + local.register_machine_axioms + by (auto simp: defs) + + moreover have "is_dioph_rel DS" + unfolding DS_def using state_bound_dioph state_mask_dioph assms dioph by auto + + ultimately show ?thesis using is_dioph_rel_def by auto +qed + +end + +end \ No newline at end of file diff --git a/thys/DPRM_Theorem/Machine_Equations/State_d_Equation.thy b/thys/DPRM_Theorem/Machine_Equations/State_d_Equation.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Machine_Equations/State_d_Equation.thy @@ -0,0 +1,433 @@ +subsubsection \State d equation\ + +theory State_d_Equation imports State_0_Equation + +begin + +context rm_eq_fixes +begin + + text \Equation 4.25\ + definition state_d :: "bool" where + "state_d \ \d>0. d\m \ s d = b*\S+ p d s + b*\S- p d (\k. s k && z (modifies (p!k))) + + b*\S0 p d (\k. s k && (e - z (modifies (p!k))))" + + text \Combining the two\ + definition state_relations_from_recursion :: "bool" where + "state_relations_from_recursion \ state_0 \ state_d" + +end + +context register_machine +begin + +lemma state_d_dioph: + fixes b e :: polynomial + fixes z s :: "polynomial list" + assumes "length z = n" "length s = Suc m" + defines "DR \ LARY (\ll. rm_eq_fixes.state_d p (ll!0!0) (ll!0!1) + (nth (ll!1)) (nth (ll!2))) + [[b, e], z, s]" + shows "is_dioph_rel DR" +proof - + + define d_domain where "d_domain \ [1..x. push_param x number_of_ex_vars) z" + "s' = map (\x. push_param x number_of_ex_vars) s" + + note e'_def = \e' = push_param e number_of_ex_vars\ + + define z0 z1 where z_def: "z0 \ map (\i. z' ! modifies (p!i)) [0.. map (\i. e' [-] z' ! modifies (p!i)) [0.. (\d. Param (d - Suc 0))" + write sum_ssub_nzero_param_of_state ("\S-'_Param _") + + define sum_ssub_zero_param_of_state where + "sum_ssub_zero_param_of_state \ (\d. Param (m + d - Suc 0))" + write sum_ssub_zero_param_of_state ("\S0'_Param _") + + define param_is_sum_ssub_nzero_term where + "param_is_sum_ssub_nzero_term \ (\d::nat. [(\S-_Param d) = \S- d (s' && z0)])" + + define param_is_sum_ssub_zero_term where + "param_is_sum_ssub_zero_term \ (\d. [(\S0_Param d) = \S0 d (s' && z1)])" + + define params_are_sum_terms where + "params_are_sum_terms \ [\ in d_domain] (\d. param_is_sum_ssub_nzero_term d + [\] param_is_sum_ssub_zero_term d)" + + define step_relation where + "step_relation \ (\d. (s'!d) [=] b' [*] ([\S+] p d (nth s')) + [+] b' [*] (\S-_Param d) + [+] b' [*] (\S0_Param d))" + + define DS where "DS \ [\number_of_ex_vars] (([\ in d_domain] (\d. step_relation d)) + [\] params_are_sum_terms)" + + have "length p > 0" + using p_nonempty by auto + hence "m \ 0" + unfolding m_def by auto + have "length s' = Suc m" and "length z0 = Suc m" and "length z1 = Suc m" + unfolding pushed_def z_def using \length s = Suc m\ m_def \length p > 0\ by auto + + have "eval DS a = eval DR a" for a + proof - + + have b'_unfold: "peval b' (push_list a ks) = peval b a" + if "length ks = number_of_ex_vars" for ks + unfolding pushed_def using push_push_simp \length d_domain = m\ by (metis that) + + have h0: "k < m \ d_domain ! k < Suc m" for k + unfolding d_domain_def apply simp + using One_nat_def Suc_pred \0 < length p\ add.commute + assms(3) d_domain_def less_diff_conv m_def nth_upt upt_Suc_append + by (smt \length d_domain = m\ less_nat_zero_code list.size(3) upt_Suc) + + have s'_unfold: "peval (s' ! (d_domain ! k)) (push_list a ks) + = peval (s ! (d_domain ! k)) a" + if "length ks = number_of_ex_vars" and "k < m" for k ks + proof - + from \k < m\ have "d_domain ! k < length s" unfolding \length s = Suc m\ + using h0 by blast + + have suc_k: "([Suc 0..k < m\) + + have "peval (map (\x. push_param x number_of_ex_vars) s ! (d_domain ! k)) (push_list a ks) + = list_eval s a (d_domain ! k)" + using push_push_map_i[of "ks" "number_of_ex_vars" "d_domain!k" s a] + using \length ks = number_of_ex_vars\ \k < m\ h0 \length s = Suc m\ by auto + also have "... = peval (s ! (d_domain ! k)) a" + unfolding list_eval_def + using nth_map [of "d_domain ! k" s "(\x. peval x a)"] \d_domain ! k < length s\ + unfolding d_domain_def using \m \ 0\ \k < m\ suc_k by auto + + finally show ?thesis unfolding pushed_def by auto + qed + + have sum_sadd_unfold: "(\S+ p (d_domain ! k) (\x. peval (s' ! x) (push_list a ks))) + = (\S+ p (d_domain ! k) ((!) (map (\P. peval P a) s)))" + if "length ks = number_of_ex_vars" for k ks + apply (rule sum_sadd_cong, auto) unfolding pushed_def + using push_push_map_i[of ks number_of_ex_vars _ s a] \length ks = number_of_ex_vars\ + unfolding list_eval_def by (simp add: \length s = Suc m\ m_def) + + have s: "peval (s' ! ka) (push_list a ks) = map (\P. peval P a) s ! ka" + if "ka < Suc m" and "length ks = number_of_ex_vars" for ka ks + unfolding pushed_def + using push_push_map_i[of ks number_of_ex_vars ka s a] \length ks = number_of_ex_vars\ + using list_eval_def \length s = Suc m\ \ka < Suc m\ by auto + + have modifies_valid: "modifies (p ! ka) < length z" if "ka < Suc m" for ka + using modifies_yields_valid_register that unfolding \length z = n\ m_def + using p_nonempty by auto + + have sum_ssub_nzero_unfold: + "(\S- p (d_domain ! k) (\k. peval (s' ! k) (push_list a ks) + && peval (z0 ! k) (push_list a ks))) + = (\S- p (d_domain ! k) (\k. map (\P. peval P a) s ! k + && map (\P. peval P a) z ! modifies (p ! k)))" + if "length ks = number_of_ex_vars" for k ks + proof- + have z0: "peval (z0 ! ka) (push_list a ks) = map (\P. peval P a) z ! modifies (p ! ka)" + if "ka < Suc m" for ka + unfolding z_def pushed_def + using push_push_map_i[of ks number_of_ex_vars "modifies (p!ka)" z a] + \length ks = number_of_ex_vars\ unfolding list_eval_def + using \length z0 = Suc m\ \ka < Suc m\ modifies_valid \0 < length p\ m_def map_nth by force + + show ?thesis apply (rule sum_ssub_nzero_cong) using s z0 le_imp_less_Suc m_def that + by presburger + qed + + have sum_ssub_zero_unfold: + "(\S0 p (d_domain ! k) (\k. peval (s' ! k) (push_list a ks) + && peval (z1 ! k) (push_list a ks))) + = (\S0 p (d_domain ! k) (\k. map (\P. peval P a) s ! k + && peval e a - map (\P. peval P a) z ! modifies (p ! k)))" + if "length ks = number_of_ex_vars" and "k < Suc m" for k ks + proof- + + have map: + "map (\i. e' [-] (z' ! (modifies (p ! i)))) [0..i. e' [-] z' ! modifies (p ! i)"] \ka < Suc m\ + by (smt (z3) One_nat_def Suc_pred \0 < length p\ \m \ 0\ le_trans length_map m_def map_nth + nth_map upt_Suc_append zero_le_one) + + have "peval (e' [-] (z' ! modifies (p ! ka))) (push_list a ks) + = peval e a - map (\P. peval P a) z ! modifies (p ! ka)" + if "ka < Suc m" for ka + unfolding z_def pushed_def apply (simp add: defs) + using push_push_simp \length ks = number_of_ex_vars\ apply auto + using push_push_map_i[of ks number_of_ex_vars "modifies (p!ka)" z a] + \length ks = number_of_ex_vars\ modifies_valid \ka < Suc m\ + unfolding list_eval_def using \length z0 = Suc m\ \0 < length p\ m_def map_nth by auto + + hence z1: "peval (z1 ! ka) (push_list a ks) + = peval e a - map (\P. peval P a) z ! modifies (p ! ka)" if "ka < Suc m" for ka + unfolding z_def using map that by auto + + show ?thesis apply (rule sum_ssub_zero_cong) using s z1 le_imp_less_Suc m_def that + by presburger + + qed + + define sum_ssub_nzero_map where + "sum_ssub_nzero_map \ \j. \S- p j (\k. map (\P. peval P a) s ! k + && map (\P. peval P a) z ! modifies (p ! k))" + define sum_ssub_zero_map where + "sum_ssub_zero_map \ \j. \S0 p j (\k. map (\P. peval P a) s ! k + && peval e a - map (\P. peval P a) z ! modifies (p ! k)) " + + define ks_ex where + "ks_ex \ map sum_ssub_nzero_map d_domain @ map sum_ssub_zero_map d_domain" + + have "length ks_ex = number_of_ex_vars" + unfolding ks_ex_def number_of_ex_vars_def using \length d_domain = m\ by auto + + have ks_ex1: + "peval (\S-_Param (d_domain ! k)) (push_list a ks_ex) + = \S- p (d_domain ! k) (\k. map (\P. peval P a) s ! k + && map (\P. peval P a) z ! modifies (p ! k))" + if "k < m" for k + proof - + have domain_at_k_bound: + "d_domain ! k - Suc 0 < length ks_ex" using that \length ks_ex = number_of_ex_vars\ + unfolding number_of_ex_vars_def using h0 by fastforce + + have "peval (\S-_Param (d_domain ! k)) (push_list a ks_ex) = ks_ex ! k" + unfolding push_list_def sum_ssub_nzero_param_of_state_def using that domain_at_k_bound + apply auto + using One_nat_def Suc_mono d_domain_def diff_Suc_1 nth_upt plus_1_eq_Suc by presburger + + also have "... = \S- p (d_domain ! k) (\k. map (\P. peval P a) s ! k + && map (\P. peval P a) z ! modifies (p ! k))" + unfolding ks_ex_def + unfolding nth_append[of "map sum_ssub_nzero_map d_domain" "map sum_ssub_zero_map d_domain" k] + using \length d_domain = m\ that unfolding sum_ssub_nzero_map_def by auto + finally show ?thesis by auto + qed + + have ks_ex2: + "peval (\S0_Param (d_domain ! k)) (push_list a ks_ex) + = \S0 p (d_domain ! k) (\k. map (\P. peval P a) s ! k + && peval e a - map (\P. peval P a) z ! modifies (p ! k))" + if "k < m" for k + proof - + have domain_at_k_bound: + "m + d_domain ! k - Suc 0 < length ks_ex" using that \length ks_ex = number_of_ex_vars\ + unfolding number_of_ex_vars_def using h0 by fastforce + + have "d_domain ! k \ 1" + unfolding d_domain_def \k < m\ + using m_def p_nonempty that by auto + + hence index_calculation: "(m + d_domain ! k - Suc 0) = k + m" + unfolding d_domain_def + by (metis (no_types, lifting) Nat.add_diff_assoc One_nat_def Suc_pred add.commute + less_diff_conv m_def nth_upt ordered_cancel_comm_monoid_diff_class.le_imp_diff_is_add + p_nonempty that) + + have "peval (\S0_Param (d_domain ! k)) (push_list a ks_ex) = ks_ex ! (k + m)" + unfolding push_list_def sum_ssub_zero_param_of_state_def using that domain_at_k_bound + by (auto simp: index_calculation) + + also have "... = \S0 p (d_domain ! k) (\k. map (\P. peval P a) s ! k + && peval e a - map (\P. peval P a) z ! modifies (p ! k))" + unfolding ks_ex_def + unfolding nth_append[of "map sum_ssub_nzero_map d_domain" "map sum_ssub_zero_map d_domain"] + using \length d_domain = m\ that unfolding sum_ssub_zero_map_def by auto + finally show ?thesis by auto + qed + + have all_quantifier_switch: "(\kd>0. d \ m \ Property d)" for Property + proof (rule) + assume asm: "\kd>0. d \ m \ Property d" + proof (auto) + fix d + assume "d > 0" "d \ m" + hence "d - Suc 0 < length d_domain" + by (metis Suc_le_eq Suc_pred \length d_domain = m\) + hence "Property (d_domain ! (d - Suc 0))" + using asm by auto + thus "Property d" + unfolding d_domain_def + by (metis One_nat_def Suc_diff_1 \0 < d\ \d \ m\ le_imp_less_Suc nth_upt plus_1_eq_Suc) + qed + next + assume asm: "\d>0. d \ m \ Property d" + show "\k 0" + unfolding d_domain_def + by (smt (z3) One_nat_def Suc_leI Suc_pred \0 < length p\ \length d_domain = m\ + add_less_cancel_left d_domain_def diff_is_0_eq' gr_zeroI le_add_diff_inverse + less_nat_zero_code less_numeral_extra(1) m_def nth_upt) + moreover have "d_domain ! k \ m" + unfolding d_domain_def using \k < length d_domain\ unfolding \length d_domain = m\ + using d_domain_def h0 less_Suc_eq_le by auto + ultimately show "Property (d_domain ! k)" + using asm by auto + qed + qed + + have "peval (s!d) a = map (\P. peval P a) s ! d" if "d > 0" and "d \ m" for d + using nth_map[of d s "\P. peval P a"] that \length s = Suc m\ by simp + + (* Start chain of equivalences *) + have "eval DS a = (\ks. number_of_ex_vars = length ks + \ (\k eval params_are_sum_terms (push_list a ks))" + unfolding DS_def by (simp add: defs) + + also have "... = (\ks. number_of_ex_vars = length ks \ + (\kS+] p (d_domain ! k) ((!) s')) (push_list a ks) + + peval b a * peval (\S-_Param (d_domain ! k)) (push_list a ks) + + peval b a * peval (\S0_Param (d_domain ! k)) (push_list a ks)) \ + eval params_are_sum_terms (push_list a ks))" + unfolding step_relation_def \length d_domain = m\ + using b'_unfold s'_unfold by (auto simp: defs) + + also have "... = (\ks. number_of_ex_vars = length ks \ + (\kS+ p (d_domain ! k) (\x. peval (s' ! x) (push_list a ks))) + + peval b a * (peval (\S-_Param (d_domain ! k)) (push_list a ks)) + + peval b a * (peval (\S0_Param (d_domain ! k)) (push_list a ks))) + \ (\kS-_Param (d_domain ! k)) (push_list a ks) + = \S- p (d_domain ! k) (\k. peval (s' ! k) (push_list a ks) + && peval (z0 ! k) (push_list a ks)) + \ peval (\S0_Param (d_domain ! k)) (push_list a ks) + = \S0 p (d_domain ! k) (\k. peval (s' ! k) (push_list a ks) + && peval (z1 ! k) (push_list a ks))))" + unfolding params_are_sum_terms_def param_is_sum_ssub_nzero_term_def + param_is_sum_ssub_zero_term_def apply (simp add: defs) + using sum_rsub_nzero_of_bit_and_eval[of s' z0] sum_rsub_zero_of_bit_and_eval[of s' z1] + \length p > 0\ \length s' = Suc m\ \length z0 = Suc m\ \length z1 = Suc m\ + unfolding \length d_domain = m\ by (simp add: defs) + + also have "... = (\ks. number_of_ex_vars = length ks \ + (\kS+ p (d_domain ! k) ((!) (map (\P. peval P a) s)) ) + + peval b a * (\S- p (d_domain ! k) (\k. map (\P. peval P a) s ! k + && map (\P. peval P a) z ! modifies (p ! k))) + + peval b a * (\S0 p (d_domain ! k) (\k. map (\P. peval P a) s ! k + && peval e a - map (\P. peval P a) z ! modifies (p ! k)))) + \ (\kS-_Param (d_domain ! k)) (push_list a ks) + = \S- p (d_domain ! k) (\k. map (\P. peval P a) s ! k + && map (\P. peval P a) z ! modifies (p ! k)) + \ peval (\S0_Param (d_domain ! k)) (push_list a ks) + = \S0 p (d_domain ! k) (\k. map (\P. peval P a) s ! k + && peval e a - map (\P. peval P a) z ! modifies (p ! k))))" + using sum_sadd_unfold sum_ssub_nzero_unfold sum_ssub_zero_unfold by auto + + also have "... = (\kS+ p (d_domain ! k) ((!) (map (\P. peval P a) s)) ) + + peval b a * (\S- p (d_domain ! k) + (\k. map (\P. peval P a) s ! k + && map (\P. peval P a) z ! modifies (p ! k))) + + peval b a * (\S0 p (d_domain ! k) + (\k. map (\P. peval P a) s ! k + && peval e a - map (\P. peval P a) z ! modifies (p ! k))))" + apply auto + apply (rule exI[of _ ks_ex]) + using \length ks_ex = number_of_ex_vars\ ks_ex1 ks_ex2 by auto + + also have "... = (\d>0. d \ m \ + peval (s ! d) a + = peval b a * \S+ p d ((!) (map (\P. peval P a) s)) + + peval b a * \S- p d (\k. map (\P. peval P a) s ! k + && map (\P. peval P a) z ! modifies (p ! k)) + + peval b a * \S0 p d (\k. map (\P. peval P a) s ! k + && peval e a - map (\P. peval P a) z ! modifies (p ! k)) )" + using all_quantifier_switch[of "\d. peval (s ! d) a = + peval b a * \S+ p d ((!) (map (\P. peval P a) s)) + + peval b a * \S- p d (\k. map (\P. peval P a) s ! k + && map (\P. peval P a) z ! modifies (p ! k)) + + peval b a * \S0 p d (\k. map (\P. peval P a) s ! k + && peval e a - map (\P. peval P a) z ! modifies (p ! k))"] + unfolding \length d_domain = m\ by auto + + finally show ?thesis + unfolding DR_def + using local.register_machine_axioms rm_eq_fixes_def[of p n] rm_eq_fixes.state_d_def[of p n] + apply (simp add: defs) + using nth_map[of _ s "\P. peval P a"] \length s = Suc m\ + by auto + qed + + moreover have "is_dioph_rel DS" + proof - + have "is_dioph_rel (param_is_sum_ssub_nzero_term d [\] param_is_sum_ssub_zero_term d)" for d + unfolding param_is_sum_ssub_nzero_term_def param_is_sum_ssub_zero_term_def + by (auto simp: dioph) + hence 1: "list_all (is_dioph_rel \ (\d. param_is_sum_ssub_nzero_term d + [\] param_is_sum_ssub_zero_term d)) d_domain" + by (simp add: list.inducts) + + have "is_dioph_rel (step_relation d)" for d + unfolding step_relation_def by (auto simp: dioph) + hence 2: "list_all (is_dioph_rel \ step_relation) d_domain" + by (simp add: list.inducts) + + show ?thesis + unfolding DS_def params_are_sum_terms_def by (auto simp: dioph 1 2) + qed + + ultimately show ?thesis using is_dioph_rel_def by auto +qed + + +lemma state_relations_from_recursion_dioph: + fixes b e :: polynomial + fixes z s :: "polynomial list" + assumes "length z = n" "length s = Suc m" + defines "DR \ LARY (\ll. rm_eq_fixes.state_relations_from_recursion p (ll!0!0) (ll!0!1) + (nth (ll!1)) (nth (ll!2))) + [[b, e], z, s]" + shows "is_dioph_rel DR" +proof - + + define DS where "DS \ (LARY (\ll. rm_eq_fixes.state_0 p (ll!0!0) (ll!0!1) + (nth (ll!1)) (nth (ll!2))) [[b, e], z, s]) + [\](LARY (\ll. rm_eq_fixes.state_d p (ll!0!0) (ll!0!1) (nth (ll!1)) + (nth (ll!2))) [[b, e], z, s])" + + have "eval DS a = eval DR a" for a + unfolding DS_def DR_def + using local.register_machine_axioms rm_eq_fixes_def + rm_eq_fixes.state_relations_from_recursion_def + using assms by (simp add: defs) + + moreover have "is_dioph_rel DS" + unfolding DS_def apply (rule and_dioph) using assms state_0_dioph state_d_dioph by blast+ + + ultimately show ?thesis using is_dioph_rel_def by auto +qed + +end + +end diff --git a/thys/DPRM_Theorem/ROOT b/thys/DPRM_Theorem/ROOT new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/ROOT @@ -0,0 +1,69 @@ +chapter AFP + +session DPRM_Theorem (AFP) = HOL + + description \The DPRM Theorem\ + options [timeout = 600] + + sessions + "HOL-Library" + "Lucas_Theorem" + "Digit_Expansions" + + directories + "Diophantine" + "Register_Machine" + "Machine_Equations" + + theories + (* Diophantine Equations *) + "Diophantine/Parametric_Polynomials" + "Diophantine/Assignments" + "Diophantine/Diophantine_Relations" + "Diophantine/Existential_Quantifier" + "Diophantine/Modulo_Divisibility" + + (* Exponentiation is Diophantine *) + "Diophantine/Exponentiation" + "Diophantine/Alpha_Sequence" + "Diophantine/Exponential_Relation" + + (* Diophantization of Digitwise Operations *) + "Diophantine/Digit_Function" + "Diophantine/Binomial_Coefficient" + "Diophantine/Binary_Masking" + "Diophantine/Binary_Orthogonal" + "Diophantine/Binary_And" + + (* Register Machines *) + "Register_Machine/RegisterMachineSpecification" + "Register_Machine/RegisterMachineProperties" + "Register_Machine/RegisterMachineSimulation" + "Register_Machine/SingleStepRegister" + "Register_Machine/SingleStepState" + "Register_Machine/MultipleStepRegister" + "Register_Machine/MultipleStepState" + "Register_Machine/MachineMasking" + + (* Arithmetization of Register Machines *) + "Register_Machine/MachineEquations" + "Register_Machine/CommutationRelations" + "Register_Machine/MultipleToSingleSteps" + "Machine_Equations/Equation_Setup" + "Machine_Equations/RM_Sums_Diophantine" + "Diophantine/Register_Machine_Sums" + "Machine_Equations/Register_Equations" + "Machine_Equations/State_0_Equation" + "Machine_Equations/State_d_Equation" + "Machine_Equations/All_State_Equations" + "Machine_Equations/Mask_Equations" + "Machine_Equations/Constants_Equations" + "Machine_Equations/All_Equations_Invariance" + "Machine_Equations/All_Equations" + "Machine_Equations/Machine_Equation_Equivalence" + + (* The DPRM Theorem *) + DPRM + + document_files + "root.tex" + "root.bib" diff --git a/thys/DPRM_Theorem/Register_Machine/CommutationRelations.thy b/thys/DPRM_Theorem/Register_Machine/CommutationRelations.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Register_Machine/CommutationRelations.thy @@ -0,0 +1,682 @@ +subsection \Preliminary commutation relations\ + +theory CommutationRelations + imports RegisterMachineSimulation MachineEquations +begin + +lemma aux_commute_bitAND_sum: + fixes N C :: nat + and fxt :: "nat \ nat" + shows "\i\N. \j\N. i \ j \ (\k. (fct i) \ k * (fct j) \ k = 0) + \ (\k \ N. fct k && C) = (\k \ N. fct k) && C" +proof (induct N) + case 0 + then show ?case by auto +next + case (Suc N) + assume Suc_assms: "\i\Suc N. \j\Suc N. i \ j \ (\k. fct i \ k * fct j \ k = 0)" + + have "(\k\Suc N. fct k && C) = (\k\N. fct k && C) + (fct (Suc N) && C)" + by (auto cong: sum.cong) + also have "... = (sum fct {..N} && C) + (fct (Suc N) && C)" + using Suc by auto + also have "... = (sum fct {..N} + fct (Suc N)) && C" + proof - + let ?a = "sum fct {..N} && C" + let ?b = "fct (Suc N) && C" + + have formula: "\d. ( ?a + ?b ) \ d = ( ?a \ d + ?b \ d + bin_carry ?a ?b d ) mod 2" + using sum_digit_formula by auto + + + (* mutual proof of both statements, they are interdependent in the inductive step *) + have nocarry4: "\n\Suc N. \d. (sum fct {..n} \ d > 0 \ (\i\n. (fct i) \ d > 0)) + \ bin_carry (sum fct {.. Suc N \ + \d. ((\i\n. (fct i) \ d = 0) + \ sum fct {..n} \ d = 0) \ bin_carry (sum fct {..d. sum fct {.. d > 0 \ (\i d = 1)" + using nth_bit_def + by (metis One_nat_def Suc_less_eq lessThan_Suc_atMost less_Suc_eq nat_less_le + not_mod2_eq_Suc_0_eq_0) + + (* then show another helper result ... *) + have h1: "\d. sum fct {.. d + fct (Suc m) \ d \ 1" + proof - + { + fix d + have "sum fct {.. d + fct (Suc m) \ d \ 1" + proof (cases "sum fct {.. d = 0") + case True + have "fct (Suc m) \ d \ 1" + using nth_bit_def by auto + then show ?thesis using True by auto + next + case False + then have "\i d > 0" + using ex by (metis neq0_conv zero_less_one) + then obtain i where i: "i fct i \ d > 0" by auto + hence "i \ Suc N" using Suc.prems by auto + hence "\j\Suc N. i \ j \ (\k. fct i \ k * fct j \ k = 0)" + using Suc_assms by auto + then have "fct (Suc m) \ d = 0" + using Suc.prems i nat_neq_iff + by (auto) (blast) + moreover from False have "sum fct {.. d = 1" + by (simp add: nth_bit_def) + ultimately show ?thesis by auto + qed + } + thus ?thesis by auto + qed + + (* ... in order to show the inductive step of the second statement *) + from h1 have h2: "\d. bin_carry (sum fct {..d. bin_carry (sum fct {..m}) (fct (Suc m)) d = 0" + by (simp add: lessThan_Suc_atMost) + + (* finally use the just proven inductive step for the 2nd statement to show the + inductive step for the first statement *) + { + fix d + assume a: "\i\Suc m. (fct i) \ d = 0" + have "sum fct {..Suc m} \ d = (sum fct {..m} + fct (Suc m)) \ d" + by auto + also have "... = + (sum fct {..m} \ d + fct (Suc m) \ d + bin_carry (sum fct {..m}) (fct (Suc m)) d) mod 2" + using sum_digit_formula[of "sum fct {..m}" "fct (Suc m)" "d"] by auto + finally have "sum fct {..Suc m} \ d = 0" + using nocarry3 Suc a by auto + } + with h2 show ?case by auto + qed + + then have "n \ Suc N \ \d. (sum fct {..n} \ d > 0 \ (\i\n. (fct i) \ d > 0)) + \ bin_carry (sum fct {..d. ?a \ d + ?b \ d \ 1" + proof - + have "\d. ?a \ d + ?b \ d = (sum fct {..N} \ d + fct (Suc N) \ d) * C \ d" + using bitAND_digit_mult add_mult_distrib by auto + then have "\d. ?a \ d + ?b \ d \ (sum fct {..N} \ d + fct (Suc N) \ d)" + using nth_bit_def by auto + thus ?thesis + using sum_carry_formula nocarry4 no_carry_mult_equiv nth_bit_bounded bitAND_digit_mult + by (metis One_nat_def add.commute add_decreasing le_Suc_eq lessThan_Suc_atMost + nat_1_eq_mult_iff) + qed + from h3 have h4: "\d. bin_carry ?a ?b d = 0" + by (metis Suc_1 Suc_n_not_le_n carry_digit_impl) + + (* Helper for proof chain below *) + have h5: "\d. bin_carry (sum fct {..N}) (fct (Suc N)) d = 0" + using nocarry4 lessThan_Suc_atMost by auto + + from formula h3 h4 have "\d. ( ?a + ?b ) \ d = ?a \ d + ?b \ d" + by (metis (no_types, lifting) add_cancel_right_left add_le_same_cancel1 add_self_mod_2 + le_zero_eq not_mod2_eq_Suc_0_eq_0 nth_bit_def one_mod_two_eq_one plus_1_eq_Suc) + + then have "\d. ( ?a + ?b ) \ d = sum fct {..N} \ d * C \ d + fct (Suc N) \ d * C \ d" + using bitAND_digit_mult by auto + + then have "\d. ( ?a + ?b ) \ d = (sum fct {..N} \ d + fct (Suc N) \ d) * C \ d" + by (simp add: add_mult_distrib) + moreover have "\d. sum fct {..N} \ d + fct (Suc N) \ d + = ( sum fct {..N} \ d + fct (Suc N) \ d + bin_carry (sum fct {..N}) (fct (Suc N)) d ) mod 2" + using h5 sum_carry_formula + by (metis add_diff_cancel_left' add_diff_cancel_right' mult_div_mod_eq mult_is_0) + ultimately have "\d. ( ?a + ?b ) \ d = (sum fct {..N} + fct (Suc N)) \ d * C \ d" + using sum_digit_formula by auto + + then have "\d. ( ?a + ?b ) \ d = ( (sum fct {..N} + fct (Suc N)) && C ) \ d" + using bitAND_digit_mult by auto + thus ?thesis using digit_wise_equiv by blast + qed + ultimately show ?case by auto +qed + + +lemma aux_commute_bitAND_sum_if: + fixes N const :: nat + assumes nocarry: "\i\N. \j\N. i \ j \ (\k. (fct i) \ k * (fct j) \ k = 0)" + shows "(\k \ N. if cond k then fct k && const else 0) + = (\k \ N. if cond k then fct k else 0) && const" +proof - + from nocarry have nocarry_if: + "\i\N. \j\N. i \ j \ (\k. (if cond i then fct i else 0) \ k * (if cond j then fct j else 0) \ k = 0)" + by (metis (full_types) aux1_digit_wise_equiv mult.commute mult_zero_left) + + have "(if cond k then fct k && const else 0) = (if cond k then fct k else 0) && const" for k + by auto + hence "(\k \ N. if cond k then fct k && const else 0) + = (\k \ N. (if cond k then fct k else 0) && const)" + by auto + also have "... = (\k \ N. if cond k then fct k else 0) && const" + using nocarry_if aux_commute_bitAND_sum[where ?fct = "\k. (if cond k then fct k else 0)"] + by blast + ultimately show ?thesis by auto +qed + +lemma mod_mod: + fixes x a b :: nat + shows "x mod 2^a mod 2^b = x mod 2^(min a b)" + by (metis min.commute take_bit_eq_mod take_bit_take_bit) + +lemma carry_gen_pow2_reduct: + assumes "c>0" + defines b: "b \ 2 ^ (Suc c)" + assumes "nth_digit x (t-1) (2^Suc c) \ c = 0" + and "nth_digit y (t-1) (2^Suc c) \ c = 0" + shows "k\c \ bin_carry (nth_digit x t b) (nth_digit y t b) k + = bin_carry x y (Suc c * t + k)" +proof (induction k) + case 0 + then show ?case + proof (cases "t=0") + case True + then show ?thesis using bin_carry_def by auto + next + case False + hence "t>0" by auto + from assms(3) have "x \ (Suc c * (t - 1) + c) = 0" + using digit_gen_pow2_reduct[of "c" "Suc c" "x" "t-1"] by auto + moreover have "y \ (Suc c * (t - 1) + c) = 0" + using assms(4) digit_gen_pow2_reduct[of "c" "Suc c" "y" "t-1"] by auto + moreover have "Suc c * (t - 1) + c = t + c * t - Suc 0" + using add.left_commute gr0_conv_Suc \t>0\ by auto + ultimately have "(x \ (t + c * t - Suc 0) + y \ (t + c * t - Suc 0) + + bin_carry x y (t + c * t - Suc 0)) \ 1" using carry_bounded by auto + hence "bin_carry x y (Suc c * t) = 0" + using sum_carry_formula[of "x" "y" "Suc c * t - 1"] \c>0\ \t>0\ by auto + + moreover have "bin_carry (nth_digit x t b) (nth_digit y t b) 0 = 0" + using 0 bin_carry_def by auto + ultimately show ?thesis by auto + qed +next + case (Suc k) + have "k x \ (Suc c * t + k) = nth_digit x t b \ k" + using digit_gen_pow2_reduct[of "k" "Suc c" "x" "t"] b by auto + moreover have "k y \ (Suc c * t + k) = nth_digit y t b \ k" + using digit_gen_pow2_reduct[of "k" "Suc c" "y" "t"] b by auto + ultimately show ?case using Suc + sum_carry_formula[of "nth_digit x t b" "nth_digit y t b" "k"] + sum_carry_formula[of "x" "y" "Suc c * t + k"] + by auto +qed + +lemma nth_digit_bound: + fixes c defines "b \ 2 ^ (Suc c)" + shows "nth_digit x t b < 2 ^ (Suc c)" + using nth_digit_def b_def by auto + +lemma digit_wise_block_additivity: + fixes c + defines "b \ 2 ^ Suc c" + assumes "nth_digit x (t-1) (2^Suc c) \ c = 0" + and "nth_digit y (t-1) (2^Suc c) \ c = 0" + and "k\c" + and "c>0" + shows "nth_digit (x+y) t b \ k = (nth_digit x t b + nth_digit y t b) \ k" +proof - + have "k < Suc c" using `k\c` by simp + have x: "nth_digit x t b \ k = x \ (Suc c * t + k)" + using digit_gen_pow2_reduct[of "k" "Suc c" "x" "t"] b_def `k k = y \ (Suc c * t + k)" + using digit_gen_pow2_reduct[of "k" "Suc c" "y" "t"] b_def `k k + = ((nth_digit x t b) \ k + (nth_digit y t b) \ k + + bin_carry (nth_digit x t b) (nth_digit y t b) k) mod 2" + using sum_digit_formula[of "nth_digit x t b" "nth_digit y t b" "k"] by auto + also have "... = (x \ (Suc c * t + k) + y \ (Suc c * t + k) + + bin_carry (nth_digit x t b) (nth_digit y t b) k) mod 2" + using x y by auto + also have "... = (x \ (Suc c * t + k) + y \ (Suc c * t + k) + + bin_carry x y (Suc c * t + k)) mod 2" + using carry_gen_pow2_reduct[of "c" "x" "t" "y" "k"] assms by auto + also have "... = (x + y) \ (Suc c * t + k)" + using sum_digit_formula by auto + also have "... = nth_digit (x+y) t b \ k" + using digit_gen_pow2_reduct[of "k" "Suc c" "x+y" "t"] b_def `k 0" + defines "b \ 2 ^ Suc c" + assumes "nth_digit x (t-1) b \ c = 0" + and "nth_digit y (t-1) b \ c = 0" + and "nth_digit x t b \ c = 0" + and "nth_digit y t b \ c = 0" + (* needs stronger assumptions than digit_wise_block_additivity *) + shows "nth_digit (x+y) t b = nth_digit x t b + nth_digit y t b" +proof - + { + have "nth_digit x t b < b" using nth_digit_bound b_def by auto + hence x_digit_bound: "\k. k\Suc c \ nth_digit x t b \ k = 0" + using nth_bit_def b_def aux_lt_implies_mask b_def by metis + + have "nth_digit y t b < b" using nth_digit_bound b_def by auto + hence y_digit_bound: "\k. k\Suc c \ nth_digit y t b \ k = 0" + using nth_bit_def b_def aux_lt_implies_mask b_def by metis + + fix k + assume k: "k\Suc c" + have carry0: "bin_carry (nth_digit x t b) (nth_digit y t b) k = 0" + proof - (* induction proof using rule nat_induct_at_least to accommodate fact k *) + (* base case *) + have base: "bin_carry (nth_digit x t b) (nth_digit y t b) (Suc c) = 0" + using sum_carry_formula[where ?k = "c"] bin_carry_bounded[where ?k = "c"] + using assms(5-6) by (metis Suc_eq_plus1 add_cancel_left_left mod_div_trivial) + + (* inductive step *) + { + fix n + assume n: "n \ Suc c" + assume IH: "bin_carry (nth_digit x t b) (nth_digit y t b) n = 0" + have "bin_carry (nth_digit x t b) (nth_digit y t b) (Suc n) + = (nth_digit x t b \ n + nth_digit y t b \ n + + bin_carry (nth_digit x t b) (nth_digit y t b) n) div 2" + using sum_carry_formula[of "nth_digit x t b" "nth_digit y t b"] by auto + also have "... = bin_carry (nth_digit x t b) (nth_digit y t b) n div 2" + using x_digit_bound y_digit_bound n by auto + also have "... = 0" using IH by auto + + finally have "bin_carry (nth_digit x t b) (nth_digit y t b) (Suc n) = 0" + by auto + } + + then show ?thesis + using k base + using nat_induct_at_least[where ?P = "\k. bin_carry (nth_digit x t b) + (nth_digit y t b) k = 0"] + by auto + qed + + have "(nth_digit x t b + nth_digit y t b) \ k + = (nth_digit x t b \ k + nth_digit y t b \ k + + bin_carry (nth_digit x t b) (nth_digit y t b) k) mod 2" + using sum_digit_formula[of "nth_digit x t b" "nth_digit y t b" "k"] by auto + hence separate_sum: "(nth_digit x t b + nth_digit y t b) \ k = 0" + using x_digit_bound y_digit_bound carry0 k by auto + + have "nth_digit (x+y) t b < b" + using nth_digit_bound b_def by auto + hence xy_sum: "nth_digit (x+y) t b \ k = 0" + using nth_bit_def b_def aux_lt_implies_mask b_def k by metis + from xy_sum separate_sum have k_ge: "nth_digit (x+y) t b \ k + = (nth_digit x t b + nth_digit y t b) \ k" + by auto + } + hence k_ge: "k\Suc c \ nth_digit (x+y) t b \ k + = (nth_digit x t b + nth_digit y t b) \ k" for k + by auto + + moreover have k_lt:"k nth_digit (x+y) t b \ k + = (nth_digit x t b + nth_digit y t b) \ k" for k + using digit_wise_block_additivity assms by auto + + ultimately have "nth_digit (x+y) t b \ k + = (nth_digit x t b + nth_digit y t b) \ k" for k + by(cases "k0" + defines b: "b \ 2 ^ (Suc c)" + assumes yltx_digits: "\t'. nth_digit y t' b \ nth_digit x t' b" + shows "y mod b^t \ x mod b^t" +proof(cases "t=0") + case True + then show ?thesis by auto +next + case False + show ?thesis using yltx_digits apply(induct t, auto) using yltx_digits + by (smt add.commute add_left_cancel add_mono_thms_linordered_semiring(1) mod_mult2_eq + mult_le_cancel2 nth_digit_def semiring_normalization_rules(7)) +qed + +lemma narry_gen_pow2_reduct: + assumes "c>0" + defines b: "b \ 2 ^ (Suc c)" + assumes yltx_digits: "\t'. nth_digit y t' b \ nth_digit x t' b" + shows "k\c \ bin_narry (nth_digit x t b) (nth_digit y t b) k + = bin_narry x y (Suc c * t + k)" +proof (induction k) + case 0 + then show ?case + proof (cases "t=0") + case True + then show ?thesis by (simp add: bin_narry_def) + next + case False + have "bin_narry x y (Suc c * t) = 0" using yltx_digits block_to_sum bin_narry_def assms + by (metis not_less power_mult) + then show ?thesis by (simp add: bin_narry_def) + qed +next + case (Suc k) + have ylex: "y\x" using yltx_digits digitwise_leq b Suc_1 lessI power_gt1 by metis + have "k x \ (Suc c * t + k) = nth_digit x t b \ k" + using digit_gen_pow2_reduct[of "k" "Suc c" "x" "t"] b by auto + moreover have "k y \ (Suc c * t + k) = nth_digit y t b \ k" + using digit_gen_pow2_reduct[of "k" "Suc c" "y" "t"] b by auto + ultimately show ?case using Suc yltx_digits + dif_narry_formula[of "(nth_digit y t b)" "(nth_digit x t b)" k] + dif_narry_formula[of y x "Suc c * t + k"] ylex by auto +qed + +lemma digit_wise_block_subtractivity: + fixes c + defines "b \ 2 ^ Suc c" + assumes yltx_digits: "\t'. nth_digit y t' b \ nth_digit x t' b" + and "k\c" + and "c>0" + shows "nth_digit (x-y) t b \ k = (nth_digit x t b - nth_digit y t b) \ k" +proof - + have x: "nth_digit x t b \ k = x \ (Suc c * t + k)" + using digit_gen_pow2_reduct[of "k" "Suc c" "x" "t"] b_def `k\c` by auto + have y: "nth_digit y t b \ k = y \ (Suc c * t + k)" + using digit_gen_pow2_reduct[of "k" "Suc c" "y" "t"] b_def `k\c` by auto + + have "b > 1" using `c>0` b_def + by (metis Suc_1 lessI power_gt1) + then have yltx: "y \ x" using digitwise_leq yltx_digits by auto + + have "(nth_digit x t b - nth_digit y t b) \ k + = ((nth_digit x t b) \ k + (nth_digit y t b) \ k + + bin_narry (nth_digit x t b) (nth_digit y t b) k) mod 2" + using dif_digit_formula yltx_digits by auto + also have "... = (x \ (Suc c * t + k) + y \ (Suc c * t + k) + + bin_narry (nth_digit x t b) (nth_digit y t b) k) mod 2" + using x y by auto + also have "... = (x \ (Suc c * t + k) + y \ (Suc c * t + k) + + bin_narry x y (Suc c * t + k)) mod 2" + using narry_gen_pow2_reduct using assms(3) assms(4) b_def yltx_digits by auto + also have "... = nth_digit (x-y) t b \ k" + using digit_gen_pow2_reduct[of "k" "Suc c" "x-y" "t"] b_def `k\c` dif_digit_formula yltx + by auto + finally show ?thesis by auto +qed + +lemma block_subtractivity: + assumes "c > 0" + defines "b \ 2 ^ Suc c" + assumes block_wise_lt: "\t'. nth_digit y t' b \ nth_digit x t' b" + shows "nth_digit (x-y) t b = nth_digit x t b - nth_digit y t b" +proof - + have "k\c \ nth_digit (x-y) t b \ k = (x - y) \ (Suc c * t + k)" for k + using digit_gen_pow2_reduct[of "k" "Suc c" "x-y" "t"] b_def by auto + have "k\c \ nth_digit x t b \ k = x \ (Suc c * t + k)" for k + using digit_gen_pow2_reduct[of "k" "Suc c" "x" "t"] b_def by auto + have "k\c \ nth_digit y t b \ k = y \ (Suc c * t + k)" for k + using digit_gen_pow2_reduct[of "k" "Suc c" "y" "t"] b_def by auto + + have k_le: "k \ c \ nth_digit (x-y) t b \ k + = (nth_digit x t b - nth_digit y t b) \ k" for k + using assms digit_wise_block_subtractivity by auto + + have "nth_digit x t b - nth_digit y t b < b" using + nth_digit_bound b_def by (simp add: less_imp_diff_less) + hence diff: "k\Suc c \ (nth_digit x t b - nth_digit y t b) \ k = 0" for k + using nth_bit_def b_def aux_lt_implies_mask b_def by metis + + have "nth_digit (x-y) t b < b" using nth_digit_bound b_def by auto + hence "k\Suc c \ nth_digit (x-y) t b \ k = 0" for k + using nth_bit_def b_def aux_lt_implies_mask b_def by metis + with diff have k_gt: "k > c \ nth_digit (x-y) t b \ k + = (nth_digit x t b - nth_digit y t b) \ k" for k + by auto + from k_le k_gt have "nth_digit (x-y) t b \ k + = (nth_digit x t b - nth_digit y t b) \ k" for k by(cases "k>c"; auto) + thus ?thesis using digit_wise_equiv[of "nth_digit x t b - nth_digit y t b" + "nth_digit (x-y) t b"] by auto +qed + +lemma bitAND_nth_digit_commute: + assumes b_def: "b = 2^(Suc c)" + shows "nth_digit (x && y) t b = nth_digit x t b && nth_digit y t b" +proof - + { + fix k + assume k: "k < Suc c" + have prod: "nth_digit (x && y) t b \ k = (x && y) \ (Suc c * t + k)" + using digit_gen_pow2_reduct[of _ "Suc c" "x && y" "t"] b_def k by auto + moreover have x: "nth_digit x t b \ k = x \ (Suc c * t + k)" + using digit_gen_pow2_reduct[of _ "Suc c" "x"] b_def k by auto + moreover have y: "nth_digit y t b \ k = y \ (Suc c * t + k)" + using digit_gen_pow2_reduct[of _ "Suc c" "y"] b_def k by auto + moreover have "(x && y) \ (Suc c * t + k) = (x \ (Suc c * t + k)) * (y \ (Suc c * t + k))" + using bitAND_digit_mult by auto + + ultimately have "nth_digit (x && y) t b \ k + = nth_digit x t b \ k * nth_digit y t b \ k" + by auto + + also have "... = (nth_digit x t b && nth_digit y t b) \ k" + using bitAND_digit_mult by auto + + finally have "nth_digit (x && y) t b \ k + = (nth_digit x t b && nth_digit y t b) \ k" + by auto + } + + (* now also consider k \ Suc c *) + then have "nth_digit (x && y) t b \ k + = (nth_digit x t b && nth_digit y t b) \ k" for k + by (metis aux_lt_implies_mask b_def bitAND_digit_mult leI mult_eq_0_iff nth_digit_bound) + + then show ?thesis using b_def digit_wise_equiv[of "nth_digit (x && y) t b"] by auto +qed + +lemma bx_aux: + shows "b>1 \ nth_digit (b^x) t' b = (if x=t' then 1 else 0)" + by (cases "t' > x", auto simp: nth_digit_def) + (metis dvd_imp_mod_0 dvd_power leI less_SucI nat_neq_iff power_diff zero_less_diff) + + +context + fixes c b :: nat + assumes b_def: "b \ 2^(Suc c)" + assumes c_gt0: "c>0" +begin + +lemma b_gt1: "b>1" using c_gt0 b_def + using one_less_numeral_iff one_less_power semiring_norm(76) by blast + +text \Commutation relations with sums\ + +(* Commute outside, need assumption for nth_digit inside sum *) + +lemma finite_sum_nth_digit_commute: + fixes M :: nat + shows "\t. \k\M. nth_digit (fct k) t b < 2^c \ + \t. (\i=0..M. nth_digit (fct i) t b) < 2^c \ + nth_digit (\i=0..M. fct i) t b = (\i=0..M. (nth_digit (fct i) t b))" +proof (induct M arbitrary: t) + case 0 thus ?case by auto +next + case (Suc M) + + have assm1: "\t. \k\Suc M. nth_digit (fct k) t b < 2^c" + using Suc.prems by auto + have assm1_reduced: "\t. \k\M. nth_digit (fct k) t b < 2^c" + using assm1 by auto + have nocarry2: "\t. \k\Suc M. nth_digit (fct k) t b \ c = 0" + using assm1 nth_bit_def by auto + + have assm2: + "\t. nth_digit (fct (Suc M)) t b + (\i=0..M. nth_digit (fct i) t b) < 2^c" + using Suc.prems by (simp add: add.commute) + hence assm2_reduced: "\t. (\i=0..M. nth_digit (fct i) t b) < 2^c" + using Suc.prems(2) add_lessD1 by fastforce + + have IH: "\t. nth_digit (\i=0..M. fct i) t b + = (\i=0..M. nth_digit (fct i) t b)" + using assm1_reduced assm2_reduced Suc.hyps by blast + then have assm2_IH_commuted: "\t. nth_digit (\i=0..M. fct i) t b < 2^c" + using assm2_reduced by auto + hence nocarry3: "\t. nth_digit (\i=0..M. fct i) t b \ c = 0" + using aux_lt_implies_mask by blast + + have 1: "nth_digit (sum fct {0..M}) (t - 1) b \ c = 0" using nocarry3 by auto + have 2: "nth_digit (fct (Suc M)) (t - 1) b \ c = 0" using nocarry2 by auto + have 3: "nth_digit (sum fct {0..M}) t b \ c = 0" using nocarry3 by auto + have 4: "nth_digit (fct (Suc M)) t b \ c = 0" using nocarry2 by auto + from block_additivity show ?case using 1 2 3 4 c_gt0 Suc b_def + using IH by auto +qed + +lemma sum_nth_digit_commute_aux: + fixes g + defines SX_def: "SX \ \l m (fct :: nat \ nat). (\k = 0..m. if g l k then fct k else 0)" + assumes nocarry: "\t. \k\M. nth_digit (fct k) t b < 2^c" + and nocarry_sum: "\t. (SX l M (\k. nth_digit (fct k) t b)) < 2^c" + shows "nth_digit (SX l M fct) t b = SX l M (\k. nth_digit (fct k) t b)" +proof - + have aux: "nth_digit (if g l i then fct i else 0) t b + = (if g l i then nth_digit (fct i) t b else 0)" for i t + using aux1_digit_wise_gen_equiv b_gt1 by auto + + from nocarry have "\t. \k\M. nth_digit (if g l k then fct k else 0) t b < 2^c" + using aux by auto + + moreover from nocarry_sum have + "\t. (\i = 0..M. nth_digit (if g l i then fct i else 0) t b) < 2 ^ c" + unfolding SX_def by (auto simp: aux) + + ultimately have "nth_digit (\k = 0..M. if g l k then fct k else 0) t b + = (\k = 0..M. nth_digit (if g l k then fct k else 0) t b)" + using finite_sum_nth_digit_commute[where ?fct = "\k. if g l k then fct k else 0"] + by (simp add: SX_def) + moreover have "\k. nth_digit (if g l k then fct k else 0) t b + = (if g l k then nth_digit (fct k) t b else 0)" + by (auto simp: nth_digit_def) + ultimately show ?thesis by (auto simp: SX_def) +qed + +lemma sum_nth_digit_commute: + fixes g + defines SX_def: "SX \ \p l (fct :: nat \ nat). (\k = 0..length p - 1. if g l k then fct k else 0)" + assumes nocarry: "\t. \k\length p - 1. nth_digit (fct k) t b < 2^c" + and nocarry_sum: "\t. (SX p l (\k. nth_digit (fct k) t b)) < 2^c" + shows "nth_digit (SX p l fct) t b = SX p l (\k. nth_digit (fct k) t b)" +proof - + let ?m = "length p - 1" + + have "\t. (\k = 0..?m. if g l k then nth_digit (fct k) t b else 0) < 2^c" + using nocarry_sum unfolding SX_def by blast + + then have "nth_digit (\k = 0..length p - 1. if g l k then fct k else 0) t b + = (\k = 0..length p - 1. if g l k then nth_digit (fct k) t b else 0)" + using nocarry + using sum_nth_digit_commute_aux[where ?M = "length p - 1"] + by auto + + then show ?thesis using SX_def by auto +qed + +text \Commute inside, need assumption for all partial sums\ +lemma finite_sum_nth_digit_commute2: + fixes M :: nat + shows "\t. \k\M. nth_digit (fct k) t b < 2^c \ + \t. \m\M. nth_digit (\i=0..m. fct i) t b < 2^c \ + nth_digit (\i=0..M. fct i) t b = (\i=0..M. (nth_digit (fct i) t b))" +proof (induct M arbitrary: t) + case 0 thus ?case by auto +next + case (Suc M) + + have assm1: "\t. \k\Suc M. nth_digit (fct k) t b < 2^c" + using Suc.prems by auto + have assm1_reduced: "\t. \k\M. nth_digit (fct k) t b < 2^c" + using assm1 by auto + have nocarry2: "\t. \k\Suc M. nth_digit (fct k) t b \ c = 0" + using assm1 nth_bit_def by auto + + have assm2: + "\t. nth_digit (\i=0..M. (fct i)) t b < 2^c" + using Suc.prems by (simp add: add.commute) + hence nocarry3: "\t. nth_digit (\i=0..M. fct i) t b \ c = 0" + using aux_lt_implies_mask by blast + + have 1: "nth_digit (sum fct {0..M}) (t - 1) b \ c = 0" using nocarry3 by auto + have 2: "nth_digit (fct (Suc M)) (t - 1) b \ c = 0" using nocarry2 by auto + have 3: "nth_digit (sum fct {0..M}) t b \ c = 0" using nocarry3 by auto + have 4: "nth_digit (fct (Suc M)) t b \ c = 0" using nocarry2 by auto + from block_additivity show ?case using 1 2 3 4 c_gt0 Suc b_def by auto +qed + +lemma sum_nth_digit_commute_aux2: + fixes g + defines SX_def: "SX \ \l m (fct :: nat \ nat). (\k = 0..m. if g l k then fct k else 0)" + assumes nocarry: "\t. \k\M. nth_digit (fct k) t b < 2^c" + and nocarry_sum: "\t. \m\M. nth_digit (SX l m fct) t b < 2^c" + shows "nth_digit (SX l M fct) t b = SX l M (\k. nth_digit (fct k) t b)" +proof - + have aux: "nth_digit (if g l i then fct i else 0) t b + = (if g l i then nth_digit (fct i) t b else 0)" for i t + using aux1_digit_wise_gen_equiv b_gt1 by auto + + from nocarry have "\t. \k\M. nth_digit (if g l k then fct k else 0) t b < 2^c" + using aux by auto + + moreover from nocarry_sum have + "\t. \m\M. nth_digit (\i = 0..m. (if g l i then fct i else 0)) t b < 2 ^ c" + unfolding SX_def by (auto simp: aux) + + ultimately have "nth_digit (\k = 0..M. if g l k then fct k else 0) t b + = (\k = 0..M. nth_digit (if g l k then fct k else 0) t b)" + using finite_sum_nth_digit_commute2[where ?fct = "\k. if g l k then fct k else 0"] + by (simp add: SX_def) + moreover have "\k. nth_digit (if g l k then fct k else 0) t b + = (if g l k then nth_digit (fct k) t b else 0)" + by (auto simp: nth_digit_def) + ultimately show ?thesis by (auto simp: SX_def) +qed + +lemma sum_nth_digit_commute2: + fixes g p + defines SX_def: "SX \ \p l (fct :: nat \ nat). (\k = 0..length p - 1. if g l k then fct k else 0)" + assumes nocarry: "\t. \k\length p - 1. nth_digit (fct k) t b < 2^c" + and nocarry_sum: "\t. \m\length p - 1. nth_digit (SX (take (Suc m) p) l fct) t b < 2^c" + shows "nth_digit (SX p l fct) t b = SX p l (\k. nth_digit (fct k) t b)" +proof - + have "\m \ length p - 1. (SX (take (Suc m) p) l fct) = (\k = 0..m. if g l k then fct k else 0)" + unfolding SX_def + by (auto) (metis (no_types, lifting) One_nat_def diff_Suc_1 min_absorb2 min_diff) + hence "\t m. m \ length p - 1 \ + nth_digit (\k = 0..m. if g l k then fct k else 0) t b < 2 ^ c" + using nocarry_sum by auto + + then have "nth_digit (\k = 0..length p - 1. if g l k then fct k else 0) t b + = (\k = 0..length p - 1. if g l k then nth_digit (fct k) t b else 0)" + using nocarry + using sum_nth_digit_commute_aux2[where ?M = "length p - 1" and ?fct = "fct" and ?g = g] + by blast + + then show ?thesis using SX_def by auto +qed + +end + +end \ No newline at end of file diff --git a/thys/DPRM_Theorem/Register_Machine/MachineEquations.thy b/thys/DPRM_Theorem/Register_Machine/MachineEquations.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Register_Machine/MachineEquations.thy @@ -0,0 +1,62 @@ +section \Arithmetization of Register Machines\ + +subsection \A first definition of the arithmetizing equations\ + +theory MachineEquations + imports MultipleStepRegister MultipleStepState MachineMasking +begin + +definition mask_equations :: "nat \ (register \ nat) \ (register \ nat) + \ nat \ nat \ nat \ nat \ bool" (* 4.15, 4.17 and 4.20 *) + where "(mask_equations n r z c d e f) = ((\l d) + \ (\l e) + \ (\l (register \ nat) \ (register \ nat) \ (state \ nat) + \ nat \ nat \ nat \ nat \ bool" where + "(reg_equations p r z s b a n q) = ( + \ \4.22\ (\l>0. l < n \ r l = b*r l + b*\R+ p l (\k. s k) - b*\R- p l (\k. s k && z l)) + \ \ \4.23\ ( r 0 = a + b*r 0 + b*\R+ p 0 (\k. s k) - b*\R- p 0 (\k. s k && z 0)) + \ (\l \Extra equation not in Matiyasevich's book. + Needed to show that all registers are empty at time q\" + +(* STATE EQUATIONS -- todo from here on *) +definition state_equations :: "program \ (state \ nat) \ (register \ nat) \ + nat \ nat \ nat \ nat \ bool" where + "state_equations p s z b e q m = ( +\ \4.24\ (\d>0. d\m \ s d = b*\S+ p d (\k. s k) + b*\S- p d (\k. s k && z (modifies (p!k))) + + b*\S0 p d (\k. s k && (e - z (modifies (p!k))))) + \ \ \4.25\ ( s 0 = 1 + b*\S+ p 0 (\k. s k) + b*\S- p 0 (\k. s k && z (modifies (p!k))) + + b*\S0 p 0 (\k. s k && (e - z (modifies (p!k))))) + \ \ \4.27\ (s m = b^q) + \ (\k\m. s k \ e) \ (\k \these equations are not from the book\ + \ (\M\m. (\k\M. s k) \ e) \ \this equation is added, too\ )" + +(* The following two equations do not appear in Matiyasevich's book, this duplicated definition + (note that they are also included just above) is, for now, only a reminder. *) +definition state_unique_equations :: "program \ (state\nat) \ nat \ nat \bool" where + "state_unique_equations p s m e = ((\k=0..m. s k) \ e \ (\k\m. s k \ e))" + + +(* RM CONSTANTS FOR DEFINITION AND CLARIFY OF EQUATIONS *) +definition rm_constants :: "nat \ nat \ nat \ nat \ nat \ nat \ nat \ bool" where + "rm_constants q c b d e f a = ( + \ \4.14\ (b = B c) + \ \ \4.16\ (d = D q c b) + \ \ \4.18\ (e = E q b) \ \4.19 left out (compare book)\ + \ \ \4.21\ (f = F q c b) + \ \ \extra equation not in the book\ c > 0 + \ \ \4.26\ (a < 2^c) \ (q>0))" + +definition rm_equations_old :: "program \ nat \ nat \ nat \ bool" where + "rm_equations_old p q a n = ( + \ b c d e f :: nat. + \ r z :: register \ nat. + \ s :: state \ nat. + mask_equations n r z c d e f + \ reg_equations p r z s b a n q + \ state_equations p s z b e q (length p - 1) + \ rm_constants q c b d e f a)" + +end diff --git a/thys/DPRM_Theorem/Register_Machine/MachineMasking.thy b/thys/DPRM_Theorem/Register_Machine/MachineMasking.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Register_Machine/MachineMasking.thy @@ -0,0 +1,659 @@ +subsection \Masking properties\ + +theory MachineMasking + imports RegisterMachineSimulation "../Diophantine/Binary_And" +begin + +(* [D] 4.18 *) +definition E :: "nat \ nat \ nat" where + "(E q b) = (\t = 0..q. b^t)" + +lemma e_geom_series: + assumes "b \ 2" + shows "(E q b = e) \ ((b-1) * e = b^(Suc q) - 1 )" (is "?P \ ?Q") +proof- + have "sum ((^) (int b)) {.. nat \ nat \ nat" where + "(D q c b) = (\t = 0..q. (2^c - 1) * b^t)" + +lemma d_geom_series: + assumes "b = 2^(Suc c)" + shows "(D q c b = d) \ ((b-1) * d = (2^c - 1) * (b^(Suc q) - 1))" (is "?P \ ?Q") +proof- + have "D q c b = (2^c - 1) * E q b" by (auto simp: E_def D_def sum_distrib_left sum_distrib_right) + moreover have "b \ 2" using assms by fastforce + ultimately show ?thesis by (smt e_geom_series mult.left_commute mult_cancel_left) +qed + +(* [D] 4.21 *) +definition F :: "nat \ nat \ nat \ nat" where + "(F q c b) = (\t = 0..q. 2^c * b^t)" + +lemma f_geom_series: + assumes "b = 2^(Suc c)" + shows "(F q c b = f) \ ( (b-1) * f = 2^c * (b^(Suc q) - 1) )" +proof- + have "F q c b = 2^c * E q b" by (auto simp: E_def F_def sum_distrib_left sum_distrib_right) + moreover have "b \ 2" using assms by fastforce + ultimately show ?thesis by (smt e_geom_series mult.left_commute mult_cancel_left) +qed + +(* AUX LEMMAS *) +lemma aux_lt_implies_mask: + assumes "a < 2^k" + shows "\r\k. a \ r = 0" + using nth_bit_def assms apply auto +proof - (* found by e *) + fix r :: nat + assume a1: "a < 2 ^ k" + assume a2: "k \ r" + have "a div 2 ^ k = 0" + using a1 Euclidean_Division.div_eq_0_iff by blast + then have "2 = (0::nat) \ a < 2 ^ r" + using a2 by (metis (no_types) div_le_mono nat_zero_less_power_iff neq0_conv not_le power_diff) + then show "a div 2 ^ r mod 2 = 0" + by simp +qed + +lemma lt_implies_mask: + fixes a b :: nat + assumes "\k. a < 2^k \ (\r b" +proof - + obtain k where assms: "a < 2^k \ (\rr r \ b \ r" using nth_bit_bounded + by (simp add: \a < 2 ^ k \ (\r r = 1)\) + hence k2: "\r\k. a \ r = 0" using aux_lt_implies_mask assms by auto + show ?thesis using masks_leq_equiv + by auto (metis k1 k2 le0 not_less) +qed + +lemma mask_conversed_shift: + fixes a b k :: nat + assumes asm: "a \ b" + shows "a * 2^k \ b * 2^k" +proof - + have shift: "x \ y \ 2*x \ 2*y" for x y by (induction x; auto) + have "a * 2 ^ k \ b * 2 ^ k \ 2 * (a * 2 ^ k) \ 2 * (b * 2 ^ k)" for k + using shift[of "a*2^k" "b*2^k"] by auto + thus ?thesis by (induction k; auto simp: asm shift algebra_simps) +qed + +lemma base_summation_bound: + fixes c q :: nat + and f :: "(nat \ nat)" + +defines b: "b \ B c" +assumes bound: "\t. f t < 2 ^ Suc c - (1::nat)" + +shows "(\t = 0..q. f t * b^t) < b^(Suc q)" +proof (induction q) + case 0 + then show ?case using B_def b bound less_imp_diff_less not_less_eq + by auto blast +next + case (Suc q) + have "(\t = 0..Suc q. f t * b ^ t) = f (Suc q) * b ^ (Suc q) + (\t = 0..q. f t * b ^ t)" + by (auto cong: sum.cong) + also have "... < (f (Suc q) + 1) * b ^ (Suc q)" + using Suc.IH by auto + also have "... < b * b ^ (Suc q)" + by (metis bound b less_diff_conv B_def mult_less_cancel2 zero_less_numeral zero_less_power) + finally show ?case by auto +qed + +lemma mask_conserved_sum: + fixes y c q :: nat + and x :: "(nat \ nat)" + +defines b: "b \ B c" +assumes mask: "\t. x t \ y" +assumes xlt: "\t. x t \ 2 ^ c - Suc 0" +assumes ylt: "y \ 2 ^ c - Suc 0" + +shows "(\t = 0..q. x t * b^t) \ (\t = 0..q. y * b^t)" +proof (induction q) + case 0 + then show ?case + using mask by auto +next + case (Suc q) + + have xb: "\t. x t < 2^Suc c - Suc 0" + using xlt + by (smt Suc_pred leI le_imp_less_Suc less_SucE less_trans n_less_m_mult_n numeral_2_eq_2 + power.simps(2) zero_less_numeral zero_less_power) + have yb: "y < 2^c" + using ylt b B_def leI order_trans by fastforce + + have sumxlt: "(\t = 0..q. x t * b ^ t) < b^(Suc q)" + using base_summation_bound xb b B_def by auto + have sumylt: "(\t = 0..q. y * b ^ t) < b^(Suc q)" + using base_summation_bound yb b B_def by auto + + have "((\t = 0..Suc q. x t * b ^ t) \ (\t = 0..Suc q. y * b ^ t)) + = (x (Suc q) * b^Suc q + (\t = 0..q. x t * b ^ t) \ + y * b^Suc q + (\t = 0..q. y * b ^ t))" + by (auto simp: atLeast0_lessThan_Suc add.commute) + also have "... = (x (Suc q) * b^Suc q \ y * b^Suc q) + \ (\t = 0..q. x t * b ^ t) \ (\t = 0..q. y * b ^ t)" + using mask_linear[where ?t = "Suc c * Suc q"] sumxlt sumylt Suc.IH b B_def + apply auto + apply (smt mask mask_conversed_shift power_Suc power_mult power_mult_distrib) + by (smt mask mask_linear power_Suc power_mult power_mult_distrib) + finally show ?case using mask_linear Suc.IH B_def + by (metis (no_types, lifting) b mask mask_conversed_shift power_mult) +qed + +lemma aux_powertwo_digits: + fixes k c :: nat + assumes "k < c" + shows "nth_bit (2^c) k = 0" +proof - + have h: "(2::nat) ^ c div 2 ^ k = 2 ^ (c - k)" + by (simp add: assms less_imp_le power_diff) + thus ?thesis + by (auto simp: h nth_bit_def assms) +qed + +lemma obtain_digit_rep: + fixes x c :: nat + shows "x && 2^c = (\t 2^c" by (simp add: lm0245) + hence "x && 2^c \ 2^c" by (simp add: masks_leq) + hence h: "x && 2^c < 2^Suc c" + by (smt Suc_lessD le_neq_implies_less lessI less_trans_Suc n_less_m_mult_n numeral_2_eq_2 + power_Suc zero_less_power) + have "\t. (x && 2^c) \ t = (nth_bit x t) * (nth_bit (2^c) t)" + using bitAND_digit_mult by auto + then show ?thesis using h digit_sum_repr[of "(x && 2^c)" "Suc c"] + by (auto) (simp add: mult.commute semiring_normalization_rules(19)) +qed + +lemma nth_digit_bitAND_equiv: + fixes x c :: nat + shows "2^c * nth_bit x c = (x && 2^c)" +proof - + have d1: "nth_bit (2^c) c = 1" + using nth_bit_def by auto + + moreover have "x && 2^c = (2::nat)^c * (x \ c) * (((2::nat)^c) \ c) + + (\t t) * (((2::nat)^c) \ t))" + using obtain_digit_rep by (auto cong: sum.cong) + + moreover have "(\t x" +assumes "x < 2 ^ Suc c" + +shows "nth_bit x c = 1" +proof - + obtain b where b: "x = 2^c + b" + using assms(1) le_Suc_ex by auto + have bb: "b < 2^c" + using assms(2) b by auto + have "(2 ^ c + b) div 2 ^ c mod 2 = (1 + b div 2 ^ c) mod 2" + by (auto simp: div_add_self1) + also have "... = 1" + by (auto simp: bb) + finally show ?thesis + by (simp only: nth_bit_def b) +qed + +lemma aux_bitAND_distrib: "2 * (a && b) = (2 * a) && (2 * b)" + by (induct a b rule: bitAND_nat.induct; auto) + +lemma bitAND_distrib: "2^c * (a && b) = (2^c * a) && (2^c * b)" +proof (induction c) + case 0 + then show ?case by auto +next + case (Suc c) + have "2 ^ Suc c * (a && b) = 2 * (2 ^ c * (a && b))" by auto + also have "... = 2 * ((2^c * a) && (2^c * b))" using Suc.IH by auto + also have "... = ((2^Suc c * a) && (2^Suc c * b))" + using aux_bitAND_distrib[of "2^c * a" "2^c * b"] + by (auto simp add: ab_semigroup_mult_class.mult_ac(1)) + finally show ?case by auto +qed + +lemma bitAND_linear_sum: + fixes x y :: "nat \ nat" + and c :: nat + and q :: nat + +defines b: "b == 2 ^ Suc c" + +assumes xb: "\t. x t < 2 ^ Suc c - 1" +assumes yb: "\t. y t < 2 ^ Suc c - 1" + +shows "(\t = 0..q. (x t && y t) * b^t) = + (\t = 0..q. x t * b^t) && (\t = 0..q. y t * b^t)" +proof (induction q) + case 0 + then show ?case + by (auto simp: b B_def) +next + case (Suc q) + have "(\t = 0..Suc q. (x t && y t) * b ^ t) = (x (Suc q) && y (Suc q)) * b ^ Suc q + + (\t = 0..q. (x t && y t) * b ^ t)" + by (auto cong: sum.cong) + + moreover have h0: "(x (Suc q) && y (Suc q)) * b ^ Suc q + = (x (Suc q) * b^Suc q) && (y (Suc q) * b^Suc q)" + using b bitAND_distrib by (auto) (smt mult.commute power_Suc power_mult) + + moreover have h1: "(\t = 0..q. (x t && y t) * b ^ t) + = (\t = 0..q. x t * b^t) && (\t = 0..q. y t * b^t)" + using Suc.IH by auto + + ultimately have h2: "(\t = 0..Suc q. (x t && y t) * b ^ t) + = ((x (Suc q) * b^Suc q) && (y (Suc q) * b^Suc q)) + + ((\t = 0..q. x t * b^t) && (\t = 0..q. y t * b^t))" + by auto + + have sumxb: "(\t = 0..q. x t * b ^ t) < b ^ Suc q" + using base_summation_bound xb b B_def by auto + have sumyb: "(\t = 0..q. y t * b ^ t) < b ^ Suc q" + using base_summation_bound yb b B_def by auto + + have h3: "((x (Suc q) * b^Suc q) && (y (Suc q) * b^Suc q)) + + ((\t = 0..q. x t * b^t) && (\t = 0..q. y t * b^t)) + = ((\t = 0..q. x t * b^t) + x (Suc q) * b^Suc q) + && ((\t = 0..q. y t * b^t) + y (Suc q) * b^Suc q)" + using sumxb sumyb bitAND_linear h2 h0 + by (auto) (smt add.commute b power_Suc power_mult) + + thus ?case using h2 by (auto cong: sum.cong) +qed + +lemma dmask_aux0: + fixes x :: nat + assumes "x > 0" + shows "(2 ^ x - Suc 0) div 2 = 2 ^ (x - 1) - Suc 0" +proof - + have 0: "(2^x - Suc 0) div 2 = (2^x - 2) div 2" + by (smt Suc_diff_Suc Suc_pred assms dvd_power even_Suc even_Suc_div_two nat_power_eq_Suc_0_iff + neq0_conv numeral_2_eq_2 zero_less_diff zero_less_power) + (* can do manual parity distinction *) + moreover have divides: "(2::nat) dvd 2^x" + by (simp add: assms dvd_power[of x "2::nat"]) + moreover have "(2^x - 2::nat) div 2 = 2^x div 2 - 2 div 2" + using div_plus_div_distrib_dvd_left[of "2" "2^x" "2"] divides + by auto + moreover have "... = 2 ^ (x - 1) - Suc 0" + by (simp add: Suc_leI assms power_diff) + ultimately have 1: "(2 ^ x - Suc 0) div 2 = 2 ^ (x - 1) - Suc 0" + by (smt One_nat_def) + thus ?thesis by simp +qed + +lemma dmask_aux: + fixes c :: nat + shows "d < c \ (2^c - Suc 0) div 2^d = 2 ^ (c - d) - Suc 0" +proof (induction d) + case 0 + then show ?case by (auto) +next + case (Suc d) + have d: "d < c" using Suc.prems by auto + have "(2 ^ c - Suc 0) div 2 ^ Suc d = (2 ^ c - Suc 0) div 2 ^ d div 2" + by (auto) (metis mult.commute div_mult2_eq) + also have "... = (2 ^ (c - d) - Suc 0) div 2" + by (subst Suc.IH) (auto simp: d) + also have "... = 2 ^ (c - Suc d) - Suc 0" + apply (subst dmask_aux0[of "c - d"]) + using d by (auto) + finally show ?case by auto +qed + + +(* REGISTERS *) +lemma register_cells_masked: + fixes l :: register + and t :: nat + and ic :: configuration + and p :: program + +assumes cells_bounded: "cells_bounded ic p c" +assumes l: "l < length (snd ic)" + +shows "R ic p l t \ 2^c - 1" +proof - + have a: "R ic p l t \ 2^c - 1" using cells_bounded less_Suc_eq_le + using l by fastforce + have b: "r < c \ nth_bit (2^c - 1) r = 1" for r + apply (auto simp: nth_bit_def) + apply (subst dmask_aux) + apply (auto) + by (metis Suc_pred dvd_power even_Suc mod_0_imp_dvd not_mod2_eq_Suc_0_eq_0 + zero_less_diff zero_less_numeral zero_less_power) + show ?thesis using lt_implies_mask cells_bounded l + by (auto) (metis One_nat_def b) +qed + +lemma lm04_15_register_masking: + fixes c :: nat + and l :: register + and ic :: configuration + and p :: program + and q :: nat + +defines "b == B c" +defines "d == D q c b" + +assumes cells_bounded: "cells_bounded ic p c" +assumes l: "l < length (snd ic)" + +defines "r == RLe ic p b q" + +shows "r l \ d" +proof - + have "\t. R ic p l t \ 2^c - 1" using cells_bounded l + by (rule register_cells_masked) + hence rmasked: "\t. R ic p l t \ 2^c - 1" + by (intro allI) + + have rlt: "\t. R ic p l t \ 2^c - 1" + using cells_bounded less_Suc_eq_le l by fastforce + + have rlmasked: "(\t = 0..q. R ic p l t * b^t) \ (\t = 0..q. (2^c - 1) * b^t)" + using rmasked rlt b_def B_def mask_conserved_sum by (auto) + + thus ?thesis + by (auto simp: r_def d_def D_def RLe_def mult.commute cong: sum.cong) +qed + +(* ZERO INDICATORS *) +lemma zero_cells_masked: + fixes l :: register + and t :: nat + and ic :: configuration + and p :: program + +assumes l: "l < length (snd ic)" + +shows "Z ic p l t \ 1" +proof - + have "nth_bit 1 0 = 1" by (auto simp: nth_bit_def) + thus ?thesis apply (auto) apply (rule lt_implies_mask) + by (metis (full_types) One_nat_def Suc_1 Z_bounded less_Suc_eq_le less_one power_one_right) +qed + +lemma lm04_15_zero_masking: + fixes c :: nat + and l :: register + and ic :: configuration + and p :: program + and q :: nat + +defines "b == B c" +defines "e == E q b" + +assumes cells_bounded: "cells_bounded ic p c" +assumes l: "l < length (snd ic)" +assumes c: "c > 0" + +defines "z == ZLe ic p b q" + +shows "z l \ e" +proof - + have "\t. Z ic p l t \ 1" using l + by (rule zero_cells_masked) + hence zmasked: "\t. Z ic p l t \ 1" + by (intro allI) + + have zlt: "\t. Z ic p l t \ 2 ^ c - 1" + using cells_bounded less_Suc_eq_le by fastforce + + have 1: "(1::nat) \ 2 ^ c - 1" using c + by (simp add: Nat.le_diff_conv2 numeral_2_eq_2 self_le_power) + + have rlmasked: "(\t = 0..q. Z ic p l t * b^t) \ (\t = 0..q. 1 * b^t)" + using zmasked zlt 1 b_def B_def mask_conserved_sum[of "Z ic p l" 1] + by (auto) + + thus ?thesis + by (auto simp: z_def e_def E_def ZLe_def mult.commute cong: sum.cong) +qed + +(* Relation between zero indicator and register *) +lemma lm04_19_zero_register_relations: + fixes c :: nat + and l :: register + and t :: nat + and ic :: configuration + and p :: program + +assumes cells_bounded: "cells_bounded ic p c" +assumes l: "l < length (snd ic)" + +defines "z == Z ic p" +defines "r == R ic p" + +shows "2^c * z l t = (r l t + 2^c - 1) && 2^c" +proof - + have a1: "R ic p l t \ 0 \ 2^c \ R ic p l t + 2^c - 1" + by auto + have a2: "R ic p l t + 2^c - 1 < 2^Suc c" using cells_bounded + by (simp add: l less_imp_diff_less) + + have "Z ic p l t = nth_bit (R ic p l t + 2^c - 1) c" + apply (cases "R ic p l t = 0") + subgoal by (auto simp add: Z_def R_def nth_bit_def) + subgoal using cells_bounded bitAND_single_digit a1 a2 Z_def + by auto + done + + also have "2^c * nth_bit (R ic p l t + 2^c - 1) c = ((R ic p l t + 2^c - 1) && 2^c)" + using nth_digit_bitAND_equiv by auto + + finally show ?thesis by (auto simp: z_def r_def) +qed + +lemma lm04_20_zero_definition: + fixes c :: nat + and l :: register + and ic :: configuration + and p :: program + and q :: nat + +defines "b == B c" +defines "f == F q c b" +defines "d == D q c b" + +assumes cells_bounded: "cells_bounded ic p c" +assumes l: "l < length (snd ic)" + +assumes c: "c > 0" + +defines "z == ZLe ic p b q" +defines "r == RLe ic p b q" + +shows "2^c * z l = (r l + d) && f" +proof - + have "\t. 2^c * Z ic p l t = (R ic p l t + 2^c - 1) && 2^c" + by (rule lm04_19_zero_register_relations cells_bounded l) + + hence raw_sums: "(\t = 0..q. 2^c * Z ic p l t * b^t) + = (\t = 0..q. ((R ic p l t + 2^c - 1) && 2^c) * b^t)" + by auto + + have "(\t = 0..q. 2^c * Z ic p l t * b^t) = 2^c * (\t = 0..q. Z ic p l t * b^t)" + by (auto simp: sum_distrib_left mult.assoc cong: sum.cong) + also have "... = 2^c * z l" + by (auto simp: z_def ZLe_def mult.commute) + finally have lhs: "(\t = 0..q. 2^c * Z ic p l t * b^t) = 2^c * z l" + by auto + + have "(\t = 0..q. (R ic p l t + (2^c - 1)) * b^t) + = (\t = 0..q. R ic p l t * b^t + (2^c - 1) * b^t)" + apply (rule sum.cong) + apply (auto simp: add.commute mult.commute) + subgoal for x using distrib_left[of "b^x" "R ic p l x" "2^c - 1"] by (auto simp: algebra_simps) + done + also have "... = (\t = 0..q. (R ic p l t * b^t)) + (\t = 0..q. (2^c - 1) * b^t)" + by (rule sum.distrib) + also have "... = r l + d" + by (auto simp: r_def RLe_def d_def D_def mult.commute) + finally have split_sums: "(\t = 0..q. (R ic p l t + (2^c - 1)) * b^t) = r l + d" + by auto + + have a1: "(2::nat) ^ c < (2::nat) ^ Suc c - 1" using c by (induct c, auto, fastforce) + have a2: "\t. R ic p l t + 2 ^ c - 1 \ 2 ^ Suc c" using cells_bounded B_def + by (simp add: less_imp_diff_less l) (simp add: Suc_leD l less_imp_le_nat numeral_Bit0) + have "(\t = 0..q. ((R ic p l t + 2^c - 1) && 2^c) * b^t) + = (\t = 0..q. (R ic p l t + 2^c - 1) * b^t) && (\t = 0..q. 2^c * b^t)" + using bitAND_linear_sum[of "\t. R ic p l t + 2^c - 1" "c" "\t. 2^c"] + cells_bounded b_def B_def a1 a2 + apply auto + by (smt One_nat_def Suc_less_eq Suc_pred a1 add.commute add_gr_0 l mult_2 + nat_add_left_cancel_less power_Suc zero_less_numeral zero_less_power) + also have "... = (\t = 0..q. (R ic p l t + 2^c - 1) * b^t) && f" + by (auto simp: f_def F_def) + also have "... = (r l + d) && f" using split_sums + by auto + finally have rhs: "(\t = 0..q. ((R ic p l t + 2^c - 1) && 2^c) * b^t) = (r l + d) && f" + by auto + + show ?thesis using raw_sums lhs rhs + by auto +qed + +lemma state_mask: +fixes c :: nat + and l :: register + and ic :: configuration + and p :: program + and q :: nat + and a :: nat + +defines "b \ B c" + and "m \ length p - 1" + +defines "e \ E q b" + +assumes is_val: "is_valid_initial ic p a" + and q: "q > 0" + and "c > 0" + +assumes terminate: "terminates ic p q" + shows "SKe ic p b q k \ e" +proof - + have "1 \ 2 ^ c - Suc 0" using \c>0\ by (metis One_nat_def Suc_leI one_less_numeral_iff + one_less_power semiring_norm(76) zero_less_diff) + have Smask: "S ic p k t \ 1" for t by (simp add: S_def) + have Sbound: "S ic p k t \ 2 ^ c - Suc 0" for t using \1\2^c-Suc 0\ by (simp add: S_def) + have rlmasked: "(\t = 0..q. S ic p k t * b^t) \ (\t = 0..q. 1 * b^t)" + using b_def B_def Smask Sbound mask_conserved_sum[of "S ic p k" 1] \1 \ 2^c-Suc 0\ by auto + + thus ?thesis using SKe_def e_def E_def by (auto simp: mult.commute) +qed + +lemma state_sum_mask: +fixes c :: nat + and l :: register + and ic :: configuration + and p :: program + and q :: nat + and a :: nat + +defines "b \ B c" + and "m \ length p - 1" + +defines "e \ E q b" + +assumes is_val: "is_valid_initial ic p a" + and q: "q > 0" + and "c > 0" + and "b > 1" + +assumes "M\m" + +assumes terminate: "terminates ic p q" +shows "(\k\M. SKe ic p b q k) \ e" +proof - + have e_aux: "nth_digit e t b = (if t\q then 1 else 0)" for t + unfolding e_def E_def b_def B_def + using `b>1` b_def nth_digit_gen_power_series[of "\k. Suc 0" "c" "q"] + by (auto simp: b_def B_def) + + have state_unique: "\k\m. S ic p k t = 1 \ (\j\k. S ic p j t = 0)" for t + using S_def by (induction t, auto) + + have h1: "\t. nth_digit (\k\M. SKe ic p b q k) t b \ (if t\q then 1 else 0)" + proof - { + fix t + have aux_bound_1: "(\k\M. S ic p k t') \ 1" for t' + proof (cases "\k\M. S ic p k t' = 1") + case True + then obtain k where k: "k\M \ S ic p k t' = 1" by auto + moreover have "\j\M. j \ k \ S ic p j t' = 0" + using state_unique `M<=m` k S_def + by (auto) (presburger) + ultimately have "(\k\M. S ic p k t') = 1" + using S_def by auto + then show ?thesis + by auto + next + case False + then show ?thesis using S_bounded + by (auto) (metis (no_types, lifting) S_def atMost_iff eq_imp_le le_SucI sum_nonpos) + qed + hence aux_bound_2: "\t'. (\k\M. S ic p k t') < 2^c" + by (metis Suc_1 `c>0` le_less_trans less_Suc_eq one_less_power) + + have h2: "(\k\M. SKe ic p b q k) = (\t = 0..q. \k\M. b ^ t * S ic p k t)" + unfolding SKe_def using sum.swap by auto + hence "(\k\M. SKe ic p b q k) = (\t = 0..q. b^t * (\k\M. S ic p k t))" + unfolding SKe_def by (simp add: sum_distrib_left) + + hence "nth_digit (\k\M. SKe ic p b q k) t b = (if t\q then (\k\M. S ic p k t) else 0)" + using `c>0` aux_bound_2 h2 unfolding SKe_def + using nth_digit_gen_power_series[of "\t. (\k\M. S ic p k t)" "c" "q" "t"] + by (smt B_def Groups.mult_ac(2) assms(7) aux_bound_1 b_def le_less_trans sum.cong) + hence "nth_digit (\k\M. SKe ic p b q k) t b \ (if t\q then 1 else 0)" + using aux_bound_1 by auto + } thus ?thesis by auto + qed + moreover have "\t>q. nth_digit (\k\M. SKe ic p b q k) t b = 0" + by (metis (full_types) h1 le_0_eq not_less) + ultimately have "\t. \ik\M. SKe ic p b q k) t b \ i + \ nth_digit e t b \ i" + using aux_lt_implies_mask linorder_neqE_nat e_aux + by (smt One_nat_def le_0_eq le_SucE less_or_eq_imp_le nat_zero_less_power_iff + numeral_2_eq_2 zero_less_Suc) + + hence "\t. \ik\M. SKe ic p b q k) \ (Suc c * t + i) \ e \ (Suc c * t + i)" + using digit_gen_pow2_reduct[where ?c = "Suc c" and ?a = "(\k\M. SKe ic p b q k)"] + using digit_gen_pow2_reduct[where ?c = "Suc c" and ?a = e] + by (simp add: b_def B_def) + moreover have "\j. \t i. i < Suc c \ j = (Suc c * t + i)" + using mod_less_divisor zero_less_Suc + by (metis add.commute mod_mult_div_eq) + ultimately have "\j. (\k\M. SKe ic p b q k) \ j \ e \ j" + by metis + + thus ?thesis + using masks_leq_equiv by auto +qed + +end diff --git a/thys/DPRM_Theorem/Register_Machine/MultipleStepRegister.thy b/thys/DPRM_Theorem/Register_Machine/MultipleStepRegister.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Register_Machine/MultipleStepRegister.thy @@ -0,0 +1,456 @@ +subsection \Multiple step relations\ + +subsubsection \Registers\ + +theory MultipleStepRegister + imports SingleStepRegister +begin + +lemma lm04_22_multiple_register: + fixes c :: nat + and l :: register + and ic :: configuration + and p :: program + and q :: nat + and a :: nat + + defines "b == B c" + and "m == length p" + and "n == length (snd ic)" + +assumes is_val: "is_valid_initial ic p a" +assumes c_gt_cells: "cells_bounded ic p c" +assumes l: "l < n" +and "0 < l" +and q: "q > 0" + +assumes terminate: "terminates ic p q" + +assumes c: "c > 1" + + defines "r == RLe ic p b q" + and "z == ZLe ic p b q" + and "s == SKe ic p b q" + +shows "r l = b * r l + + b * (\R+ p l s) + - b * (\R- p l (\k. z l && s k))" +proof - + have 0: "snd ic ! l = 0" using assms(4, 7) by (cases "ic"; auto simp add: is_valid_initial_def) + + have "b^(Suc t) * (\R- p l (\k. (Z ic p l t) * S ic p k t)) \ b^(Suc t) * R ic p l t " for t + proof (cases "t=0") + case True + hence "R ic p l 0 = 0" by (auto simp add: 0 R_def) + thus ?thesis by (auto simp add: True Z_def sum_rsub.simps) + next + case False + define cs where "cs \ fst (steps ic p t)" + have sub: "(\R- p l (\k. Z ic p l t * S ic p k t)) + = (if issub (p!cs) \ l = modifies (p!cs) then Z ic p l t else 0)" + using single_step_sub Z_def R_def is_val l n_def cs_def by auto + show ?thesis using sub by (auto simp add: sum_rsub.simps R_def Z_def) + qed + + from this have positive: "b^(Suc t) * (\R- p l (\k. (Z ic p l t) * S ic p k t)) + \ b^(Suc t) * R ic p l t + + b^(Suc t) * (\R+ p l (\k. S ic p k t))" for t + by (auto simp add: Nat.trans_le_add1) + + have commute_add: "(\t=0..q-1. \R+ p l (\k. b^t * S ic p k t)) + = \R+ p l (\k. \t=0..q-1. (b^t * S ic p k t))" + using sum_radd_commutative[of "p" "l" "\k t. b^t * S ic p k t " "q-1"] by auto + + have r_q: "l R ic p l q = 0" + using terminate terminates_def correct_halt_def by (auto simp: n_def R_def) + hence z_q: "l Z ic p l q = 0" + using terminate terminates_def correct_halt_def by (auto simp: Z_def) + have "\k ishalt (p!k)" + using is_val is_valid_initial_def[of "ic" "p" "a"] is_valid_def[of "ic" "p"] + program_includes_halt.simps by blast + hence s_q: "\k < length p - 1. S ic p k q = 0" + using terminate terminates_def correct_halt_def S_def by auto + + from r_q have rq: "(\x = 0..q - 1. int b ^ x * int (snd (steps ic p x) ! l)) = + (\x = 0..q. int b ^ x * int (snd (steps ic p x) ! l))" + by (auto simp: r_q R_def l; + smt Suc_pred mult_0_right of_nat_0 of_nat_mult power_mult_distrib q sum.atLeast0_atMost_Suc zero_power) + + have "(\t = 0..q - 1. b ^ t * (Z ic p l t * S ic p k t)) + + (b^(Suc (q-1)) * (Z ic p l (Suc (q-1)) * S ic p k (Suc (q-1)))) + = (\t = 0..Suc (q-1). b ^ t * (Z ic p l t * S ic p k t)) " for k + using comm_monoid_add_class.sum.atLeast0_atMost_Suc by auto + hence zq: "(\t = 0..q - 1. b ^ t * (Z ic p l t * S ic p k t)) + = (\t = 0..q. b ^ t * (Z ic p l t * S ic p k t)) " for k + using z_q q l by auto + + have "(if isadd (p ! k) \ l = modifies (p ! k) then \t = 0..q - Suc 0. b ^ t * S ic p k t else 0) + = (if isadd (p ! k) \ l = modifies (p ! k) then \t = 0..q. b ^ t * S ic p k t else 0)" for k + proof (cases "p!k") + case (Add x11 x12) + have sep: "(\t = 0..q-1. b ^ t * S ic p k t) + b^q * S ic p k q + = (\t = 0..(Suc (q-1)). b ^ t * S ic p k t)" + using comm_monoid_add_class.sum.atLeast0_atMost_Suc[of "\t. b^t * S ic p k t" "q-1"] q + by auto + have "ishalt (p ! (fst (steps ic p q)))" + using terminates_halt_state[of "ic" "p"] is_val terminate by auto + hence "S ic p k q = 0" using Add S_def[of "ic" "p" "k" "q"] by auto + with sep q have "(\t = 0..q - Suc 0. b ^ t * S ic p k t) = (\t = 0..q. b ^ t * S ic p k t)" + by auto + thus ?thesis by auto + next + case (Sub x21 x22 x23) + then show ?thesis by auto + next + case Halt + then show ?thesis by auto + qed + + hence add_q: "\R+ p l (\k. \t=0..(q-1). b^t * S ic p k t) + = \R+ p l (\k. \t=0..q. b^t * S ic p k t)" + using sum_radd.simps single_step_add[of "ic" "p" "a" "l" "snd ic"] is_val l n_def by auto + + (* Start *) + have "r l = (\t = 0..q. b^t * R ic p l t)" using r_def RLe_def by auto + also have "... = R ic p l 0 + (\t = 1..q. b^t * R ic p l t)" + by (auto simp: q comm_monoid_add_class.sum.atLeast_Suc_atMost) + also have "... = (\t \ {1..q}. b^t * R ic p l t)" + by (simp add: R_def 0) + also have "... = (\t \ (Suc ` {0..(q-1)}). b^t * R ic p l t)" using q by auto + also have "... = (sum ((\t. b^t * R ic p l t) \ Suc)) {0..(q-1)}" + using comm_monoid_add_class.sum.reindex[of "Suc" "{0..(q-1)}" "(\t. b^t * R ic p l t)"] by auto + also have "... = (\t = 0..(q-1). b^(Suc t) *(R ic p l t + + (\R+ p l (\k. S ic p k t)) + - (\R- p l (\k. (Z ic p l t) * S ic p k t))))" + using lm04_06_one_step_relation_register[of "ic" "p" "a" "l"] is_val l + by (simp add: n_def m_def) + + also have "... = (\t \ {0..(q-1)}. b^(Suc t) * R ic p l t + + b^(Suc t) * (\R+ p l (\k. S ic p k t)) + - b^(Suc t) * (\R- p l (\k. (Z ic p l t) * S ic p k t)))" + by (auto simp add: algebra_simps) + + finally have "int (r l) = (\t \ {0..(q-1)}. int( + b^(Suc t) * R ic p l t + + b^(Suc t) * (\R+ p l (\k. S ic p k t)) + - b^(Suc t) * (\R- p l (\k. (Z ic p l t) * S ic p k t))))" + by auto + + also have "... = (\t \ {0..(q-1)}. int (b^(Suc t) * R ic p l t) + + int (b^(Suc t) * (\R+ p l (\k. S ic p k t))) + - int (b^(Suc t) * (\R- p l (\k. (Z ic p l t) * S ic p k t))))" + by (simp only: sum_int positive) + + also have "... = (\t \ {0..(q-1)}. int (b^(Suc t) * R ic p l t) + + (int (b^(Suc t) * (\R+ p l (\k. S ic p k t))) + - int (b^(Suc t) * (\R- p l (\k. (Z ic p l t) * S ic p k t)))))" + by (simp add: add_diff_eq) + + also have "... = (\t \ {0..(q-1)}. int( b^(Suc t) * R ic p l t )) + + (\t \ {0..(q-1)}. int( b^(Suc t) * (\R+ p l (\k. S ic p k t))) + - int( b^(Suc t) * (\R- p l (\k. (Z ic p l t) * S ic p k t ))))" + by (auto simp only: sum.distrib) + + also have "... = int b * int (\t \ {0..(q-1)}. b^t * R ic p l t ) + + int b * (\t \ {0..(q-1)}. int(b^t * (\R+ p l (\k. S ic p k t))) + - int(b^t * (\R- p l (\k. (Z ic p l t) * S ic p k t ))))" + by (auto simp: sum_distrib_left mult.assoc right_diff_distrib) + + also have "... = int b * int (\t \ {0..(q-1)}. b^t * R ic p l t ) + + int b * (\t \ {0..(q-1)}. int(b^t * (\R+ p l (\k. S ic p k t)))) + - int b * (\t \ {0..(q-1)}. int(b^t * (\R- p l (\k. (Z ic p l t) * S ic p k t))))" + by (auto simp add: sum.distrib int_distrib(4) sum_subtractf) + + also have "... = int b * int (\t \ {0..(q-1)}. b^t * R ic p l t ) + + int b * (\t \ {0..(q-1)}. int(\R+ p l (\k. b^t * S ic p k t))) + - int b * (\t \ {0..(q-1)}. int(\R- p l (\k. b^t * (Z ic p l t * S ic p k t))))" + using sum_radd_distrib sum_rsub_distrib by auto + + also have "... = int b * int (\t = 0..q-1. b^t * R ic p l t) + + int b * int (\t = 0..q-1. \R+ p l (\k. b^t * S ic p k t)) + - int b * int (\t = 0..q-1. \R- p l (\k. b^t * (Z ic p l t * S ic p k t)))" + by auto + + also have "... = int b * int (\t = 0..q-1. b^t * R ic p l t) + + int b * int (\R+ p l (\k. \t=0..q-1. b^t * S ic p k t)) + - int b * int (\R- p l (\k. \t=0..q-1. b^t * (Z ic p l t * S ic p k t)))" + using sum_rsub_commutative[of "p" "l" "\k t. b^t * (Z ic p l t * S ic p k t)" "q-1"] + using sum_radd_commutative[of "p" "l" "\k t. b^t * S ic p k t " "q-1"] by auto + + (* extend all sums until q, because values don't change at time q *) + also have "... = int b * int (\t=0..q. b^t * R ic p l t) + + int b * int (\R+ p l (\k. \t=0..q-1. b^t * S ic p k t)) + - int b * int (\R- p l (\k. \t=0..q-1. b^t * (Z ic p l t * S ic p k t)))" + by (auto simp: rq R_def; smt One_nat_def rq) + + also have "... = int b * int (\t=0..q. b^t * R ic p l t) + + int b * int (\R+ p l (\k. \t=0..q. b^t * S ic p k t)) + - int b * int (\R- p l (\k. \t=0..q. b^t * (Z ic p l t * S ic p k t)))" + using zq add_q by auto + + also have "... = int b * int (RLe ic p b q l) + + int b * int (\R+ p l (SKe ic p b q)) + - int b * int (\R- p l (\k. \t=0..q. b^t * (Z ic p l t * S ic p k t)))" + by (auto simp: RLe_def; metis SKe_def) + + (* prove multiplication turns into bitAND *) + also have "... = int b * int (RLe ic p b q l) + + int b * int (\R+ p l (SKe ic p b q)) + - int b * int (\R- p l (\k. ZLe ic p b q l && SKe ic p b q k))" + using mult_to_bitAND c_gt_cells b_def c by auto + + finally have "int(r l) = int b * int (r l) + + int b * int (\R+ p l s) + - int b * int (\R- p l (\k. z l && s k))" + by (auto simp: r_def s_def z_def) + + hence "r l = b * r l + + b * \R+ p l s + - b * \R- p l (\k. z l && s k)" + using int_ops(5) int_ops(7) nat_int nat_minus_as_int by presburger + + thus ?thesis by simp +qed + +lemma lm04_23_multiple_register1: + fixes c :: nat + and l :: register + and ic :: configuration + and p :: program + and q :: nat + and a :: nat + + defines "b == B c" + and "m == length p" + and "n == length (snd ic)" + +assumes is_val: "is_valid_initial ic p a" +assumes c_gt_cells: "cells_bounded ic p c" +assumes l: "l = 0" +and q: "q > 0" + +assumes c: "c > 1" + + assumes terminate: "terminates ic p q" + + defines "r == RLe ic p b q" + and "z == ZLe ic p b q" + and "s == SKe ic p b q" + +shows "r l = a + b * r l + + b * (\R+ p l s) + - b * (\R- p l (\k. z l && s k))" +proof - + have n: "n > 0" using is_val + by (auto simp add: is_valid_initial_def n_def) + + have 0: "snd ic ! l = a" + using assms by (cases "ic"; auto simp add: is_valid_initial_def List.hd_conv_nth) + + find_theorems "hd ?l = ?l ! 0" + + have bound_fst_ic: "(if fst ic \ length p-1 then 1 else 0) \ Suc 0" by auto + have "(if issub (p ! k) \ l = modifies (p ! k) then if fst ic = k then 1 else 0 else 0) + = (if k = fst ic \ issub (p ! k) \ l = modifies (p ! k) then 1 else 0)" for k by auto + hence "(if issub (p ! k) \ l = modifies (p ! k) then if fst ic = k then Suc 0 else 0 else 0) + \ (if k = fst ic then 1 else 0)" for k + apply (cases "p!k") + apply (cases "modifies (p!k)") + by auto + hence sub: " (\k = 0..length p-1. if issub (p ! k) \ l = modifies (p ! k) + then if fst ic = k then Suc 0 else 0 else 0) \ Suc 0" + using Groups_Big.ordered_comm_monoid_add_class.sum_mono[of "{0..length p-1}" + "\k. (if issub (p ! k) \ l = modifies (p ! k) then if fst ic = k then Suc 0 else 0 else 0)" + "\k. (if k = fst ic then 1 else 0)"] bound_fst_ic Orderings.ord_class.ord_eq_le_trans by auto + + have "b^(Suc t) * (\R- p l (\k. (Z ic p l t) * S ic p k t)) \ b^(Suc t) * R ic p l t " for t + proof (cases "t=0") + case True + hence "a = R ic p l 0" by (auto simp add: 0 R_def) + thus ?thesis + apply (cases "a=0") + subgoal by (auto simp add: True R_def Z_def sum_rsub.simps) + subgoal + apply (auto simp add: True R_def Z_def sum_rsub.simps S_def) + using sub by auto + done + next + case False + define cs where "cs \ fst (steps ic p t)" + have sub: "(\R- p l (\k. Z ic p l t * S ic p k t)) + = (if issub (p!cs) \ l = modifies (p!cs) then Z ic p l t else 0)" + using single_step_sub Z_def R_def is_val l n_def cs_def n by auto + show ?thesis using sub by (auto simp add: sum_rsub.simps R_def Z_def) + qed + + from this have positive: "b^(Suc t) * (\R- p l (\k. (Z ic p l t) * S ic p k t)) + \ b^(Suc t) * R ic p l t + + b^(Suc t) * (\R+ p l (\k. S ic p k t))" for t + by (auto simp add: Nat.trans_le_add1) + + have distrib_add: "\t. b^t * \R+ p l (\k. S ic p k t) = \R+ p l (\k. b ^ t * S ic p k t)" + by (simp add: sum_radd_distrib) + have distrib_sub: "\t. b^t * \R- p l (\k. Z ic p l t * S ic p k t) + = \R- p l (\k. b ^ t * (Z ic p l t * S ic p k t))" + by (simp add: sum_rsub_distrib) + + have commute_add: "(\t=0..q-1. \R+ p l (\k. b^t * S ic p k t)) + = \R+ p l (\k. \t=0..q-1. (b^t * S ic p k t))" + using sum_radd_commutative[of "p" "l" "\k t. b^t * S ic p k t " "q-1"] by auto + + have "length (snd ic) > 0" using is_val is_valid_initial_def[of "ic" "p" "a"] by auto + hence r_q: "R ic p l q = 0" + using terminate terminates_def correct_halt_def l by (auto simp: n_def R_def) + hence z_q: "Z ic p l q = 0" + using terminate by (auto simp: Z_def) + + have "\k ishalt (p!k)" + using is_val is_valid_initial_def[of "ic" "p" "a"] is_valid_def[of "ic" "p"] + program_includes_halt.simps by blast + hence s_q: "\k < length p - 1. S ic p k q = 0" + using terminate terminates_def correct_halt_def by (auto simp: S_def) + + from r_q have rq: "(\x = 0..q - 1. int b ^ x * int (snd (steps ic p x) ! l)) = + (\x = 0..q. int b ^ x * int (snd (steps ic p x) ! l))" + by (auto simp: r_q R_def; smt Suc_pred mult_0_right of_nat_0 of_nat_mult power_mult_distrib q + sum.atLeast0_atMost_Suc zero_power) + + have "(\t = 0..q - 1. b ^ t * (Z ic p l t * S ic p k t)) + + (b^(Suc (q-1)) * (Z ic p l (Suc (q-1)) * S ic p k (Suc (q-1)))) + = (\t = 0..Suc (q-1). b ^ t * (Z ic p l t * S ic p k t)) " for k + using comm_monoid_add_class.sum.atLeast0_atMost_Suc by auto + hence zq: "(\t = 0..q - 1. b ^ t * (Z ic p l t * S ic p k t)) + = (\t = 0..q. b ^ t * (Z ic p l t * S ic p k t)) " for k + using z_q q by auto + + have "(if isadd (p ! k) \ l = modifies (p ! k) then \t = 0..q - Suc 0. b ^ t * S ic p k t else 0) + = (if isadd (p ! k) \ l = modifies (p ! k) then \t = 0..q. b ^ t * S ic p k t else 0)" for k + proof (cases "p!k") + case (Add x11 x12) + have sep: "(\t = 0..q-1. b ^ t * S ic p k t) + b^q * S ic p k q + = (\t = 0..(Suc (q-1)). b ^ t * S ic p k t)" + using comm_monoid_add_class.sum.atLeast0_atMost_Suc[of "\t. b^t * S ic p k t" "q-1"] q + by auto + have "ishalt (p ! (fst (steps ic p q)))" + using terminates_halt_state[of "ic" "p"] is_val terminate by auto + hence "S ic p k q = 0" using Add S_def[of "ic" "p" "k" "q"] by auto + with sep q have "(\t = 0..q - Suc 0. b ^ t * S ic p k t) = (\t = 0..q. b ^ t * S ic p k t)" + by auto + thus ?thesis by auto + next + case (Sub x21 x22 x23) + then show ?thesis by auto + next + case Halt + then show ?thesis by auto + qed + + hence add_q: "\R+ p l (\k. \t=0..(q-1). b^t * S ic p k t) + = \R+ p l (\k. \t=0..q. b^t * S ic p k t)" + using sum_radd.simps single_step_add[of "ic" "p" "a" "l" "snd ic"] is_val l n_def by auto + + (* Start *) + have "r l = (\t = 0..q. b^t * R ic p l t)" using r_def RLe_def by auto + also have "... = R ic p l 0 + (\t = 1..q. b^t * R ic p l t)" + by (auto simp: q comm_monoid_add_class.sum.atLeast_Suc_atMost) + also have "... = a + (\t \ {1..q}. b^t * R ic p l t)" + by (simp add: R_def 0) + also have "... = a + (\t = (Suc 0)..(Suc (q-1)). b^t * R ic p l t)" using q by auto + also have "... = a + (\t \ (Suc ` {0..(q-1)}). b^t * R ic p l t)" by auto + also have "... = a + (sum ((\t. b^t * R ic p l t) \ Suc)) {0..(q-1)}" + using comm_monoid_add_class.sum.reindex[of "Suc" "{0..(q-1)}" "(\t. b^t * R ic p l t)"] by auto + also have "... = a + (\t = 0..(q-1). b^(Suc t) * R ic p l (Suc t))" by auto + also have "... = a + (\t = 0..(q-1). b^(Suc t) *(R ic p l t + + (\R+ p l (\k. S ic p k t)) + - (\R- p l (\k. (Z ic p l t) * S ic p k t))))" + using lm04_06_one_step_relation_register[of "ic" "p" "a" "l"] is_val n n_def l + by (auto simp add: n_def m_def) + + also have "... = a + (\t \ {0..(q-1)}. b^(Suc t) * R ic p l t + + b^(Suc t) * (\R+ p l (\k. S ic p k t)) + - b^(Suc t) * (\R- p l (\k. (Z ic p l t) * S ic p k t)))" + by (auto simp add: algebra_simps) + + finally have "int (r l) = int a +(\t \ {0..(q-1)}. int( + b^(Suc t) * R ic p l t + + b^(Suc t) * (\R+ p l (\k. S ic p k t)) + - b^(Suc t) * (\R- p l (\k. (Z ic p l t) * S ic p k t))))" + by auto + + also have "... = int a + (\t \ {0..(q-1)}. int (b^(Suc t) * R ic p l t) + + int (b^(Suc t) * (\R+ p l (\k. S ic p k t))) + - int (b^(Suc t) * (\R- p l (\k. (Z ic p l t) * S ic p k t))))" + by (simp only: sum_int positive) + + also have "... = int a + (\t \ {0..(q-1)}. int (b^(Suc t) * R ic p l t) + + (int (b^(Suc t) * (\R+ p l (\k. S ic p k t))) + - int (b^(Suc t) * (\R- p l (\k. (Z ic p l t) * S ic p k t)))))" + by (simp add: add_diff_eq) + + also have "... = int a + (\t \ {0..(q-1)}. int( b^(Suc t) * R ic p l t )) + + (\t \ {0..(q-1)}. int( b^(Suc t) * (\R+ p l (\k. S ic p k t))) + - int( b^(Suc t) * (\R- p l (\k. (Z ic p l t) * S ic p k t ))))" + by (auto simp only: sum.distrib) + + also have "... = int a + int b * int (\t \ {0..(q-1)}. b^t * R ic p l t ) + + int b * (\t \ {0..(q-1)}. int(b^t * (\R+ p l (\k. S ic p k t))) + - int(b^t * (\R- p l (\k. (Z ic p l t) * S ic p k t ))))" + by (auto simp: sum_distrib_left mult.assoc right_diff_distrib) + + also have "... = int a + int b * int (\t \ {0..(q-1)}. b^t * R ic p l t ) + + int b * (\t \ {0..(q-1)}. int(b^t * (\R+ p l (\k. S ic p k t)))) + - int b * (\t \ {0..(q-1)}. int(b^t * (\R- p l (\k. (Z ic p l t) * S ic p k t))))" + by (auto simp add: sum.distrib int_distrib(4) sum_subtractf) + + also have "... = int a + int b * int (\t \ {0..(q-1)}. b^t * R ic p l t ) + + int b * (\t \ {0..(q-1)}. int(\R+ p l (\k. b^t * S ic p k t))) + - int b * (\t \ {0..(q-1)}. int(\R- p l (\k. b^t * (Z ic p l t * S ic p k t))))" + using distrib_add distrib_sub by auto + + also have "... = int a + int b * int (\t = 0..q-1. b^t * R ic p l t) + + int b * int (\t = 0..q-1. \R+ p l (\k. b^t * S ic p k t)) + - int b * int (\t = 0..q-1. \R- p l (\k. b^t * (Z ic p l t * S ic p k t)))" + by auto + + also have "... = int a + int b * int (\t = 0..q-1. b^t * R ic p l t) + + int b * int (\R+ p l (\k. \t=0..q-1. b^t * S ic p k t)) + - int b * int (\R- p l (\k. \t=0..q-1. b^t * (Z ic p l t * S ic p k t)))" + using sum_rsub_commutative[of "p" "l" "\k t. b^t * (Z ic p l t * S ic p k t)" "q-1"] + using sum_radd_commutative[of "p" "l" "\k t. b^t * S ic p k t " "q-1"] by auto + + (* extend all sums until q, because values don't change at time q *) + also have "... = int a + int b * int (\t=0..q. b^t * R ic p l t) + + int b * int (\R+ p l (\k. \t=0..q-1. b^t * S ic p k t)) + - int b * int (\R- p l (\k. \t=0..q-1. b^t * (Z ic p l t * S ic p k t)))" + by (auto simp: rq R_def; smt One_nat_def rq) + + also have "... = int a + int b * int (\t=0..q. b^t * R ic p l t) + + int b * int (\R+ p l (\k. \t=0..q. b^t * S ic p k t)) + - int b * int (\R- p l (\k. \t=0..q. b^t * (Z ic p l t * S ic p k t)))" + using zq add_q by auto + + also have "... = int a + int b * int (RLe ic p b q l) + + int b * int (\R+ p l (SKe ic p b q)) + - int b * int (\R- p l (\k. \t=0..q. b^t * (Z ic p l t * S ic p k t)))" + by (auto simp: RLe_def; metis SKe_def) + + (* prove multiplication turns into bitAND *) + also have "... = int a + int b * int (RLe ic p b q l) + + int b * int (\R+ p l (SKe ic p b q)) + - int b * int (\R- p l (\k. ZLe ic p b q l && SKe ic p b q k))" + using mult_to_bitAND c_gt_cells b_def c by auto + + finally have "int(r l) = int a + int b * int (r l) + + int b * int (\R+ p l s) + - int b * int (\R- p l (\k. z l && s k))" + by (auto simp: r_def s_def z_def) + + hence "r l = a + b * r l + + b * \R+ p l s + - b * \R- p l (\k. z l && s k)" + using int_ops(5) int_ops(7) nat_int nat_minus_as_int by presburger + + thus ?thesis by simp +qed + +end diff --git a/thys/DPRM_Theorem/Register_Machine/MultipleStepState.thy b/thys/DPRM_Theorem/Register_Machine/MultipleStepState.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Register_Machine/MultipleStepState.thy @@ -0,0 +1,404 @@ +subsubsection \States\ + +theory MultipleStepState + imports SingleStepState +begin + +lemma lm04_24_multiple_step_states: +fixes c :: nat + and l :: register + and ic :: configuration + and p :: program + and q :: nat + and a :: nat + + defines "b == B c" + and "m == length p" + +assumes is_val: "is_valid_initial ic p a" +assumes c_gt_cells: "cells_bounded ic p c" +assumes d: "d \ m-1" and "0 < d" + and q: "q > 0" + +assumes terminate: "terminates ic p q" + +assumes c: "c > 1" + + defines "r \ RLe ic p b q" + and "z \ ZLe ic p b q" + and "s \ SKe ic p b q" + and "e \ \t = 0..q. b^t" + +shows "s d = b * (\S+ p d s) + + b * (\S- p d (\k. z (modifies (p!k)) && s k)) + + b * (\S0 p d (\k. (e - z (modifies (p!k))) && s k))" +proof - + have "program_includes_halt p" + using is_val is_valid_initial_def[of "ic" "p" "a"] is_valid_def[of "ic" "p"] by auto + + have halt_term0: "t \ q-1 \ (if ishalt (p!(fst (steps ic p t))) \ d = fst (steps ic p t) + then Suc 0 else 0) = 0" for t + using terminate terminates_def by auto + + have single_step: "S ic p d (Suc t) = (\S+ p d (\k. S ic p k t)) + + (\S- p d (\k. Z ic p (modifies (p!k)) t * S ic p k t)) + + (\S0 p d (\k. (1 - Z ic p (modifies (p!k)) t) * S ic p k t)) + + (if ishalt (p!(fst (steps ic p t))) \ d = fst (steps ic p t) then Suc 0 else 0)" for t + using lm04_07_one_step_relation_state[of "ic" "p" "a" "d"] is_val `d>0` d + by (simp add: m_def) + + have b: "b > 0" using b_def B_def by auto + have halt: "ishalt (p!fst(steps ic p q))" using terminate terminates_def correct_halt_def by auto + have add_conditions: "(if isadd (p ! k) \ d = goes_to (p ! k) + then (\t = 0..q - Suc 0. b ^ t * S ic p k t) + b ^ q * S ic p k q else 0) + = (if isadd (p ! k) \ d = goes_to (p ! k) + then \t = 0..q - Suc 0. b ^ t * S ic p k t else 0)" for k + apply (cases "p!k"; cases "d = goes_to (p!k)") using q S_def b halt by auto + have "b * b ^ (q - Suc 0) = b ^ (q - Suc 0 + Suc 0)" using q + by (simp add: power_eq_if) + have "(\k. (\t = 0..(q-1). b^t * S ic p k t) + b^(Suc (q-1)) * S ic p k (Suc (q-1))) + = (\k. (\t = 0..(Suc (q-1)). b^t * S ic p k t))" by auto + hence "\S+ p d (\k. (\t = 0..(q-1). b^t * S ic p k t) + b^q * S ic p k (Suc (q-1))) + = \S+ p d (\k. \t = 0..(Suc (q-1)). b^t * S ic p k t)" using q + by auto + hence add_q: "\S+ p d (\k. \t = 0..(q-1). b^t * S ic p k t) + = \S+ p d (\k. \t = 0..q. b^t * S ic p k t)" + by (auto simp add: sum_sadd.simps q add_conditions) + + have "issub (p!k) \ b ^ (Suc (q-1)) * (Z ic p (modifies (p ! k)) (Suc (q-1)) * + (if fst (steps ic p (Suc (q-1))) = k then Suc 0 else 0)) = 0" for k + by (auto simp: q halt) + hence sum_equiv_nzero: "issub (p!k) \ + (\t = 0..q-1. b ^ t * (Z ic p (modifies (p ! k)) t * + (if fst (steps ic p t) = k then Suc 0 else 0))) + = (\t = 0..(Suc (q-1)). b ^ t * (Z ic p (modifies (p ! k)) t * + (if fst (steps ic p t) = k then Suc 0 else 0)))" for k + using sum.atLeast0_atMost_Suc[of "\t. b ^ t * (Z ic p (modifies (p ! k)) t + * (if fst (steps ic p t) = k then Suc 0 else 0))" "q-1"] by auto + hence sub_nzero_conditions: "(if issub (p ! k) \ d = goes_to (p ! k) then + \t = 0..q - Suc 0. b ^ t * (Z ic p (modifies (p ! k)) t * S ic p k t) else 0) + = (if issub (p ! k) \ d = goes_to (p ! k) then + \t = 0..q. b ^ t * (Z ic p (modifies (p ! k)) t * S ic p k t) else 0)" for k + apply (cases "issub (p!k)") using q S_def halt b by auto + have "(\k. (\t=0..(q-1). b^t * (Z ic p (modifies (p!k)) t * S ic p k t)) + + b^(Suc (q-1)) * (Z ic p (modifies (p!k)) (Suc (q-1)) * S ic p k (Suc (q-1)))) + = (\k. \t=0..(Suc (q-1)). b^t * (Z ic p (modifies (p!k)) t * S ic p k t))" by auto + hence sub_nzero_q: "(\S- p d (\k. \t=0..(q-1). b^t * (Z ic p (modifies (p!k)) t * S ic p k t))) + = (\S- p d (\k. \t=0..q. b^t * (Z ic p (modifies (p!k)) t * S ic p k t)))" + by (auto simp: sum_ssub_nzero.simps q sub_nzero_conditions) + + have "issub (p!k) \ b ^ (Suc (q-1)) * ((Suc 0 - Z ic p (modifies (p ! k)) (Suc (q-1))) + * S ic p k (Suc (q-1))) = 0" for k using q halt S_def by auto + hence sum_equiv_zero: "issub (p!k) \ + (\t = 0..q-1. b ^ t * ((Suc 0 - Z ic p (modifies (p ! k)) t) * S ic p k t)) + = (\t = 0..Suc (q-1). b ^ t * ((Suc 0 - Z ic p (modifies (p ! k)) t) * S ic p k t))" for k + using sum.atLeast0_atMost_Suc[of "\t. b ^ t * ((Suc 0 - Z ic p (modifies (p ! k)) t) + * S ic p k t)" "q-1"] by auto + have "(if issub (p ! k) \ d = goes_to_alt (p ! k) then + \t = 0..q - Suc 0. b ^ t * ((Suc 0 - Z ic p (modifies (p ! k)) t) * S ic p k t) else 0) + = (if issub (p ! k) \ d = goes_to_alt (p ! k) then + \t = 0..q. b ^ t * ((Suc 0 - Z ic p (modifies (p ! k)) t) * S ic p k t) else 0)" for k + apply (cases "issub (p!k)") using sum_equiv_zero[of "k"] q by auto + hence sub_zero_q: "(\S0 p d (\k.\t=0..q-1. b^t * ((1 - Z ic p (modifies(p!k)) t) * S ic p k t))) + = (\S0 p d (\k.\t=0..q. b^t * ((1 - Z ic p (modifies (p!k)) t) * S ic p k t)))" + using sum_ssub_zero.simps q by auto + + (* Start *) + have "s d = (\t = 0..q. b^t * S ic p d t)" using s_def SKe_def by auto + also have "... = S ic p d 0 + (\t = 1..q. b^t * S ic p d t)" + by (auto simp: q comm_monoid_add_class.sum.atLeast_Suc_atMost) + also have "... = (\t = 1..q. b^t * S ic p d t)" + using S_def `d>0` is_val is_valid_initial_def[of "ic" "p" "a"] by auto + also have "... = (\t \ (Suc ` {0..(q-1)}). b^t * S ic p d t)" using q by auto + also have "... = (sum ((\t. b^t * S ic p d t) \ Suc)) {0..(q-1)}" + using comm_monoid_add_class.sum.reindex[of "Suc" "{0..(q-1)}" "(\t. b^t * S ic p d t)"] by auto + also have "... = (\t = 0..(q-1). b^(Suc t) *((\S+ p d (\k. S ic p k t)) + + (\S- p d (\k. Z ic p (modifies (p!k)) t * S ic p k t)) + + (\S0 p d (\k. (1 - Z ic p (modifies (p!k)) t) * S ic p k t)) + + (if ishalt (p!(fst (steps ic p t))) \ d = fst (steps ic p t) + then Suc 0 else 0)))" + using single_step by auto + also have "... = (\t = 0..(q-1). b^(Suc t) *((\S+ p d (\k. S ic p k t)) + + (\S- p d (\k. Z ic p (modifies (p!k)) t * S ic p k t)) + + (\S0 p d (\k. (1 - Z ic p (modifies (p!k)) t) * S ic p k t))))" + using halt_term0 by auto + also have "... = (\t = 0..(q-1). (b^(Suc t) * (\S+ p d (\k. S ic p k t)) + + b^(Suc t) * (\S- p d (\k. Z ic p (modifies (p!k)) t * S ic p k t)) + + b^(Suc t) * (\S0 p d (\k. (1 - Z ic p (modifies (p!k)) t) * S ic p k t))))" + by (simp add: algebra_simps) + also have "... = (\t = 0..(q-1). (b^(Suc t) * (\S+ p d (\k. S ic p k t)))) + +(\t=0..(q-1). b^(Suc t)*(\S- p d (\k. Z ic p (modifies (p!k)) t * S ic p k t))) + +(\t=0..(q-1). b^(Suc t)*(\S0 p d (\k. (1 - Z ic p (modifies (p!k)) t) * S ic p k t)))" + by (auto simp only: sum.distrib) + also have "... = b * (\t = 0..(q-1). (b^t * (\S+ p d (\k. S ic p k t)))) + + b*(\t=0..(q-1). b^t*(\S- p d (\k. Z ic p (modifies (p!k)) t * S ic p k t))) + + b*(\t=0..(q-1). b^t*(\S0 p d (\k. (1 - Z ic p (modifies (p!k)) t) * S ic p k t)))" + by (auto simp: algebra_simps sum_distrib_left) + also have "... = b * (\t = 0..(q-1). (\S+ p d (\k. b^t * S ic p k t))) + + b*(\t=0..(q-1). (\S- p d (\k. b^t * (Z ic p (modifies (p!k)) t * S ic p k t)))) + + b*(\t=0..(q-1). (\S0 p d (\k. b^t * ((1 - Z ic p (modifies (p!k)) t) * S ic p k t))))" + using sum_sadd_distrib sum_ssub_nzero_distrib sum_ssub_zero_distrib by auto + also have "... = b * (\S+ p d (\k. \t = 0..(q-1). b^t * S ic p k t)) + + b*(\S- p d (\k. \t=0..(q-1). b^t * (Z ic p (modifies (p!k)) t * S ic p k t))) + + b*(\S0 p d (\k. \t=0..(q-1). b^t * ((1 - Z ic p (modifies (p!k)) t) * S ic p k t)))" + using sum_sadd_commutative sum_ssub_nzero_commutative sum_ssub_zero_commutative by auto + (* extend all sums until q *) + finally have eq1: "s d = b * (\S+ p d (\k. \t = 0..q. b^t * S ic p k t)) + + b*(\S- p d (\k. \t=0..q. b^t * (Z ic p (modifies (p!k)) t * S ic p k t))) + + b*(\S0 p d (\k. \t=0..q. b^t * ((1 - Z ic p (modifies (p!k)) t) * S ic p k t)))" + using add_q sub_nzero_q sub_zero_q by auto + also have "... = b * (\S+ p d (\k. s k)) + + b*(\S- p d (\k. \t=0..q. b^t * (Z ic p (modifies (p!k)) t * S ic p k t))) + + b*(\S0 p d (\k. \t=0..q. b^t * ((1 - Z ic p (modifies (p!k)) t) * S ic p k t)))" + using SKe_def s_def by auto + finally have "s d = b * (\S+ p d s) + + b*(\S- p d (\k. \t=0..q. b^t * (Z ic p (modifies (p!k)) t * S ic p k t))) + + b*(\S0 p d (\k. \t=0..q. b^t * ((1 - Z ic p (modifies (p!k)) t) * S ic p k t)))" + by auto + also have "... = b * (\S+ p d s) + + b*(\S- p d (\k. ZLe ic p b q (modifies (p!k)) && SKe ic p b q k)) + + b*(\S0 p d (\k. \t=0..q. b^t * ((1 - Z ic p (modifies (p!k)) t) * S ic p k t)))" + using mult_to_bitAND c_gt_cells b_def c by auto + finally have "s d = b * (\S+ p d s) + + b*(\S- p d (\k. ZLe ic p b q (modifies (p!k)) && SKe ic p b q k)) + + b*(\S0 p d (\k. (e - ZLe ic p b q (modifies (p!k))) && SKe ic p b q k))" + using mult_to_bitAND_state c_gt_cells b_def c e_def by auto + thus ?thesis using s_def z_def by auto +qed + +lemma lm04_25_multiple_step_state1: +fixes c :: nat + and l :: register + and ic :: configuration + and p :: program + and q :: nat + and a :: nat + + defines "b == B c" + and "m == length p" + +assumes is_val: "is_valid_initial ic p a" +assumes c_gt_cells: "cells_bounded ic p c" +assumes d: "d=0" + and q: "q > 0" + +assumes terminate: "terminates ic p q" + +assumes c: "c > 1" + + defines "r \ RLe ic p b q" + and "z \ ZLe ic p b q" + and "s \ SKe ic p b q" + and "e \ \t = 0..q. b^t" + +shows "s d = 1 + b * (\S+ p d s) + + b * (\S- p d (\k. z (modifies (p!k)) && s k)) + + b * (\S0 p d (\k. (e - z (modifies (p!k))) && s k))" +proof - + have "program_includes_halt p" + using is_val is_valid_initial_def[of "ic" "p" "a"] is_valid_def[of "ic" "p"] by auto + hence "p \ []" by auto + have "\ ishalt (p!d)" using d m_def `program_includes_halt p` by auto + hence "(if ishalt (p ! fst (steps ic p t)) \ d = fst (steps ic p t) then Suc 0 else 0) = 0" for t + by auto + hence single_step: "\t. S ic p d (Suc t) = (\S+ p d (\k. S ic p k t)) + + (\S- p d (\k. Z ic p (modifies (p!k)) t * S ic p k t)) + + (\S0 p d (\k. (1 - Z ic p (modifies (p!k)) t) * S ic p k t))" + using lm04_07_one_step_relation_state[of "ic" "p" "a" "d"] is_val d `p \ []` by (simp add: m_def) + + have b: "b > 0" using b_def B_def by auto + have halt: "ishalt (p!fst(steps ic p q))" using terminate terminates_def correct_halt_def by auto + have add_conditions: "(if isadd (p ! k) \ d = goes_to (p ! k) + then (\t = 0..q - Suc 0. b ^ t * S ic p k t) + b ^ q * S ic p k q else 0) + = (if isadd (p ! k) \ d = goes_to (p ! k) + then \t = 0..q - Suc 0. b ^ t * S ic p k t else 0)" for k + apply (cases "p!k"; cases "d = goes_to (p!k)") using q S_def b halt by auto + have "b * b ^ (q - Suc 0) = b ^ (q - Suc 0 + Suc 0)" using q + by (simp add: power_eq_if) + have "(\k. (\t = 0..(q-1). b^t * S ic p k t) + b^(Suc (q-1)) * S ic p k (Suc (q-1))) + = (\k. (\t = 0..(Suc (q-1)). b^t * S ic p k t))" by auto + hence "\S+ p d (\k. (\t = 0..(q-1). b^t * S ic p k t) + b^q * S ic p k (Suc (q-1))) + = \S+ p d (\k. \t = 0..(Suc (q-1)). b^t * S ic p k t)" using q + by auto + hence add_q: "\S+ p d (\k. \t = 0..(q-1). b^t * S ic p k t) + = \S+ p d (\k. \t = 0..q. b^t * S ic p k t)" + by (auto simp add: sum_sadd.simps q add_conditions) + + have "issub (p!k) \ b ^ (Suc (q-1)) * (Z ic p (modifies (p ! k)) (Suc (q-1)) * + (if fst (steps ic p (Suc (q-1))) = k then Suc 0 else 0)) = 0" for k + by (auto simp: q halt) + hence sum_equiv_nzero: "issub (p!k) \ + (\t = 0..q-1. b ^ t * (Z ic p (modifies (p ! k)) t * + (if fst (steps ic p t) = k then Suc 0 else 0))) + = (\t = 0..(Suc (q-1)). b ^ t * (Z ic p (modifies (p ! k)) t * + (if fst (steps ic p t) = k then Suc 0 else 0)))" for k + using sum.atLeast0_atMost_Suc[of "\t. b ^ t * (Z ic p (modifies (p ! k)) t + * (if fst (steps ic p t) = k then Suc 0 else 0))" "q-1"] by auto + hence sub_nzero_conditions: "(if issub (p ! k) \ d = goes_to (p ! k) then + \t = 0..q - Suc 0. b ^ t * (Z ic p (modifies (p ! k)) t * S ic p k t) else 0) + = (if issub (p ! k) \ d = goes_to (p ! k) then + \t = 0..q. b ^ t * (Z ic p (modifies (p ! k)) t * S ic p k t) else 0)" for k + apply (cases "issub (p!k)") using q S_def halt b by auto + have "(\k. (\t=0..(q-1). b^t * (Z ic p (modifies (p!k)) t * S ic p k t)) + + b^(Suc (q-1)) * (Z ic p (modifies (p!k)) (Suc (q-1)) * S ic p k (Suc (q-1)))) + = (\k. \t=0..(Suc (q-1)). b^t * (Z ic p (modifies (p!k)) t * S ic p k t))" by auto + hence sub_nzero_q: "(\S- p d (\k. \t=0..(q-1). b^t * (Z ic p (modifies (p!k)) t * S ic p k t))) + = (\S- p d (\k. \t=0..q. b^t * (Z ic p (modifies (p!k)) t * S ic p k t)))" + by (auto simp: sum_ssub_nzero.simps q sub_nzero_conditions) + + have "issub (p!k) \ b ^ (Suc (q-1)) * ((Suc 0 - Z ic p (modifies (p ! k)) (Suc (q-1))) + * S ic p k (Suc (q-1))) = 0" for k using q halt S_def by auto + hence sum_equiv_zero: "issub (p!k) \ + (\t = 0..q-1. b ^ t * ((Suc 0 - Z ic p (modifies (p ! k)) t) * S ic p k t)) + = (\t = 0..Suc (q-1). b ^ t * ((Suc 0 - Z ic p (modifies (p ! k)) t) * S ic p k t))" for k + using sum.atLeast0_atMost_Suc[of "\t. b ^ t * ((Suc 0 - Z ic p (modifies (p ! k)) t) + * S ic p k t)" "q-1"] by auto + have "(if issub (p ! k) \ d = goes_to_alt (p ! k) then + \t = 0..q - Suc 0. b ^ t * ((Suc 0 - Z ic p (modifies (p ! k)) t) * S ic p k t) else 0) + = (if issub (p ! k) \ d = goes_to_alt (p ! k) then + \t = 0..q. b ^ t * ((Suc 0 - Z ic p (modifies (p ! k)) t) * S ic p k t) else 0)" for k + apply (cases "issub (p!k)") using sum_equiv_zero[of "k"] q by auto + hence sub_zero_q: "(\S0 p d (\k.\t=0..q-1. b^t * ((1 - Z ic p (modifies(p!k)) t) * S ic p k t))) + = (\S0 p d (\k.\t=0..q. b^t * ((1 - Z ic p (modifies (p!k)) t) * S ic p k t)))" + using sum_ssub_zero.simps q by auto + + have S0: "S ic p d 0 = 1" using S_def is_val is_valid_initial_def[of "ic" "p" "a"] d by auto + + (* Start *) + have "s d = (\t = 0..q. b^t * S ic p d t)" using s_def SKe_def by auto + also have "... = S ic p d 0 + (\t = 1..q. b^t * S ic p d t)" + by (auto simp: q comm_monoid_add_class.sum.atLeast_Suc_atMost) + also have "... = b^0 * S ic p d 0 + (\t = 1..q. b^t * S ic p d t)" + using S_def is_val is_valid_initial_def[of "ic" "p" "a"] by auto + also have "... = 1 + (\t \ (Suc ` {0..(q-1)}). b^t * S ic p d t)" using q S0 by auto + also have "... = 1 + (sum ((\t. b^t * S ic p d t) \ Suc)) {0..(q-1)}" + using comm_monoid_add_class.sum.reindex[of "Suc" "{0..(q-1)}" "(\t. b^t * S ic p d t)"] by auto + also have "... = 1 + (\t = 0..(q-1). b^(Suc t) *((\S+ p d (\k. S ic p k t)) + + (\S- p d (\k. Z ic p (modifies (p!k)) t * S ic p k t)) + + (\S0 p d (\k. (1 - Z ic p (modifies (p!k)) t) * S ic p k t))))" + using single_step by auto + also have "... = 1 + (\t = 0..(q-1). (b^(Suc t) * (\S+ p d (\k. S ic p k t)) + + b^(Suc t) * (\S- p d (\k. Z ic p (modifies (p!k)) t * S ic p k t)) + + b^(Suc t) * (\S0 p d (\k. (1 - Z ic p (modifies (p!k)) t) * S ic p k t))))" + by (simp add: algebra_simps) + also have "... = 1 + (\t = 0..(q-1). (b^(Suc t) * (\S+ p d (\k. S ic p k t)))) + +(\t=0..(q-1). b^(Suc t)*(\S- p d (\k. Z ic p (modifies (p!k)) t * S ic p k t))) + +(\t=0..(q-1). b^(Suc t)*(\S0 p d (\k. (1 - Z ic p (modifies (p!k)) t) * S ic p k t)))" + by (auto simp only: sum.distrib) + also have "... = 1 + b * (\t = 0..(q-1). (b^t * (\S+ p d (\k. S ic p k t)))) + + b*(\t=0..(q-1). b^t*(\S- p d (\k. Z ic p (modifies (p!k)) t * S ic p k t))) + + b*(\t=0..(q-1). b^t*(\S0 p d (\k. (1 - Z ic p (modifies (p!k)) t) * S ic p k t)))" + by (auto simp: algebra_simps sum_distrib_left) + also have "... = 1 + b * (\t = 0..(q-1). (\S+ p d (\k. b^t * S ic p k t))) + + b*(\t=0..(q-1). (\S- p d (\k. b^t * (Z ic p (modifies (p!k)) t * S ic p k t)))) + + b*(\t=0..(q-1). (\S0 p d (\k. b^t * ((1 - Z ic p (modifies (p!k)) t) * S ic p k t))))" + using sum_sadd_distrib sum_ssub_nzero_distrib sum_ssub_zero_distrib by auto + also have "... = 1 + b * (\S+ p d (\k. \t = 0..(q-1). b^t * S ic p k t)) + + b*(\S- p d (\k. \t=0..(q-1). b^t * (Z ic p (modifies (p!k)) t * S ic p k t))) + + b*(\S0 p d (\k. \t=0..(q-1). b^t * ((1 - Z ic p (modifies (p!k)) t) * S ic p k t)))" + using sum_sadd_commutative sum_ssub_nzero_commutative sum_ssub_zero_commutative by auto + (* extend all sums until q *) + finally have eq1: "s d = 1 + b * (\S+ p d (\k. \t = 0..q. b^t * S ic p k t)) + + b*(\S- p d (\k. \t=0..q. b^t * (Z ic p (modifies (p!k)) t * S ic p k t))) + + b*(\S0 p d (\k. \t=0..q. b^t * ((1 - Z ic p (modifies (p!k)) t) * S ic p k t)))" + using add_q sub_nzero_q sub_zero_q by auto + also have "... = 1 + b * (\S+ p d (\k. s k)) + + b*(\S- p d (\k. \t=0..q. b^t * (Z ic p (modifies (p!k)) t * S ic p k t))) + + b*(\S0 p d (\k. \t=0..q. b^t * ((1 - Z ic p (modifies (p!k)) t) * S ic p k t)))" + using SKe_def s_def by auto + finally have "s d = 1 + b * (\S+ p d s) + + b*(\S- p d (\k. \t=0..q. b^t * (Z ic p (modifies (p!k)) t * S ic p k t))) + + b*(\S0 p d (\k. \t=0..q. b^t * ((1 - Z ic p (modifies (p!k)) t) * S ic p k t)))" + by auto + also have "... = 1 + b * (\S+ p d s) + + b*(\S- p d (\k. ZLe ic p b q (modifies (p!k)) && SKe ic p b q k)) + + b*(\S0 p d (\k. \t=0..q. b^t * ((1 - Z ic p (modifies (p!k)) t) * S ic p k t)))" + using mult_to_bitAND c_gt_cells b_def c by auto + finally have "s d = 1 + b * (\S+ p d s) + + b*(\S- p d (\k. ZLe ic p b q (modifies (p!k)) && SKe ic p b q k)) + + b*(\S0 p d (\k. (e - ZLe ic p b q (modifies (p!k))) && SKe ic p b q k))" + using mult_to_bitAND_state c_gt_cells b_def c e_def by auto + thus ?thesis using s_def z_def by auto +qed + +lemma halting_condition_04_27: +fixes c :: nat + and l :: register + and ic :: configuration + and p :: program + and q :: nat + and a :: nat + +defines "b == B c" + and "m == length p - 1" + +assumes is_val: "is_valid_initial ic p a" + and q: "q > 0" + +assumes terminate: "terminates ic p q" + +shows "SKe ic p b q m = b ^ q" +proof - + have halt: "ishalt (p ! (fst (steps ic p q)))" + using terminate terminates_def correct_halt_def by auto + have "\k ishalt (p!k)" using is_val is_valid_initial_def[of "ic" "p" "a"] + is_valid_def[of "ic" "p"] program_includes_halt.simps by blast + hence "ishalt (p!k) \ k \ length p - 1" for k using not_le_imp_less by auto + hence gt: "fst (steps ic p q) \ m" using halt m_def by auto + have "fst (steps ic p q) \ m" + using p_contains[of "ic" "p" "a" "q"] is_val m_def by auto + hence q_steps_m: "fst (steps ic p q) = m" using gt by auto + hence 1: "S ic p m q = 1" using S_def by auto + + have "ishalt (p!m)" using q_steps_m halt by auto + have "\t ishalt (p ! (fst (steps ic p t)))" using terminate terminates_def by auto + hence "\t (fst (steps ic p t) = m)" using `ishalt (p!m)` by auto + hence 0: "t \ q-1 \ S ic p m t = 0" for t using q S_def[of "ic" "p" "m" "t"] by auto + + have "SKe ic p b q m = (\t = 0..(Suc (q-1)). b ^ t * S ic p m t)" by (auto simp: q SKe_def) + also have "... = (\t = 0..(q-1). b^t * S ic p m t) + b ^ (Suc (q-1)) * S ic p m (Suc (q-1))" + by auto + also have "... = b ^ q" using 0 1 q by auto + finally show ?thesis by auto +qed + +lemma state_q_bound: +fixes c :: nat + and l :: register + and ic :: configuration + and p :: program + and q :: nat + and a :: nat + +defines "b == B c" + and "m == length p - 1" + +assumes is_val: "is_valid_initial ic p a" + and q: "q > 0" + and terminate: "terminates ic p q" + and c: "c > 0" + +assumes "k1" using B_def apply auto + by (metis One_nat_def one_less_numeral_iff power_gt1_lemma semiring_norm(76)) + hence b: "b > 2" using c b_def B_def + by (smt One_nat_def Suc_le_lessD less_Suc_eq_le less_trans_Suc linorder_neqE_nat + numeral_2_eq_2 power_Suc0_right power_inject_exp) + from \k have "\ ishalt (p!k)" using is_val + by (simp add: is_valid_def is_valid_initial_def is_val m_def) + hence "S ic p k q = 0" using terminate terminates_def correct_halt_def S_def by auto + hence "SKe ic p b q k = (\t = 0..q-1. b ^ t * S ic p k t)" + using \q>0\ apply (auto cong: sum.cong simp: SKe_def) by (metis (no_types, lifting) Suc_pred + add_cancel_right_right mult_0_right sum.atLeast0_atMost_Suc) + also have "... \ (\t = 0..q-1. b^t)" by (auto simp add: S_def gr_implies_not0 sum_mono) + also have "... < b ^ q" + using `q>0` sum_bt + by (metis Suc_diff_1 b) + + finally show ?thesis by auto +qed + +end diff --git a/thys/DPRM_Theorem/Register_Machine/MultipleToSingleSteps.thy b/thys/DPRM_Theorem/Register_Machine/MultipleToSingleSteps.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Register_Machine/MultipleToSingleSteps.thy @@ -0,0 +1,1148 @@ +subsection \From multiple to single step relations\ + +theory MultipleToSingleSteps + imports MachineEquations CommutationRelations "../Diophantine/Binary_And" +begin + +text \This file contains lemmas that are needed to prove the <-- direction of conclusion4.5 in the +file MachineEquationEquivalence. In particular, it is shown that single step equations follow +from the multiple step relations. The key idea of Matiyasevich's proof is to code all register +and state values over the time into one large number. A further central statement in this file +shows that the decoding of these numbers back to the single cell contents is indeed correct.\ + +context + fixes a :: nat + and ic:: configuration + and p :: program + and q :: nat + and r z :: "register \ nat" + and s :: "state \ nat" + and b c d e f :: nat + and m n :: nat + and Req Seq Zeq (* These variables are to store the single cell entries obtained from the + rows r s and z in the protocol *) + + assumes m_def: "m \ length p - 1" + and n_def: "n \ length (snd ic)" + + assumes is_val: "is_valid_initial ic p a" + + (* rm_equations *) + assumes m_eq: "mask_equations n r z c d e f" + and r_eq: "reg_equations p r z s b a n q" + and s_eq: "state_equations p s z b e q m" + and c_eq: "rm_constants q c b d e f a" + + assumes Seq_def: "Seq = (\k t. nth_digit (s k) t b)" + and Req_def: "Req = (\l t. nth_digit (r l) t b)" + and Zeq_def: "Zeq = (\l t. nth_digit (z l) t b)" + +begin + +text \Basic properties\ +lemma n_gt0: "n>0" + using n_def is_val is_valid_initial_def[of "ic" "p" "a"] is_valid_def + by auto + +lemma f_def: "f = (\t = 0..q. 2^c * b^t)" + using c_eq by (simp add: rm_constants_def F_def) +lemma e_def: "e = (\t = 0..q. b^t)" + using c_eq by (simp add: rm_constants_def E_def) +lemma d_def: "d = (\t = 0..q. (2^c - 1) * b^t)" + using c_eq by (simp add: D_def rm_constants_def) +lemma b_def: "b = 2^(Suc c)" + using c_eq by (simp add: B_def rm_constants_def) + +lemma b_gt1: "b > 1" using c_eq B_def rm_constants_def by auto + +lemma c_gt0: "c > 0" using rm_constants_def c_eq by auto +lemma h0: "1 < (2::nat)^c" + using c_gt0 one_less_numeral_iff one_less_power semiring_norm(76) by blast +(* this even implies b > 2, in fact b \ 4 *) + +lemma rl_fst_digit_zero: + assumes "l < n" + shows "nth_digit (r l) t b \ c = 0" +proof - + have "2 ^ c - (Suc 0) < 2 ^ Suc c" using c_gt0 by (simp add: less_imp_diff_less) + hence "\t. nth_digit d t b = (if t\q then 2^c - 1 else 0)" + using nth_digit_gen_power_series[of "\x. 2^c - 1" "c"] d_def c_gt0 b_def + by (simp add: d_def) + then have d_lead_digit_zero: "\t. (nth_digit d t b) \ c = 0" + by (auto simp: nth_bit_def) + + from m_eq have "(r l) \ d" + by (simp add: mask_equations_def assms) + then have "\k. (r l) \ k \ d \ k" + by (auto simp: masks_leq_equiv) + then have "\t. (nth_digit (r l) t b \ c) \ (nth_digit d t b \ c)" + using digit_gen_pow2_reduct[where ?c = "Suc c"] by (auto simp: b_def) + thus ?thesis + by (auto simp: d_lead_digit_zero) +qed + +lemma e_mask_bound: + assumes "x \ e" + shows "nth_digit x t b \ 1" +proof - + have x_bounded: "nth_digit x t' b \ nth_digit e t' b" for t' + proof - + have "\t'. x \ t' \ e \ t'" using assms masks_leq_equiv by auto + then have k_lt_c: "\t'. \k' k' + \ nth_digit e t' b \ k'" + using digit_gen_pow2_reduct by (auto simp: b_def) (metis power_Suc) + + have "k\Suc c \ x mod (2 ^ Suc c) div 2 ^ k = 0" for k x::nat + by (metis Euclidean_Division.div_eq_0_iff One_nat_def Suc_le_lessD b_gt1 b_def less_le_trans + mod_less_divisor nat_power_less_imp_less not_less not_less0 not_less_eq_eq + numeral_2_eq_2 zero_less_Suc) + then have "\k\Suc c. nth_digit x y b \ k = 0" for x y + using b_def nth_bit_def nth_digit_def by auto + then have k_gt_c: "\t'. \k'\Suc c. nth_digit x t' b \ k' + \ nth_digit e t' b \ k'" + by auto + + from k_lt_c k_gt_c have "nth_digit x t' b \ nth_digit e t' b" for t' + using bitwise_leq by (meson not_le) + thus ?thesis by auto + qed + + have "\k. Suc 0 < 2^c" using c_gt0 h0 by auto + hence e_aux: "nth_digit e tt b \ 1" for tt + using e_def b_def c_gt0 nth_digit_gen_power_series[of "\k. Suc 0" "c" "q"] by auto + + show ?thesis using e_aux[of "t"] x_bounded[of "t"] using le_trans by blast +qed + +(* --- individual bounds --- *) +lemma sk_bound: + shows "\t k. k\length p - 1 \ nth_digit (s k) t b \ 1" +proof - + have "\k\length p - 1. s k \ e" using s_eq state_equations_def m_def by auto + thus ?thesis using e_mask_bound by auto +qed + +lemma sk_bitAND_bound: + shows "\t k. k\length p - 1 \ nth_digit (s k && x k) t b \ 1" + using bitAND_nth_digit_commute sk_bound bitAND_lt masks_leq + by (auto simp: b_def) (meson dual_order.trans) + +lemma s_bound: + shows "\jM\m. (\k\M. s k) \ e" + using s_eq state_equations_def by auto + +lemma sk_sum_bound: + shows "\t M. M\length p - 1 \ nth_digit (\k\M. s k) t b \ 1" + using sk_sum_masked e_mask_bound m_def by auto + +lemma sum_sk_bound: + shows "(\k\length p - 1. nth_digit (s k) t b) \ 1" +proof - + have "\t m. m \ length p - 1 \ nth_digit (sum s {0..m}) t b < 2 ^ c" + using sk_sum_bound b_def c_gt0 h0 + by (metis atLeast0AtMost le_less_trans) + moreover have "\t k. k \ length p - 1 \ nth_digit (s k) t b < 2 ^ c" + using sk_bound b_def c_gt0 h0 + by (metis le_less_trans) + + ultimately have "nth_digit (\k\length p - 1. s k) t b + = (\k\length p - 1. nth_digit (s k) t b)" + using b_def c_gt0 + using finite_sum_nth_digit_commute2[where ?M = "length p - 1"] + by (simp add: atMost_atLeast0) + + thus ?thesis using sk_sum_bound by (metis order_refl) +qed + +lemma bitAND_sum_lt: "(\k\length p - 1. nth_digit (s k && x k) t b) + \ (\k\length p - 1. Seq k t)" +proof - + have "(\k\length p - 1. nth_digit (s k && x k) t b) + = (\k\length p - 1. nth_digit (s k) t b && nth_digit (x k) t b)" + using bitAND_nth_digit_commute b_def by auto + also have "... \ (\k\length p - 1. nth_digit (s k) t b)" + using bitAND_lt by (simp add: sum_mono) + finally show ?thesis using Seq_def by auto +qed + + +lemma states_unique_RAW: + "\k\m. Seq k t = 1 \ (\j\m. j \ k \ Seq j t = 0)" +proof - + { + fix k + assume "k\m" + assume skt_1: "Seq k t = 1" + + have "\j\m. j \ k \ Seq j t = 0" + proof - + { + fix j + assume "j\m" + assume "j \ k" + let ?fct = "(\k. Seq k t)" + + have "Seq j t = 0" + proof (rule ccontr) + assume "Seq j t \ 0" + then have "Seq j t + Seq k t > 1" + using skt_1 by auto + moreover have "sum ?fct {0..m} \ sum ?fct {j, k}" + using `j\m` `k\m` sum_mono2 + by (metis atLeastAtMost_iff empty_subsetI finite_atLeastAtMost insert_subset le0) + ultimately have "(\k\m. Seq k t) > 1" + by (simp add: \j \ k\ atLeast0AtMost) + thus "False" + using sum_sk_bound[where ?t = t] + by (auto simp: Seq_def m_def) + qed + } + thus ?thesis by auto + qed + } + thus ?thesis by auto +qed + +lemma block_sum_radd_bound: + shows "\t. (\R+ p l (\k. nth_digit (s k) t b)) \ 1" +proof - + { + fix t + have "(\R+ p l (\k. Seq k t)) \ (\k\length p - 1. Seq k t)" + unfolding sum_radd.simps + by (simp add: atMost_atLeast0 sum_mono) + hence "(\R+ p l (\k. Seq k t)) \ 1" + using sum_sk_bound[of t] Seq_def + using dual_order.trans by blast + } + thus ?thesis using Seq_def by auto +qed + +lemma block_sum_rsub_bound: + shows "\t. (\R- p l (\k. nth_digit (s k && z l) t b)) \ 1" +proof - + { + fix t + have "(\R- p l (\k. nth_digit (s k && z l) t b)) + \ (\k\length p - 1. nth_digit (s k && z l) t b)" + unfolding sum_rsub.simps + by (simp add: atMost_atLeast0 sum_mono) + also have "... \ (\k\length p - 1. Seq k t)" + using bitAND_sum_lt[where ?x = "\k. z l"] by blast + finally have "(\R- p l (\k. nth_digit (s k && z l) t b)) \ 1" + using sum_sk_bound[of t] Seq_def + using dual_order.trans by blast + } + thus ?thesis using Seq_def by auto +qed + +lemma block_sum_rsub_special_bound: + shows "\t. (\R- p l (\k. nth_digit (s k) t b)) \ 1" +proof - + { + fix t + have "(\R- p l (\k. nth_digit (s k) t b)) + \ (\k\length p - 1. nth_digit (s k) t b)" + unfolding sum_rsub.simps + by (simp add: atMost_atLeast0 sum_mono) + then have "(\R- p l (\k. nth_digit (s k) t b)) \ 1" + using sum_sk_bound[of t] + using dual_order.trans by blast + } + thus ?thesis using Seq_def by auto +qed + +lemma block_sum_sadd_bound: + shows "\t. (\S+ p j (\k. nth_digit (s k) t b)) \ 1" +proof - + { + fix t + have "(\S+ p j (\k. Seq k t)) \ (\k\length p - 1. Seq k t)" + unfolding sum_sadd.simps + by (simp add: atMost_atLeast0 sum_mono) + hence "(\S+ p j (\k. Seq k t)) \ 1" + using sum_sk_bound[of t] Seq_def + using dual_order.trans by blast + } + thus ?thesis using Seq_def by auto +qed + +lemma block_sum_ssub_bound: + shows "\t. (\S- p j (\k. nth_digit (s k && z (l k)) t b)) \ 1" +proof - + { + fix t + have "(\S- p j (\k. nth_digit (s k && z (l k)) t b)) + \ (\k\length p - 1. nth_digit (s k && z (l k)) t b)" + unfolding sum_ssub_nzero.simps + by (simp add: atMost_atLeast0 sum_mono) + also have "... \ (\k\length p - 1. Seq k t)" + using bitAND_sum_lt[where ?x = "\k. z (l k)"] by blast + finally have "(\S- p j (\k. nth_digit (s k && z (l k)) t b)) \ 1" + using sum_sk_bound[of t] Seq_def + using dual_order.trans by blast + } + thus ?thesis using Seq_def by auto +qed + +lemma block_sum_szero_bound: + shows "\t. (\S0 p j (\k. nth_digit (s k && (e - z (l k))) t b)) \ 1" +proof - + { + fix t + have "(\S0 p j (\k. nth_digit (s k && e - z (l k)) t b)) + \ (\k\length p - 1. nth_digit (s k && e - z (l k)) t b)" + unfolding sum_ssub_zero.simps + by (simp add: atMost_atLeast0 sum_mono) + also have "... \ (\k\length p - 1. Seq k t)" + using bitAND_sum_lt[where ?x = "\k. e - z (l k)"] by blast + finally have "(\S0 p j (\k. nth_digit (s k && e - z (l k)) t b)) \ 1" + using sum_sk_bound[of t] Seq_def + using dual_order.trans by blast + } + thus ?thesis using Seq_def by auto +qed + +lemma sum_radd_nth_digit_commute: + shows "nth_digit (\R+ p l (\k. s k)) t b = \R+ p l (\k. nth_digit (s k) t b)" +proof - + have a1: "\t. \k\length p - 1. nth_digit (s k) t b < 2^c" + using sk_bound h0 by (meson le_less_trans) + + have a2: "\t. (\R+ p l (\k. nth_digit (s k) t b)) < 2^c" + using block_sum_radd_bound h0 by (meson le_less_trans) + + show ?thesis + using a1 a2 c_gt0 b_def unfolding sum_radd.simps + using sum_nth_digit_commute[where ?g = "\l k. isadd (p ! k) \ l = modifies (p ! k)"] + by blast +qed + +lemma sum_rsub_nth_digit_commute: + shows "nth_digit (\R- p l (\k. s k && z l)) t b + = \R- p l (\k. nth_digit (s k && z l) t b)" +proof - + have a1: "\t. \k\length p - 1. nth_digit (s k && z l) t b < 2^c" + using sk_bitAND_bound[where ?x = "\k. z l"] h0 le_less_trans by blast + + have a2: "\t. (\R- p l (\k. nth_digit (s k && z l) t b)) < 2^c" + using block_sum_rsub_bound h0 by (meson le_less_trans) + + show ?thesis + using a1 a2 c_gt0 b_def unfolding sum_rsub.simps + using sum_nth_digit_commute + [where ?g = "\l k. issub (p ! k) \ l = modifies (p ! k)" and ?fct = "\k. s k && z l"] + by blast +qed + +lemma sum_sadd_nth_digit_commute: + shows "nth_digit (\S+ p j (\k. s k)) t b = \S+ p j (\k. nth_digit (s k) t b)" +proof - + have a1: "\t. \k\length p - 1. nth_digit (s k) t b < 2^c" + using sk_bound h0 by (meson le_less_trans) + + have a2: "\t. (\S+ p j (\k. nth_digit (s k) t b)) < 2^c" + using block_sum_sadd_bound h0 by (meson le_less_trans) + + show ?thesis + using a1 a2 b_def c_gt0 unfolding sum_sadd.simps + using sum_nth_digit_commute[where ?g = "\j k. isadd (p ! k) \ j = goes_to (p ! k)"] + by blast +qed + +lemma sum_ssub_nth_digit_commute: + shows "nth_digit (\S- p j (\k. s k && z (l k))) t b + = \S- p j (\k. nth_digit (s k && z (l k)) t b)" +proof - + have a1: "\t. \k\length p - 1. nth_digit (s k && z (l k)) t b < 2^c" + using sk_bitAND_bound[where ?x = "\k. z (l k)"] h0 le_less_trans by blast + + have a2: "\t. (\S- p j (\k. nth_digit (s k && z (l k)) t b)) < 2^c" + using block_sum_ssub_bound h0 by (meson le_less_trans) + + show ?thesis + using a1 a2 b_def c_gt0 unfolding sum_ssub_nzero.simps + using sum_nth_digit_commute + [where ?g = "\j k. issub (p ! k) \ j = goes_to (p ! k)" and ?fct = "\k. s k && z (l k)"] + by blast +qed + +lemma sum_szero_nth_digit_commute: + shows "nth_digit (\S0 p j (\k. s k && (e - z (l k)))) t b + = \S0 p j (\k. nth_digit (s k && (e - z (l k))) t b)" +proof - + have a1: "\t. \k\length p - 1. nth_digit (s k && (e - z (l k))) t b < 2^c" + using sk_bitAND_bound[where ?x = "\k. e - z (l k)"] h0 le_less_trans by blast + + have a2: "\t. (\S0 p j (\k. nth_digit (s k && (e - z (l k))) t b)) < 2^c" + using block_sum_szero_bound h0 by (meson le_less_trans) + + show ?thesis + using a1 a2 b_def c_gt0 unfolding sum_ssub_zero.simps + using sum_nth_digit_commute + [where ?g = "\j k. issub (p ! k) \ j = goes_to_alt (p ! k)" and ?fct = "\k. s k && e - z (l k)"] + by blast +qed + +lemma block_bound_impl_fst_digit_zero: + assumes "nth_digit x t b \ 1" + shows "(nth_digit x t b) \ c = 0" + using assms apply (auto simp: nth_bit_def) + by (metis (no_types, opaque_lifting) c_gt0 div_le_dividend le_0_eq le_Suc_eq mod_0 mod_Suc + mod_div_trivial numeral_2_eq_2 power_eq_0_iff power_mod) + +lemma sum_radd_block_bound: + shows "nth_digit (\R+ p l (\k. s k)) t b \ 1" + using block_sum_radd_bound sum_radd_nth_digit_commute by auto +lemma sum_radd_fst_digit_zero: + shows "(nth_digit (\R+ p l s) t b) \ c = 0" + using sum_radd_block_bound block_bound_impl_fst_digit_zero by auto + +lemma sum_sadd_block_bound: + shows "nth_digit (\S+ p j (\k. s k)) t b \ 1" + using block_sum_sadd_bound sum_sadd_nth_digit_commute by auto +lemma sum_sadd_fst_digit_zero: + shows "(nth_digit (\S+ p j s) t b) \ c = 0" + using sum_sadd_block_bound block_bound_impl_fst_digit_zero by auto + +lemma sum_ssub_block_bound: + shows "nth_digit (\S- p j (\k. s k && z (l k))) t b \ 1" + using block_sum_ssub_bound sum_ssub_nth_digit_commute by auto +lemma sum_ssub_fst_digit_zero: + shows "(nth_digit (\S- p j (\k. s k && z (l k))) t b) \ c = 0" + using sum_ssub_block_bound block_bound_impl_fst_digit_zero by auto + +lemma sum_szero_block_bound: + shows "nth_digit (\S0 p j (\k. s k && (e - z (l k)))) t b \ 1" + using block_sum_szero_bound sum_szero_nth_digit_commute by auto +lemma sum_szero_fst_digit_zero: + shows "(nth_digit (\S0 p j (\k. s k && (e - z (l k)))) t b) \ c = 0" + using sum_szero_block_bound block_bound_impl_fst_digit_zero by auto + +lemma sum_rsub_special_block_bound: + shows "nth_digit (\R- p l (\k. s k)) t b \ 1" +proof - + have a1: "\t k. k \ length p - 1 \ nth_digit (s k) t b < 2^c" + using sk_bound h0 le_less_trans by blast + have a2: "\t. \R- p l (\k. nth_digit (s k) t b) < 2^c" + using block_sum_rsub_special_bound h0 le_less_trans by blast + + have "nth_digit (\R- p l (\k. s k)) t b = \R- p l (\k. nth_digit (s k) t b)" + using a1 a2 b_def c_gt0 unfolding sum_rsub.simps + using sum_nth_digit_commute[where ?g = "\l k. issub (p ! k) \ l = modifies (p ! k)"] + by blast + + thus ?thesis + using block_sum_rsub_special_bound by auto +qed + +lemma sum_state_special_block_bound: + shows "nth_digit (\S+ p j (\k. s k) + + \S0 p j (\k. s k && (e - z (l k)))) t b \ 1" +proof - + have aux_sum_zero: + "\S0 p j (\k. nth_digit (s k) t b && nth_digit (e - z (l k)) t b) + \ \S0 p j (\k. nth_digit (s k) t b)" + unfolding sum_ssub_zero.simps + by (auto simp: bitAND_lt sum_mono) + + have aux_addsub_excl: "(if isadd (p ! k) then Seq k t else 0) + + (if issub (p ! k) then Seq k t else 0) + = (if isadd (p ! k) \ issub (p ! k) then Seq k t else 0)" for k + by auto + + have aux_sum_add_lt: + "\S+ p j (\k. Seq k t) \ (\k = 0..length p - 1. if isadd (p ! k) then Seq k t else 0)" + unfolding sum_sadd.simps by (simp add: sum_mono) + have aux_sum_sub_lt: + "\S0 p j (\k. Seq k t) \ (\k = 0..length p - 1. if issub (p ! k) then Seq k t else 0)" + unfolding sum_ssub_zero.simps by (simp add: sum_mono) + + have "nth_digit (\S+ p j (\k. s k) + + \S0 p j (\k. s k && e - z (l k))) t b + = nth_digit (\S+ p j (\k. s k)) t b + + nth_digit (\S0 p j (\k. s k && e - z (l k))) t b" + using sum_sadd_fst_digit_zero sum_szero_fst_digit_zero block_additivity + by (auto simp: c_gt0 b_def) + also have "... = \S+ p j (\k. nth_digit (s k) t b) + + \S0 p j (\k. nth_digit (s k && e - z (l k)) t b)" + by (simp add: sum_sadd_nth_digit_commute sum_szero_nth_digit_commute) + also have "... \ \S+ p j (\k. Seq k t) + \S0 p j (\k. Seq k t)" + using bitAND_nth_digit_commute aux_sum_zero + unfolding Seq_def by (simp add: b_def) + also have "... \ (\k = 0..length p - 1. if isadd (p ! k) then Seq k t else 0) + + (\k = 0..length p - 1. if issub (p ! k) then Seq k t else 0)" + using aux_sum_add_lt aux_sum_sub_lt by auto + also have "... = (\k\length p - 1. if (isadd (p ! k) \ issub (p ! k)) + then Seq k t else 0)" + using aux_addsub_excl + using sum.distrib[where ?g = "\k. if isadd (p ! k) then Seq k t else 0" + and ?h = "\k. if issub (p ! k) then Seq k t else 0"] + by (auto simp: aux_addsub_excl atMost_atLeast0) + also have "... \ (\k\length p - 1. Seq k t)" + by (smt eq_iff le0 sum_mono) + + finally show ?thesis using sum_sk_bound[of t] Seq_def by auto +qed +lemma sum_state_special_fst_digit_zero: + shows "(nth_digit (\S+ p j (\k. s k) + + \S0 p j (\k. s k && (e - z (modifies (p!k))))) t b) \ c = 0" + using sum_state_special_block_bound block_bound_impl_fst_digit_zero by auto + +text \Main three redution lemmas: Zero Indicators, Registers, States\ + +lemma Z: + assumes "l 0 then Suc 0 else 0)" +proof - + have cond: "2^c * (z l) = (r l + d) && f" using m_eq mask_equations_def assms by auto + have d_block: "\t\q. nth_digit d t b = 2^c - 1" using d_def b_def + less_imp_diff_less nth_digit_gen_power_series[of "\_. 2^c-1" "c"] c_gt0 by auto + have rl_bound: "t\q \ nth_digit (r l) t b \ c= 0" for t by (simp add: assms rl_fst_digit_zero) + have f_block: "\t\q. nth_digit f t b = 2^c" + using f_def b_def less_imp_diff_less nth_digit_gen_power_series[of "\_. 2^c" c] c_gt0 by auto + then have "\t\q.\k k = 0" by (simp add: aux_powertwo_digits) + + moreover have AND_gen: "\t\q.\k\c. nth_digit ((r l + d) && f) t b \ k = (nth_digit (r l + d) t b \ k) * nth_digit f t b \ k" + using b_def digit_gen_pow2_reduct bitAND_digit_mult digit_gen_pow2_reduct le_imp_less_Suc by presburger + ultimately have "\t\q.\k k = 0" using f_def by auto + moreover have "(r l + d) && f < b ^ Suc q" using lm0245[of "r l + d" f] masks_leq[of "(r l + d) && f" f] f_def + proof- + have "2 < b" using b_def c_gt0 gr0_conv_Suc not_less_iff_gr_or_eq by fastforce + then have "b^u + b^u < b* b^ u" for u using zero_less_power[of b u] mult_less_mono1[of 2 b "b^u"] by linarith + then have "(\t\{..t\{..q. b^q" "{..t\{.. f" using lm0245[of "r l + d" f] masks_leq[of "(r l + d) && f" f] by auto + ultimately show ?thesis by auto + qed + then have rldf0: "t>q \ nth_digit ((r l + d) && f) t b = 0" for t using nth_digit_def[of "r l + d && f" t b] + div_less[of "r l + d && f" "b^t"] b_def power_increasing[of "Suc q" t b] by auto + moreover have "\t>q.\k k = 0" using aux_lt_implies_mask rldf0 by fastforce + ultimately have AND_zero: "\t.\k k = 0" using leI by blast + + have "0 k< Suc c \ nth_digit (z l) t b \ k = nth_digit ((r l + d) && f) (Suc t) b \ (k - 1)" + for k using b_def nth_digit_bound digit_gen_pow2_reduct[of k "Suc c" "z l" t] aux_digit_shift[of "z l" c "t + c * t + k"] + digit_gen_pow2_reduct[of "k-1" "Suc c" "z l * 2^c" "Suc t"] cond by (simp add: add.commute add.left_commute mult.commute) + then have aux: "0 k< Suc c \ nth_digit (z l) t b \ k = 0" for k using AND_zero by auto + have zl_formula: "nth_digit (z l) t b = nth_digit (z l) t b \ 0 " + using b_def digit_sum_repr[of "nth_digit (z l) t b" "Suc c"] + proof - + have "nth_digit (z l) t b < 2 ^ Suc c + \ nth_digit (z l) t b = (\k\{0.. k * 2 ^ k)" + using b_def digit_sum_repr[of "nth_digit (z l) t b" "Suc c"] + by (simp add: atLeast0LessThan) + hence "nth_digit (z l) t b < 2 ^ Suc c + \ nth_digit (z l) t b = nth_digit (z l) t b \ 0 + + (\k\{0<.. k * 2 ^ k)" + by (metis One_nat_def atLeastSucLessThan_greaterThanLessThan mult_numeral_1_right + numeral_1_eq_Suc_0 power_0 sum.atLeast_Suc_lessThan zero_less_Suc) + thus ?thesis using aux nth_digit_bound b_def by auto + qed + + consider (tleq) "t\q" |(tgq) "t>q" by linarith + then show ?thesis + proof cases + case tleq + then have t_bound: "t \ q" by auto + have "nth_digit ((r l + d) && f) t b \ c = (nth_digit (r l + d) t b \ c)" + using f_block bitAND_single_digit AND_gen t_bound by auto + moreover have "nth_digit (r l + d && f) t b < 2 ^ Suc c" using nth_digit_def b_def by simp + ultimately have AND_all:"nth_digit ((r l + d) && f) t b = (nth_digit (r l + d) t b \ c) * 2^c" using AND_gen AND_zero + using digit_sum_repr[of "nth_digit ((r l + d) && f) t b" "Suc c"] by auto + + then have "\k k = 0" using cond AND_zero by metis + moreover have "nth_digit (2^c * (z l)) t b \ c = nth_digit (z l) t b \ 0" + using digit_gen_pow2_reduct[of c "Suc c" " (2^c * (z l))" t] + digit_gen_pow2_reduct[of 0 "Suc c" "z l" t] b_def by (simp add: aux_digit_shift mult.commute) + ultimately have zl0: "nth_digit (2^c * (z l)) t b = 2^c * nth_digit (z l) t b \ 0" + using digit_sum_repr[of "nth_digit (2^c * (z l)) t b" "Suc c"] nth_digit_bound b_def by auto + + have "nth_digit (2^c * (z l)) t b = 2^c * nth_digit (z l) t b" using zl0 zl_formula by auto + then have zl_block: "nth_digit (z l) t b = nth_digit (r l + d) t b \ c" using AND_all cond by auto + + consider (g0) "Req l t > 0" | (e0) "Req l t = 0 " by auto + then show ?thesis + proof cases + case e0 + show ?thesis using e0 apply(auto simp add: Req_def Zeq_def) subgoal + proof- + assume asm: "nth_digit (r l) t b = 0" + have add:"((nth_digit d t b) + (nth_digit (r l) t b)) \ c = 0" by (simp add: asm d_block nth_bit_def t_bound) + from d_block t_bound have "nth_digit d (t-1) b \ c = 0" + using add asm by auto + then have "(nth_digit (r l + d) t b) \ c = 0" + using add digit_wise_block_additivity[of "r l" "t" "c" "d" "c"] rl_bound[of "t-1"] b_def asm t_bound c_gt0 by auto + then show ?thesis using zl_block by simp + qed done + next + case g0 + show ?thesis using g0 apply(auto simp add: Req_def Zeq_def) subgoal + proof - + assume "0 < nth_digit (r l) t b" + then obtain k0 where k0_def: "nth_digit (r l) t b \ k0 = 1" using aux0_digit_wise_equiv by auto + then have "k0\c" using nth_digit_bound[of "r l" t c] b_def aux_lt_implies_mask by (metis Suc_leI leI zero_neq_one) + then have k0bound: "k0k k = 1" using d_block t_bound nth_bit_def[of "nth_digit d t b"] + by (metis One_nat_def Suc_1 Suc_diff_Suc Suc_pred dmask_aux even_add even_power odd_iff_mod_2_eq_one + one_mod_two_eq_one plus_1_eq_Suc zero_less_Suc zero_less_power) + ultimately have "nth_digit d t b \ k0 = 1" by simp + then have "bin_carry (nth_digit d t b) (nth_digit (r l) t b) (Suc k0) = 1" using + k0_def sum_carry_formula carry_bounded less_eq_Suc_le by simp + moreover have "\n. Suc k0 \ n \ n < c \ bin_carry (nth_digit d t b) (nth_digit (r l) t b) n = + Suc 0 \ bin_carry (nth_digit d t b) (nth_digit (r l) t b) (Suc n) = Suc 0" subgoal for n + proof- + assume "n c = 1" + using sum_digit_formula[of "nth_digit d t b" "nth_digit (r l) t b" c] + d_block nth_bit_def t_bound assms rl_fst_digit_zero by auto + + from d_block t_bound have "nth_digit d (t-1) b \ c = 0" + by (smt aux_lt_implies_mask diff_le_self diff_less le_eq_less_or_eq le_trans + zero_less_numeral zero_less_one zero_less_power) + then have "(nth_digit (r l + d) t b) \ c = 1" using add b_def t_bound + block_additivity assms rl_fst_digit_zero c_gt0 d_block by (simp add: add.commute) + then show ?thesis using zl_block by simp + qed done + qed + next + case tgq + then have t_bound: "qk k = 0" using cond AND_zero by simp + (* the next two statements are already proven above but here they rely on rl0 *) + moreover have "nth_digit (2^c * (z l)) t b \ c = nth_digit (z l) t b \ 0" + using digit_gen_pow2_reduct[of c "Suc c" " (2^c * (z l))" t] + digit_gen_pow2_reduct[of 0 "Suc c" "z l" t] b_def by (simp add: aux_digit_shift mult.commute) + ultimately have zl0: "nth_digit (2^c * (z l)) t b = 2^c * nth_digit (z l) t b \ 0" + using digit_sum_repr[of "nth_digit (2^c * (z l)) t b" "Suc c"] nth_digit_bound b_def by auto + have "0 k< Suc c \ nth_digit (z l) t b \ k = nth_digit ((r l + d) && f) (Suc t) b \ (k - 1)" + for k using b_def nth_digit_bound digit_gen_pow2_reduct[of k "Suc c" "z l" t] aux_digit_shift[of "z l" c "t + c * t + k"] + digit_gen_pow2_reduct[of "k-1" "Suc c" "z l * 2^c" "Suc t"] cond by (simp add: add.commute add.left_commute mult.commute) + (* Using some simplification lemmas, the result follows *) + then show ?thesis using Zeq_def Req_def cond rl0 zl0 rldf0 zl_formula t_bound by auto + qed +qed + +lemma zl_le_rl: "l z l \ r l" for l +proof - + assume l: "l < n" + have "Zeq l t \ Req l t" for t using Z l by auto + hence "nth_digit (z l) t b \ nth_digit (r l) t b" for t + using Zeq_def Req_def by auto + thus ?thesis using digitwise_leq b_gt1 by auto +qed + + +lemma modifies_valid: "\k\m. modifies (p!k) < n" +proof - + have reg_check: "program_register_check p n" + using is_val by (cases ic, auto simp: is_valid_initial_def n_def is_valid_def) + { + fix k + assume "k \ m" + then have "p ! k \ set p" + by (metis \k \ m\ add_eq_if diff_le_self is_val le_antisym le_trans m_def + n_not_Suc_n not_less not_less0 nth_mem p_contains) + then have "instruction_register_check n (p ! k)" + using reg_check by (auto simp: list_all_def) + then have "modifies (p!k) < n" by (cases "p ! k", auto simp: n_gt0) + } + thus ?thesis by auto +qed + +lemma seq_bound: "k \ length p - 1 \ Seq k t \ 1" + using sk_bound Seq_def by blast + +lemma skzl_bitAND_to_mult: + assumes "k \ length p - 1" + assumes "l < n" + shows "nth_digit (z l) t b && nth_digit (s k) t b = (Zeq l t) * Seq k t" +proof - + have "nth_digit (z l) t b && nth_digit (s k) t b = (Zeq l t) && Seq k t" + using Zeq_def Seq_def by simp + also have "... = (Zeq l t) * Seq k t" + using bitAND_single_bit_mult_equiv[of "(Zeq l t)" "Seq k t"] seq_bound Z assms by auto + finally show ?thesis by auto +qed + +lemma skzl_bitAND_to_mult2: + assumes "k \ length p - 1" + assumes "\k \ length p - 1. l k < n" + shows "(1 - nth_digit (z (l k)) t b) && nth_digit (s k) t b + = (1 - Zeq (l k) t) * Seq k t" +proof - + have "(1 - nth_digit (z (l k)) t b) && nth_digit (s k) t b + = (1 - Zeq (l k) t) && Seq k t" + using Zeq_def Seq_def by simp + also have "... = (1 - Zeq (l k) t) * Seq k t" + using bitAND_single_bit_mult_equiv[of "(1 - Zeq (l k) t)" "Seq k t"] seq_bound Z assms by auto + finally show ?thesis by auto +qed + +lemma state_equations_digit_commute: + assumes "t < q" and "j \ m" + defines "l \ \k. modifies (p!k)" + shows "nth_digit (s j) (Suc t) b = + (\S+ p j (\k. Seq k t)) + + (\S- p j (\k. Zeq (l k) t * Seq k t)) + + (\S0 p j (\k. (1 - Zeq (l k) t) * Seq k t))" +proof - + define o' :: nat where "o' \ if j = 0 then 1 else 0" + have o'_div: "o' div b = 0" using b_gt1 by (auto simp: o'_def) + + have l: "\k\length p - 1. (l k) < n" + using l_def by (auto simp: m_def modifies_valid) + + have "\k. Suc 0 < 2^c" using c_gt0 h0 by auto + hence e_aux: "\tt. nth_digit e tt b = (if tt\q then Suc 0 else 0)" + using e_def b_def c_gt0 nth_digit_gen_power_series[of "\k. Suc 0" "c" "q"] by auto + have zl_bounded: "k\m \ \t'. nth_digit (z (l k)) t' b \ nth_digit e t' b" for k + proof - + assume "k\m" + from m_eq have "\l e" using mask_equations_def by auto + then have "\lt'. (z l) \ t' \ e \ t'" using masks_leq_equiv by auto + then have k_lt_c: "\lt'. \k' k' + \ nth_digit e t' b \ k'" + using digit_gen_pow2_reduct by (auto simp: b_def) (metis power_Suc) + + have "k\Suc c \ x mod (2 ^ Suc c) div 2 ^ k = 0" for k x::nat + by (metis Euclidean_Division.div_eq_0_iff One_nat_def Suc_le_lessD b_gt1 b_def less_le_trans + mod_less_divisor nat_power_less_imp_less not_less not_less0 not_less_eq_eq + numeral_2_eq_2 zero_less_Suc) + then have "\k\Suc c. nth_digit x y b \ k = 0" for x y + using b_def nth_bit_def nth_digit_def by auto + then have k_gt_c: "\lt'. \k'\Suc c. nth_digit (z l) t' b \ k' + \ nth_digit e t' b \ k'" + by auto + from k_lt_c k_gt_c have "\lt'. nth_digit (z l) t' b \ nth_digit e t' b" + using bitwise_leq by (meson not_le) + thus ?thesis by (auto simp: modifies_valid l_def `k\m`) + qed + + have "\t k. k\m \ nth_digit (e - z (l k)) t b = + nth_digit e t b - nth_digit (z (l k)) t b" + using zl_bounded block_subtractivity by (auto simp: c_gt0 b_def l_def) + then have sum_szero_aux: + "\t k. t k\m \ nth_digit (e - z (l k)) t b = 1-nth_digit (z (l k)) t b" + using e_aux by auto + + have skzl_bound2: "\k\length p - 1. (l k) < n \ + \t. \k\length p - 1. nth_digit (s k && (e - z (l k))) t b < 2^c" + proof - + assume l: "\k\length p - 1. (l k) < n" + have "\t. \k\length p - 1. nth_digit (s k && (e - z (l k))) t b + = nth_digit (s k) t b && nth_digit (e - z (l k)) t b" + using bitAND_nth_digit_commute Zeq_def b_def by auto + + moreover have "\tk\length p - 1. + nth_digit (s k) t b && nth_digit (e - z (l k)) t b + = nth_digit (s k) t b && (1 - nth_digit (z (l k)) t b)" + using sum_szero_aux by (simp add: m_def) + + moreover have "\t. \k\length p - 1. + nth_digit (s k) t b && (1 - nth_digit (z (l k)) t b) + \ nth_digit (s k) t b" + using Z l using lm0245 masks_leq by (simp add: lm0244) + + moreover have "\t. \k\length p - 1. nth_digit (s k) t b < 2^c" + using sk_bound h0 by (meson le_less_trans) + + ultimately show ?thesis + using le_less_trans by (metis lm0244 masks_leq) + qed + + (* START *) + have "s j = o' + b*\S+ p j (\k. s k) + b*\S- p j (\k. s k && z (modifies (p!k))) + + b*\S0 p j (\k. s k && (e - z (modifies (p!k))))" + using s_eq state_equations_def \j\m\ by (auto simp: o'_def) + then have "s j div b^Suc t mod b = + ( o' + b*\S+ p j (\k. s k) + + b*\S- p j (\k. s k && z (modifies (p!k))) + + b*\S0 p j (\k. s k && (e - z (modifies (p!k)))) ) div b div b^t mod b" + by (auto simp: algebra_simps div_mult2_eq) + also have "... = (\S+ p j (\k. s k) + + \S- p j (\k. s k && z (modifies (p!k))) + + \S0 p j (\k. s k && (e - z (modifies (p!k)))) ) div b^t mod b" + using o'_div + by (auto simp: algebra_simps div_mult2_eq) + (smt Nat.add_0_right add_mult_distrib2 b_gt1 div_mult_self2 gr_implies_not0) + (* Interchanged order in the following with add.commute to ease next step *) + also have "... = nth_digit (\S- p j (\k. s k && z (l k)) + + \S+ p j (\k. s k) + + \S0 p j (\k. s k && (e - z (l k)))) t b" + by (auto simp: nth_digit_def l_def add.commute) + + (* STEP 2: Commute nth_digit to the inside of all expressions. *) + also have "... = nth_digit (\S- p j (\k. s k && z (l k))) t b + + nth_digit (\S+ p j (\k. s k) + + \S0 p j (\k. s k && (e - z (l k)))) t b" + using block_additivity sum_ssub_fst_digit_zero sum_state_special_fst_digit_zero + by (auto simp: l_def c_gt0 b_def add.assoc) + also have "... = nth_digit (\S+ p j (\k. s k)) t b + + nth_digit (\S- p j (\k. s k && z (l k))) t b + + nth_digit (\S0 p j (\k. s k && (e - z (l k)))) t b" + using block_additivity sum_sadd_fst_digit_zero sum_szero_fst_digit_zero + by (auto simp: l_def c_gt0 b_def) + also have "... = nth_digit (\S+ p j (\k. s k)) t b + + (\S- p j (\k. nth_digit (s k && z (l k)) t b)) + + nth_digit (\S0 p j (\k. s k && (e - z (l k)))) t b" + using sum_ssub_nth_digit_commute by auto + also have "... = nth_digit (\S+ p j (\k. s k)) t b + + \S- p j (\k. nth_digit (s k && z (l k)) t b) + + \S0 p j (\k. nth_digit (s k && (e - z (l k))) t b)" + using l_def l skzl_bound2 sum_szero_nth_digit_commute by (auto) + also have "... = \S+ p j (\k. nth_digit (s k) t b) + + \S- p j (\k. nth_digit (s k && z (l k)) t b) + + \S0 p j (\k. nth_digit (s k && (e - z (l k))) t b)" + using sum_sadd_nth_digit_commute by auto + also have "... = \S+ p j (\k. nth_digit (s k) t b) + + \S- p j (\k. nth_digit (z (l k)) t b && nth_digit (s k) t b) + + \S0 p j (\k. (nth_digit (e - z (l k)) t b) && nth_digit (s k) t b)" + using bitAND_nth_digit_commute b_def by (auto simp: bitAND_commutes) + also have "... = (\S+ p j (\k. nth_digit (s k) t b)) + + (\S- p j (\k. nth_digit (z (l k)) t b && nth_digit (s k) t b)) + + (\S0 p j (\k. (1 - nth_digit (z (l k)) t b) && nth_digit (s k) t b))" + using sum_szero_aux assms sum_ssub_zero.simps m_def \t + apply (auto) using sum.cong atLeastAtMost_iff by smt + (* this ATP proof ensures that the \S0 only uses the sum_szero_aux lemma with k \ m + because the summation of sum_ssub_zero ranges from k = 0 .. (length p - 1) *) + ultimately have "s j div (b ^ Suc t) mod b = + (\S+ p j (\k. nth_digit (s k) t b)) + + (\S- p j (\k. nth_digit (z (l k)) t b && nth_digit (s k) t b)) + + (\S0 p j (\k. (1 - nth_digit (z (l k)) t b) && nth_digit (s k) t b))" + by auto + + moreover have "(\S- p j (\k. nth_digit (z (l k)) t b && nth_digit (s k) t b)) + = (\S- p j (\k. Zeq (l k) t * Seq k t))" + using skzl_bitAND_to_mult sum_ssub_nzero.simps l + by (smt atLeastAtMost_iff sum.cong) + + moreover have "(\S0 p j (\k. (1 - nth_digit (z (l k)) t b) && nth_digit (s k) t b)) + = (\S0 p j (\k. (1 - Zeq (l k) t) * Seq k t))" + using skzl_bitAND_to_mult2 sum_ssub_zero.simps l + by (auto) (smt atLeastAtMost_iff sum.cong) + + ultimately have "nth_digit (s j) (Suc t) b = + (\S+ p j (\k. Seq k t)) + + (\S- p j (\k. Zeq (l k) t * Seq k t)) + + (\S0 p j (\k. (1 - Zeq (l k) t) * Seq k t))" + using Seq_def nth_digit_def by auto + + thus ?thesis by auto +qed + +lemma aux_nocarry_sk: + assumes "t\q" + shows "i \ j \ i\m \ j\m \ nth_digit (s i) t b * nth_digit (s j) t b = 0" +proof (cases "t=q") + case True + have "j < m \ Seq j q = 0" for j using s_bound Seq_def nth_digit_def by auto + then show ?thesis using True Seq_def apply auto by (metis le_less less_nat_zero_code) +next + case False + hence "k\m \ nth_digit (s k) t b = 1 \ + (\j\m. j \ k \ nth_digit (s j) t b = 0)" for k + using states_unique_RAW[of "t"] Seq_def assms by auto + thus ?thesis + by (auto) (metis One_nat_def le_neq_implies_less m_def not_less_eq sk_bound) +qed + +lemma nocarry_sk: + assumes "i \ j" and "i\m" and "j\m" + shows "(s i) \ k * (s j) \ k = 0" +proof - + have reduct: "(s i) \ k = nth_digit (s i) (k div Suc c) b \ (k mod Suc c)" for i + using digit_gen_pow2_reduct[of "k mod Suc c" "Suc c" "s i" "k div Suc c"] b_def + using mod_less_divisor zero_less_Suc by presburger + have "k div Suc c \ q \ + nth_digit (s i) (k div Suc c) b * nth_digit (s j) (k div Suc c) b = 0" + using aux_nocarry_sk assms by auto + moreover have "k div Suc c > q \ + nth_digit (s i) (k div Suc c) b * nth_digit (s j) (k div Suc c) b = 0" + using nth_digit_def s_bound apply auto + using b_gt1 div_greater_zero_iff leD le_less less_trans mod_less neq0_conv power_increasing_iff + by (smt assms) + ultimately have "nth_digit (s i) (k div Suc c) b \ (k mod Suc c) + * nth_digit (s j) (k div Suc c) b \ (k mod Suc c) = 0" + using nth_bit_def by auto + thus ?thesis using reduct[of "i"] reduct[of "j"] by auto +qed + + +lemma commute_sum_rsub_bitAND: "\R- p l (\k. s k && z l) = \R- p l (\k. s k) && z l" +proof - + show ?thesis apply (auto simp: sum_rsub.simps) + using m_def nocarry_sk aux_commute_bitAND_sum_if[of "m" + "s" "\k. issub (p ! k) \ l = modifies (p ! k)" "z l"] + by (auto simp add: atMost_atLeast0) +qed + +lemma sum_rsub_bound: "l \R- p l (\k. s k && z l) \ r l + \R+ p l s" +proof - + assume "lR- p l (\k. s k) && z l \ z l" by (auto simp: lm0245 masks_leq) + also have "... \ r l" using zl_le_rl `lObtaining single step register relations from multiple step register relations \ +lemma mult_to_single_reg: + "c>0 \ l Req l (Suc t) = Req l t + (\R+ p l (\k. Seq k t)) + - (\R- p l (\k. (Zeq l t) * Seq k t))" for l t +proof - + assume l: "l 0" + + have a_div: "a div b = 0" using c_eq rm_constants_def B_def by auto + + have subtract_bound: "\t'. nth_digit (\R- p l (\k. s k && z l)) t' b + \ nth_digit (r l + \R+ p l (\k. s k)) t' b" + proof - + { + fix t' + have "nth_digit (z l) t' b \ nth_digit (r l) t' b" + using Zeq_def Req_def Z l by auto + then have "nth_digit (\R- p l (\k. s k)) t' b && nth_digit (z l) t' b + \ nth_digit (r l) t' b" + using sum_rsub_special_block_bound + by (meson dual_order.trans lm0245 masks_leq) + then have "nth_digit (\R- p l (\k. s k && z l)) t' b + \ nth_digit (r l) t' b" + using commute_sum_rsub_bitAND bitAND_nth_digit_commute b_def by auto + then have "nth_digit (\R- p l (\k. s k && z l)) t' b + \ nth_digit (r l) t' b + nth_digit (\R+ p l (\k. s k)) t' b" + by auto + then have "nth_digit (\R- p l (\k. s k && z l)) t' b + \ nth_digit (r l + \R+ p l (\k. s k)) t' b" + using block_additivity rl_fst_digit_zero sum_radd_fst_digit_zero + by (auto simp: b_def c l) + } + then show ?thesis by auto + qed + + define a' where "a' \ (if l = 0 then a else 0)" + have a'_div: "a' div b = 0" using a_div a'_def by auto + + (* START PROOF CHAIN *) + (* STEP 1: Remove a' using divisibility properties. *) + have "r l div (b * b ^ t) mod b = + (a' + b*r l + b*\R+ p l (\k. s k) - b*\R- p l (\k. s k && z l)) div (b * b ^ t) mod b" + using r_eq reg_equations_def by (auto simp: a'_def l) + also have "... = + (a' + b * (r l + \R+ p l (\k. s k) - \R- p l (\k. s k && z l))) div b div b^t mod b" + by (auto simp: algebra_simps div_mult2_eq) + (metis Nat.add_diff_assoc add_mult_distrib2 mult_le_mono2 sum_rsub_bound l) + also have "... = + ((r l + \R+ p l (\k. s k) - \R- p l (\k. s k && z l)) + a' div b) div b ^ t mod b" + using b_gt1 by auto + also have "... = (r l + \R+ p l (\k. s k) - \R- p l (\k. s k && z l)) div b ^ t mod b" + using a'_div by auto + also have "... = nth_digit (r l + \R+ p l (\k. s k) - \R- p l (\k. s k && z l)) t b" + using nth_digit_def by auto + + (* STEP 2: Commute nth_digit to the inside of all expressions. *) + also have "... = nth_digit (r l + \R+ p l (\k. s k)) t b + - nth_digit (\R- p l (\k. s k && z l)) t b" + using block_subtractivity subtract_bound + by (auto simp: c b_def) + also have "... = nth_digit (r l) t b + + nth_digit (\R+ p l (\k. s k)) t b + - nth_digit (\R- p l (\k. s k && z l)) t b" + using block_additivity rl_fst_digit_zero sum_radd_fst_digit_zero + by (auto simp: l b_def c) + also have "... = nth_digit (r l) t b + + \R+ p l (\k. nth_digit (s k) t b) + - \R- p l (\k. nth_digit (s k && z l) t b)" + using sum_radd_nth_digit_commute + using sum_rsub_nth_digit_commute + by auto + ultimately have "r l div (b * b ^ t) mod b = + (nth_digit (r l) t b) + + \R+ p l (\k. nth_digit (s k) t b) + - \R- p l (\k. nth_digit (z l) t b && nth_digit (s k) t b)" + using bitAND_nth_digit_commute b_def by (simp add: bitAND_commutes) + + then show ?thesis using Req_def Seq_def nth_digit_def skzl_bitAND_to_mult l + by (auto simp: sum_rsub.simps) (smt atLeastAtMost_iff sum.cong) +qed + +text \Obtaining single step state relations from multiple step state relations\ +lemma mult_to_single_state: + fixes t j :: nat + defines "l \ \k. modifies (p!k)" + shows "j\m \ t Seq j (Suc t) = (\S+ p j (\k. Seq k t)) + + (\S- p j (\k. Zeq (l k) t * Seq k t)) + + (\S0 p j (\k. (1 - Zeq (l k) t) * Seq k t))" +proof - + assume "j \ m" + assume "t < q" + + have "nth_digit (s j) (Suc t) b = + (\S+ p j (\k. Seq k t)) + + (\S- p j (\k. Zeq (l k) t * Seq k t)) + + (\S0 p j (\k. (1 - Zeq (l k) t) * Seq k t))" + using state_equations_digit_commute \j\m\ \t l_def by auto + + then show ?thesis using nth_digit_def l_def Seq_def by auto +qed + +text \Conclusion: + The central equivalence showing that the cell entries obtained from r s z indeed coincide with + the correct cell values when executing the register machine. This statement is proven by + induction using the single step relations for Req and Seq as well as the statement for Zeq. \ + +lemma rzs_eq: + "l j\m \ t\q \ R ic p l t = Req l t \ Z ic p l t = Zeq l t \ S ic p j t = Seq j t" +proof (induction t arbitrary: j l) + have "m>0" using m_def is_val is_valid_initial_def[of "ic" "p"] is_valid_def[of "ic" "p"] by auto + + case 0 + (* STATES *) + have mod_aux0: "Suc (b * k) mod b = 1" for k + using euclidean_semiring_cancel_class.mod_mult_self2[of "1" "b" "k"] b_gt1 by auto + have step_state0: "s 0 = 1 + b*\S+ p 0 (\k. s k) + b*\S- p 0 (\k. s k && z (modifies (p!k))) + + b*\S0 p 0 (\k. s k && (e - z (modifies (p!k))))" + using s_eq state_equations_def by auto + hence "Seq 0 0 = 1" using Seq_def by (auto simp: nth_digit_def mod_aux0) + hence S00: "Seq 0 0 = S ic p 0 0" using S_def is_val is_valid_initial_def[of "ic"] by auto + + have "s m = b^q" using s_eq state_equations_def by auto + hence "Seq m 0 = 0" using Seq_def nth_digit_def c_eq rm_constants_def by auto + hence Sm0: "S ic p m 0 = Seq m 0" + using is_val is_valid_initial_def[of "ic" "p" "a"] S_def `m>0` by auto + + have step_states: "\d>0. d s d = b*\S+ p d (\k. s k) + + b*\S- p d (\k. s k && z (modifies (p!k))) + + b*\S0 p d (\k. s k && (e - z (modifies (p!k))))" + using s_eq state_equations_def by auto + hence "\k>0. k Seq k 0 = 0" using Seq_def by (auto simp: nth_digit_def) + hence "k>0 \ k Seq k 0 = S ic p k 0" for k + using S_def is_val is_valid_initial_def[of "ic"] by auto + + with S00 Sm0 have Sid: "k\m \ Seq k 0 = S ic p k 0" for k + by (cases "k=0"; cases "k=m"; auto) + + (* REGISTERS *) + have "b * (r 0 + \R+ p 0 s - \R- p 0 (\k. s k && z 0)) + = b * (r 0 + \R+ p 0 s) - b * \R- p 0 (\k. s k && z 0)" + using Nat.diff_mult_distrib2[of "b" "r 0 + \R+ p 0 s" "\R- p 0 (\k. s k && z 0)"] by auto + also have "... = b * r 0 + b* \R+ p 0 s - b * \R- p 0 (\k. s k && z 0)" + using Nat.add_mult_distrib2[of "b" "r 0" "\R+ p 0 s"] by auto + ultimately have distrib: "a + b * (r 0 + \R+ p 0 s - \R- p 0 (\k. s k && z 0)) + = a + b * r 0 + b * \R+ p 0 s - b * \R- p 0 (\k. s k && z 0)" + by (auto simp: algebra_simps) + (metis Nat.add_diff_assoc add_mult_distrib2 mult_le_mono2 n_gt0 sum_rsub_bound) + + hence "Req 0 0 = (a + b * r 0 + b * \R+ p 0 s - b * \R- p 0 (\k. s k && z 0)) mod b" + using Req_def nth_digit_def r_eq reg_equations_def by auto + also have "... = (a + b * (r 0 + \R+ p 0 s - \R- p 0 (\k. s k && z 0))) mod b" + using distrib by auto + finally have "Req 0 0 = a" using c_eq rm_constants_def B_def by auto + hence R00: "R ic p 0 0 = Req 0 0" + using R_def is_val is_valid_initial_def[of "ic" "p" "a"] by auto + + have rl_transform: "l>0 \ r l = b * r l + b * \R+ p l s - b * \R- p l (\k. s k && z l)" + using reg_equations_def r_eq `l0 \ (b * r l + b * \R+ p l s - b * \R- p l (\k. s k && z l)) mod b = 0" + using Req_def nth_digit_def reg_equations_def r_eq by auto + hence "l>0 \ Req l 0 = 0" + using Req_def rl_transform nth_digit_def by auto + hence "l>0 \ Req l 0 = R ic p l 0" using is_val is_valid_initial_def[of "ic"] R_def by auto + hence Rid: "R ic p l 0 = Req l 0" using R00 by (cases "l=0"; auto) + + (* ZERO INDICATORS *) + hence Zid: "Z ic p l 0 = Zeq l 0" using Z Z_def 0 by auto + + show ?case using Sid Rid Zid `lm` by auto +next + case (Suc t) + + (* INDUCTIVE HYPOTHESIS *) + have "Suc t\q" using Suc by auto + then have "tm \ S ic p k t = Seq k t" for k using Suc m_def by auto + have Z_IH: "\l::nat. l Z ic p l t = Zeq l t" using Suc by auto + + (* REGISTERS *) + from S_IH have S1: "k\m \ + (if isadd (p ! k) \ l = modifies (p ! k) then Seq k t else 0) + = (if isadd (p ! k) \ l = modifies (p ! k) then S ic p k t else 0)" for k + by auto + have S2: "k \ {0..length p-1} \ + (if issub (p ! k) \ l = modifies (p ! k) then Zeq l t * Seq k t else 0) + = (if issub (p ! k) \ l = modifies (p ! k) then Zeq l t * S ic p k t else 0)" for k + using Suc m_def by auto + + have "Req l (Suc t) = Req l t + (\R+ p l (\k. Seq k t)) - (\R- p l (\k. (Zeq l t) * Seq k t))" + using mult_to_single_reg[of "l"] `lR+ p l (\k. S ic p k t)) + - (\R- p l (\k. (Z ic p l t) * S ic p k t))" + using Suc sum_radd.simps sum_rsub.simps S1 S2 m_def by auto + finally have R: "Req l (Suc t) = R ic p l (Suc t)" + using is_val `l Suc m" by (simp add: m_def) + + (* STATES *) + have "s m = b ^ q" using s_eq state_equations_def by auto + hence "Seq m t = 0" using Seq_def \t nth_digit_def apply auto + using b_gt1 bx_aux by auto + hence "S ic p m t = 0" using Suc by auto + hence "fst (steps ic p t) \ m" using S_def by auto + hence "fst (steps ic p t) < m" using is_val m_def + by (metis less_Suc_eq less_le_trans p_contains plength) + hence nohalt: "\ ishalt (p ! fst (steps ic p t))" using is_valid_def[of "ic" "p"] + is_valid_initial_def[of "ic" "p" "a"] m_def is_val by auto + + have "jm` m_def + by (metis (full_types) diff_less is_val length_greater_0_conv less_imp_diff_less less_one + list.size(3) nat_less_le not_less not_less_zero p_contains) + have "Seq j (Suc t) = (\S+ p j (\k. Seq k t)) + + (\S- p j (\k. Zeq (modifies (p!k)) t * Seq k t)) + + (\S0 p j (\k. (1 - Zeq (modifies (p!k)) t) * Seq k t))" + using mult_to_single_state `j\m` `tS+ p j (\k. Seq k t)) + + (\S- p j (\k. Z ic p (modifies (p!k)) t * Seq k t)) + + (\S0 p j (\k. (1 - Z ic p (modifies (p!k)) t) * Seq k t))" + using Z_IH modifies_valid sum_ssub_zero.simps sum_ssub_nzero.simps + by (auto simp: m_def, smt atLeastAtMost_iff sum.cong) + also have "... = (\S+ p j (\k. S ic p k t)) + + (\S- p j (\k. Z ic p (modifies (p!k)) t * S ic p k t)) + + (\S0 p j (\k. (1 - Z ic p (modifies (p!k)) t) * S ic p k t))" + using S_IH sum_ssub_zero.simps sum_ssub_nzero.simps sum_sadd.simps + by (auto simp: m_def, smt atLeastAtMost_iff sum.cong) + finally have S: "Seq j (Suc t) = S ic p j (Suc t)" + using is_val lm04_07_one_step_relation_state[of "ic" "p" "a" "j"] `jSimple Properties of Register Machines\ + +theory RegisterMachineProperties + imports "RegisterMachineSpecification" +begin + +lemma step_commutative: "steps (step c p) p t = step (steps c p t) p" + by (induction t; auto) + +lemma step_fetch_correct: + fixes t :: nat + and c :: configuration + and p :: program + assumes "is_valid c p" + defines "ct \ (steps c p t)" + shows "fst (steps (step c p) p t) = fetch (fst ct) p (read (snd ct) p (fst ct))" + using ct_def step_commutative step_def by auto + +subsubsection \From Configurations to a Protocol\ + +text \Register Values\ + +definition R :: "configuration \ program \ nat \ nat \ nat" + where "R c p n t = (snd (steps c p t)) ! n" + +fun RL :: "configuration \ program \ nat \ nat \ nat \ nat" where + "RL c p b 0 l = ((snd c) ! l)" | + "RL c p b (Suc t) l = ((snd c) ! l) + b * (RL (step c p) p b t l)" + +lemma RL_simp_aux: + \snd c ! l + b * RL (step c p) p b t l = + RL c p b t l + b * (b ^ t * snd (step (steps c p t) p) ! l)\ + by (induction t arbitrary: c) + (auto simp: step_commutative algebra_simps) + +declare RL.simps[simp del] +lemma RL_simp: + "RL c p b (Suc t) l = (snd (steps c p (Suc t)) ! l) * b ^ (Suc t) + (RL c p b t l)" +proof (induction t arbitrary: p c b) + case 0 + thus ?case by (auto simp: RL.simps) +next + case (Suc t p c b) + show ?case + by (subst RL.simps) (* \unfold one level\ *) + (auto simp: Suc step_commutative algebra_simps RL_simp_aux) +qed + +text \State Values\ + +definition S :: "configuration \ program \ nat \ nat \ nat" + where "S c p k t = (if (fst (steps c p t) = k) then (Suc 0) else 0)" + +definition S2 :: "configuration \ nat \ nat" + where "S2 c k = (if (fst c) = k then 1 else 0)" + +fun SK :: "configuration \ program \ nat \ nat \ nat \ nat" + where "SK c p b 0 k = (S2 c k)" | + "SK c p b (Suc t) k = (S2 c k) + b * (SK (step c p) p b t k)" + +lemma SK_simp_aux: + \SK c p b (Suc (Suc t)) k = + S2 (steps c p (Suc (Suc t))) k * b ^ Suc (Suc t) + SK c p b (Suc t) k\ + by (induction t arbitrary: c) (auto simp: step_commutative algebra_simps) + +declare SK.simps[simp del] +lemma SK_simp: + "SK c p b (Suc t) k = (S2 (steps c p (Suc t)) k) * b ^ (Suc t) + (SK c p b t k)" +proof (induction t arbitrary: p c b k) + case 0 + thus ?case by (auto simp: SK.simps) +next + case (Suc t p c b k) + show ?case + by (auto simp: Suc algebra_simps step_commutative SK_simp_aux) +qed + +text \Zero-Indicator Values\ + +definition Z :: "configuration \ program \ nat \ nat \ nat" where + "Z c p n t = (if (R c p n t > 0) then 1 else 0)" + +definition Z2 :: "configuration \ nat \ nat" where + "Z2 c n = (if (snd c) ! n > 0 then 1 else 0)" + +fun ZL :: "configuration \ program \ nat \ nat \ nat \ nat" + where "ZL c p b 0 l = (Z2 c l)" | + "ZL c p b (Suc t) l = (Z2 c l) + b * (ZL (step c p) p b t l)" + +lemma ZL_simp_aux: +"Z2 c l + b * ZL (step c p) p b t l = + ZL c p b t l + b * (b ^ t * Z2 (step (steps c p t) p) l)" + by (induction t arbitrary: c) (auto simp: step_commutative algebra_simps) + +declare ZL.simps[simp del] +lemma ZL_simp: + "ZL c p b (Suc t) l = (Z2 (steps c p (Suc t)) l) * b ^ (Suc t) + (ZL c p b t l)" +proof (induction t arbitrary: p c b) + case 0 + thus ?case by (auto simp: ZL.simps) +next + case (Suc t p c b) + show ?case + by (subst ZL.simps) (auto simp: Suc step_commutative algebra_simps ZL_simp_aux) +qed + +subsubsection \Protocol Properties\ + +lemma Z_bounded: "Z c p l t \ 1" + by (auto simp: Z_def) + +lemma S_bounded: "S c p k t \ 1" + by (auto simp: S_def) + +lemma S_unique: "\k\length p. (k \ fst (steps c p t) \ S c p k t = 0)" + by (auto simp: S_def) + + +(* takes c :: nat, the exponent defining the base b *) +fun cells_bounded :: "configuration \ program \ nat \ bool" where + "cells_bounded conf p c = ((\l<(length (snd conf)). \t. 2^c > R conf p l t) + \ (\k t. 2^c > S conf p k t) + \ (\l t. 2^c > Z conf p l t))" + +lemma steps_tape_length_invar: "length (snd (steps c p t)) = length (snd c)" + by (induction t; auto simp add: step_def update_def) + +lemma step_is_valid_invar: "is_valid c p \ is_valid (step c p) p" + by (auto simp add: step_def update_def is_valid_def) + +fun fetch_old + where + "(fetch_old p s (Add r next) _) = next" + | "(fetch_old p s (Sub r next nextalt) val) = (if val = 0 then nextalt else next)" + | "(fetch_old p s Halt _) = s" + +lemma fetch_equiv: + assumes "i = p!s" + shows "fetch s p v = fetch_old p s i v" + by (cases i; auto simp: assms fetch_def) + +(* Corollary: All states have instructions in the program list *) +lemma p_contains: "is_valid_initial ic p a \ (fst (steps ic p t)) < length p" +proof - + assume asm: "is_valid_initial ic p a" + hence "fst ic = 0" using is_valid_initial_def is_valid_def by blast + hence 0: "ic = (0, snd ic)" by (metis prod.collapse) + show ?thesis using 0 asm + apply (induct t) apply auto[1] + subgoal by (auto simp add: is_valid_initial_def is_valid_def) + apply (cases "p ! fst (steps ic p t)") + apply (auto simp add: list_all_length fetch_equiv step_def + is_valid_initial_def is_valid_def fetch_old.elims) + by (metis RegisterMachineSpecification.isc_add RegisterMachineSpecification.isc_sub + fetch_old.elims) + +qed + +lemma steps_is_valid_invar: "is_valid c p \ is_valid (steps c p t) p" + by (induction t; auto simp add: step_def update_def is_valid_def) + +lemma terminates_halt_state: "terminates ic p q \ is_valid_initial ic p a + \ ishalt (p ! (fst (steps ic p q)))" +proof - + assume terminate: "terminates ic p q" + assume is_val: "is_valid_initial ic p a" + have "1 < length p" using is_val is_valid_initial_def[of "ic" "p" "a"] + is_valid_def[of "ic" "p"] program_includes_halt.simps + by blast + hence "p \ []" by auto + hence "p ! (length p - 1) = last p" using List.last_conv_nth[of "p"] by auto + thus ?thesis + using terminate terminates_def correct_halt_def is_val is_valid_def[of "ic" "p"] by auto +qed + +lemma R_termination: + fixes l :: register and ic :: configuration + assumes is_val: "is_valid ic p" and terminate: "terminates ic p q" and l: "l < length (snd ic)" + shows "\t\q. R ic p l t = 0" +proof - + have ishalt: "ishalt (p ! fst (steps ic p q))" + using terminate terminates_def correct_halt_def is_valid_def is_val by auto + have halt: "ishalt (p ! fst (steps ic p (q + t)))" for t + apply (induction t) + using terminate terminates_def ishalt step_def fetch_def by auto + have "l<(length (snd ic)) \R ic p l (q+t) = 0" for t + apply (induction t arbitrary: l) + subgoal using terminate terminates_def correct_halt_def R_def by auto + subgoal using R_def step_def halt update_def by auto + done + thus ?thesis using le_Suc_ex l by force +qed + +lemma terminate_c_exists: "is_valid ic p \ terminates ic p q \ \c>1. cells_bounded ic p c" +proof - + assume is_val: "is_valid ic p" + assume terminate: "terminates ic p q" + define n where "n \ length (snd ic)" + define rmax where "rmax \ Max ({k. \lt {2})" + have "\lt {k. \lttl rmax" using rmax_def by auto + moreover have "\t\q. \l rmax" + using rmax_def R_termination terminate n_def is_val by auto + ultimately have r: "\lt. R ic p l t \ rmax" using not_le_imp_less by blast + have gt2: "rmax \ 2" using rmax_def by auto + hence sz: "(\k t. rmax > S ic p k t) \ (\l t. rmax > Z ic p l t)" + using S_bounded Z_bounded S_def Z_def by auto + have "(\lt. R ic p l t < 2^rmax) \ (\k t. S ic p k t < 2^rmax) + \ (\l t. Z ic p l t < 2^rmax)" + using less_exp[of "rmax"] r sz by (metis le_neq_implies_less dual_order.strict_trans) + moreover have "rmax > 1" using gt2 by auto + ultimately show ?thesis using n_def by auto +qed + +end diff --git a/thys/DPRM_Theorem/Register_Machine/RegisterMachineSimulation.thy b/thys/DPRM_Theorem/Register_Machine/RegisterMachineSimulation.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Register_Machine/RegisterMachineSimulation.thy @@ -0,0 +1,333 @@ +subsection \Simulation of a Register Machine\ + +theory RegisterMachineSimulation + imports RegisterMachineProperties "Digit_Expansions.Binary_Operations" +begin + +definition B :: "nat \ nat" where + "(B c) = 2^(Suc c)" + +definition "RLe c p b q l = (\t = 0..q. b^t * R c p l t)" +definition "SKe c p b q k = (\t = 0..q. b^t * S c p k t)" +definition "ZLe c p b q l = (\t = 0..q. b^t * Z c p l t)" + +fun sum_radd :: "program \ register \ (nat \ nat) \ nat" + where "sum_radd p l f = (\k = 0..length p-1. if isadd (p!k) \ l = modifies (p!k) then f k else 0)" + +abbreviation sum_radd_abbrev ("\R+ _ _ _" [999, 999, 999] 1000) + where "(\R+ p l f) \ (sum_radd p l f)" + +fun sum_rsub :: "program \ register \ (nat \ nat) \ nat" + where "sum_rsub p l f = (\k = 0..length p-1. if issub (p!k) \ l = modifies (p!k) then f k else 0)" + +abbreviation sum_rsub_abbrev ("\R- _ _ _ " [999, 999, 999] 1000) + where "(\R- p l f) \ (sum_rsub p l f)" + +(* Note: different naming convention for sums compared to original paper *) +fun sum_sadd :: "program \ state \ (nat \ nat) \ nat" + where "sum_sadd p d f = (\k = 0..length p-1. if isadd (p!k) \ d = goes_to (p!k) then f k else 0)" + +abbreviation sum_sadd_abbrev ("\S+ _ _ _ " [999, 999, 999] 1000) + where "(\S+ p d f) \ (sum_sadd p d f)" + +(* careful: f needs be passed the program so that z l t can be properly called with l = modifies (p!k) *) +fun sum_ssub_nzero :: "program \ state \ (nat \ nat) \ nat" + where "sum_ssub_nzero p d f = (\k = 0..length p-1. if issub (p!k) \ d = goes_to (p!k) then f k else 0)" + +abbreviation sum_ssub_nzero_abbrev ("\S- _ _ _ " [999, 999, 999] 1000) + where "(\S- p d f) \ (sum_ssub_nzero p d f)" + +fun sum_ssub_zero :: "program \ state \ (nat \ nat) \ nat" + where "sum_ssub_zero p d f = (\k = 0..length p-1. if issub (p!k) \ d = goes_to_alt (p!k) then f k else 0)" + +abbreviation sum_ssub_zero_abbrev ("\S0 _ _ _ " [999, 999, 999] 1000) + where "(\S0 p d f) \ (sum_ssub_zero p d f)" + +declare sum_radd.simps[simp del] +declare sum_rsub.simps[simp del] +declare sum_sadd.simps[simp del] +declare sum_ssub_nzero.simps[simp del] +declare sum_ssub_zero.simps[simp del] + +text \Special sum cong lemmas\ + +lemma sum_sadd_cong: + assumes "\k. k \ length p-1 \ isadd (p!k) \ l = goes_to (p!k) \ f k = g k" + shows "\S+ p l f = \S+ p l g" + unfolding sum_sadd.simps + by (rule sum.cong, simp) (rule if_cong, simp_all add: assms) + +lemma sum_ssub_nzero_cong: + assumes "\k. k \ length p - 1 \ issub (p!k) \ l = goes_to (p!k) \ f k = g k" + shows "\S- p l f = \S- p l g" + unfolding sum_ssub_nzero.simps + by (rule sum.cong, simp) (rule if_cong, simp_all add: assms) + +lemma sum_ssub_zero_cong: + assumes "\k. k \ length p - 1 \ issub (p!k) \ l = goes_to_alt (p!k) \ f k = g k" + shows "\S0 p l f = \S0 p l g" + unfolding sum_ssub_zero.simps + by (rule sum.cong, simp) (rule if_cong, simp_all add: assms) + +lemma sum_radd_cong: + assumes "\k. k \ length p - 1 \ isadd (p!k) \ l = modifies (p!k) \ f k = g k" + shows "\R+ p l f = \R+ p l g" + unfolding sum_radd.simps + by (rule sum.cong, simp) (rule if_cong, simp_all add: assms) + +lemma sum_rsub_cong: + assumes "\k. k \ length p - 1 \ issub (p!k) \ l = modifies (p!k) \ f k = g k" + shows "\R- p l f = \R- p l g" + unfolding sum_rsub.simps + by (rule sum.cong, simp) (rule if_cong, simp_all add: assms) + + + +text \Properties and simple lemmas\ +lemma RLe_equivalent: "RL c p b q l = RLe c p b q l" + by (induction q arbitrary: c) (auto simp add: RLe_def R_def RL.simps(1) RL_simp) + +lemma SKe_equivalent: "SK c p b q k = SKe c p b q k" + by (induction q arbitrary: c) (auto simp add: SKe_def S_def SK.simps(1) S2_def SK_simp) + +lemma ZLe_equivalent: "ZL c p b q l = ZLe c p b q l" + by (induction q arbitrary: c) (auto simp add: ZLe_def ZL.simps(1) R_def Z2_def Z_def ZL_simp) + + +lemma sum_radd_distrib: "a * (\R+ p l f) = (\R+ p l (\k. a * f k))" + by (auto simp add: sum_radd.simps sum_distrib_left; smt mult_is_0 sum.cong) + +lemma sum_rsub_distrib: "a * (\R- p l f) = (\R- p l (\k. a * f k))" + by (auto simp add: sum_rsub.simps sum_distrib_left; smt mult_is_0 sum.cong) + +lemma sum_sadd_distrib: "a * (\S+ p d f) = (\S+ p d (\k. a * f k))" for a + by (auto simp add: sum_sadd.simps sum_distrib_left; smt mult_is_0 sum.cong) + +lemma sum_ssub_nzero_distrib: "a * (\S- p d f) = (\S- p d (\k. a * f k))" for a + by (auto simp add: sum_ssub_nzero.simps sum_distrib_left; smt mult_is_0 sum.cong) + +lemma sum_ssub_zero_distrib: "a * (\S0 p d f) = (\S0 p d (\k. a * f k))" for a + by (auto simp add: sum_ssub_zero.simps sum_distrib_left; smt mult_is_0 sum.cong) + +lemma sum_distrib: + fixes SX :: "program \ nat \ (nat \ nat) \ nat" + and p :: program + +assumes SX_simps: "\h. SX p x h = (\k = 0..length p-1. if g x k then h k else 0)" + +shows "SX p x h1 + SX p x h2 = SX p x (\k. h1 k + h2 k)" + by (subst SX_simps)+ (auto simp: sum.distrib[symmetric] intro: sum.cong) + +lemma sum_commutative: + fixes SX :: "program \ nat \ (nat \ nat) \ nat" + and p :: program + +assumes SX_simps: "\h. SX p x h = (\k = 0..length p-1. if g x k then h k else 0)" + + shows "(\t=0..q::nat. SX p x (\k. f k t)) + = (SX p x (\k. \t=0..q. f k t))" +proof (induction q) + case 0 + then show ?case by (auto) +next + case (Suc q) + have SX_add: "SX p x h1 + SX p x h2 = SX p x (\k. h1 k + h2 k)" for h1 h2 + by (subst sum_distrib[where ?h1.0 = "h1"]) (auto simp: SX_simps) + + + have h1: "(\t\(Suc q). SX p x (\k. f k t)) = SX p x (\k. f k (Suc q)) + sum (\t. SX p x (\k. f k t)) {0..q}" + by (auto simp add: sum.atLeast0_atMost_Suc add.commute atMost_atLeast0) + also have h2: "... = SX p x (\k. f k (Suc q)) + SX p x (\k. sum (f k){0..q})" + using Suc.IH Suc.prems by auto + also have h3: "... = SX p x (\k. sum (f k) {0..(Suc q)})" + by (subst SX_add) (auto simp: atLeast0_atMost_Suc) + finally show ?case using Suc.IH by (simp add: atMost_atLeast0) +qed + +lemma sum_radd_commutative: "(\t=0..(q::nat). \R+ p l (\k. f k t)) = (\R+ p l (\k. \t=0..q. f k t))" + by (rule sum_commutative sum_radd.simps) + +lemma sum_rsub_commutative: "(\t=0..(q::nat). \R- p l (\k. f k t)) = (\R- p l (\k. \t=0..q. f k t))" + by (rule sum_commutative sum_rsub.simps) + +lemma sum_sadd_commutative: "(\t=0..(q::nat). \S+ p l (\k. f k t)) = (\S+ p l (\k. \t=0..q. f k t))" + by (rule sum_commutative sum_sadd.simps) + +lemma sum_ssub_nzero_commutative: "(\t=0..(q::nat). \S- p l (\k. f k t)) = (\S- p l (\k. \t=0..q. f k t))" + by (rule sum_commutative sum_ssub_nzero.simps) + +lemma sum_ssub_zero_commutative: "(\t=0..(q::nat). \S0 p l (\k. f k t)) = (\S0 p l (\k. \t=0..q. f k t))" + by (rule sum_commutative sum_ssub_zero.simps) + + +lemma sum_int: "c \ a + b \ int(a + b - c) = int(a) + int(b) - int(c)" + by (simp add: SMT.int_plus) + +lemma ZLe_bounded: "b > 2 \ ZLe c p b q l < b ^ (Suc q)" + using Z_bounded ZLe_def +proof (induction q) + case 0 + then show ?case by (simp add: Z_bounded ZLe_def Z_def) +next + case (Suc q) + have "ZLe c p b (Suc q) l = b ^ (Suc q) * Z c p l (Suc q) + ZLe c p b q l" + by (auto simp: ZLe_def) + also have "ZLe c p b q l < b ^ (Suc q)" using Suc.IH + by (auto simp: ZLe_def Z_def Suc.prems(1)) + also have "b ^ (Suc q) * Z c p l (Suc q) \ b ^ (Suc q)" using Suc.prems(1) + by (auto simp: Z_def) + finally have "ZLe c p b (Suc q) l < 2 * b ^ (Suc q)" + by auto + also have "... < b ^ Suc (Suc q)" + using Suc.prems(1) by auto + finally show ?case by simp +qed + +lemma SKe_bounded: "b > 2 \ SKe c p b q k < b ^ (Suc q)" + proof (induction q) + case 0 + then show ?case by (auto simp add: SKe_def S_bounded S_def) +next + case (Suc q) + have "SKe c p b (Suc q) k = b ^ (Suc q) * S c p k (Suc q) + SKe c p b q k" + by (auto simp: SKe_def) + also have "SKe c p b q k < b ^ (Suc q)" using Suc.IH + by (auto simp: Suc.prems(1)) + also have "b ^ (Suc q) * S c p k (Suc q) \ b ^ (Suc q)" using Suc.prems(1) + by (auto simp: S_def) + finally have "SKe c p b (Suc q) k < 2 * b ^ (Suc q)" + by auto + also have "... < b ^ Suc (Suc q)" + using Suc.prems(1) by auto + finally show ?case by simp +qed + +lemma mult_to_bitAND: + assumes cells_bounded: "cells_bounded ic p c" +and "c > 1" +and "b = B c" + +shows "(\t=0..q. b^t * (Z ic p l t * S ic p k t)) + = ZLe ic p b q l && SKe ic p b q k" +proof (induction q arbitrary: ic p c l k) + case 0 + then show ?case using S_bounded Z_bounded + by (auto simp add: SKe_def ZLe_def bitAND_single_bit_mult_equiv) +next + case (Suc q) + + have b4: "b > 2" using assms(2-3) apply (auto simp add: B_def) + by (metis One_nat_def Suc_less_eq2 lessI numeral_2_eq_2 power_gt1) + + have ske: "SKe ic p b q k < b ^ (Suc q)" using SKe_bounded b4 by auto + have zle: "ZLe ic p b q l < b ^ (Suc q)" using ZLe_bounded b4 by auto + + have ih: "(\t = 0..q. b ^ t * (Z ic p l t * S ic p k t)) = ZLe ic p b q l && SKe ic p b q k" + using Suc.IH by auto + + have "(\t = 0..Suc q. b ^ t * (Z ic p l t * S ic p k t)) + = b ^(Suc q) * (Z ic p l (Suc q) * S ic p k (Suc q)) + (\t = 0..q. b ^ t * (Z ic p l t * S ic p k t))" + by (auto simp: sum.atLeast0_atMost_Suc add.commute) + + also have "... = b ^ (Suc q) * (Z ic p l (Suc q) * S ic p k (Suc q)) + (ZLe ic p b q l && SKe ic p b q k)" + by (auto simp add: ih) + + also have "... = b ^ (Suc q) * (Z ic p l (Suc q) && S ic p k (Suc q)) + (ZLe ic p b q l && SKe ic p b q k)" + using bitAND_single_bit_mult_equiv S_bounded Z_bounded by (auto) + + also have "... = (b ^(Suc q) * Z ic p l (Suc q) + ZLe ic p b q l) && (b ^ (Suc q) * S ic p k (Suc q) + SKe ic p b q k)" + using bitAND_linear ske zle + by (auto) (smt B_def assms(3) bitAND_linear mult.commute power_Suc power_mult) + + also have "... = (ZLe ic p b (Suc q) l && SKe ic p b (Suc q) k)" + by (auto simp: ZLe_def SKe_def add.commute) + + finally show ?case by simp +qed + +lemma sum_bt: + fixes b q :: nat + assumes "b > 2" + shows "(\t = 0..q. b^t) < b ^ (Suc q)" + using assms +proof (induction q, auto) + fix qb :: nat + assume "sum ((^) b) {0..qb} < b * b ^ qb" +then have f1: "sum ((^) b) {0..qb} < b ^ Suc qb" + by fastforce +have "b ^ Suc qb * 2 < b ^ Suc (Suc qb)" + using assms by force +then have "2 * b ^ Suc qb < b ^ Suc (Suc qb)" + by simp +then have "b ^ Suc qb + sum ((^) b) {0..qb} < b ^ Suc (Suc qb)" + using f1 by linarith +then show "sum ((^) b) {0..qb} + b * b ^ qb < b * (b * b ^ qb)" + by simp +qed + +lemma mult_to_bitAND_state: + assumes cells_bounded: "cells_bounded ic p c" +and c: "c > 1" +and b: "b = B c" + +shows "(\t=0..q. b^t * ((1 - Z ic p l t) * S ic p k t)) + = ((\t = 0..q. b^t) - ZLe ic p b q l) && SKe ic p b q k" +proof (induction q arbitrary: ic p c l k) + case 0 + show ?case using Z_def S_def ZLe_def SKe_def by auto +next + case (Suc q) + + have b4: "b > 2" using assms(2-3) apply (auto simp add: B_def) + by (metis One_nat_def Suc_less_eq2 lessI numeral_2_eq_2 power_gt1) + + have ske: "SKe ic p b q k < b ^ (Suc q)" using SKe_bounded b4 by auto + have zle: "ZLe ic p b q l < b ^ (Suc q)" using ZLe_bounded b4 by auto + define cst where "cst \ Suc q" + define e where "e \ \t = 0..Suc q. b^t" + + have "(\t = 0..q. b^t) < b ^ (Suc q)" + using sum_bt b4 by auto + hence zle2: "(\t = 0..q. b^t) - ZLe ic p b q l < b ^ (Suc q)" + using less_imp_diff_less by blast + + have "(\t = 0..x. b^t) - ZLe ic p b x l = (\t=0..x. b^t - b^t * Z ic p l t)" for x + unfolding ZLe_def + using Z_bounded sum_subtractf_nat[where ?f = "(^) b" and ?g = "\t. b ^ t * Z ic p l t"] + by auto + hence aux_sum: "(\t = 0..x. b^t) - ZLe ic p b x l = (\t=0..x. b^t * (1 - Z ic p l t))" for x + using diff_Suc_1 diff_mult_distrib2 by auto + + have aux1: "b ^(Suc q) * (1 - Z ic p l (Suc q)) + (\t=0..q. b^t * (1 - Z ic p l t)) + = (\t = 0..cst. b ^ t * (1 - Z ic p l t))" + by (auto simp: sum.atLeast0_atMost_Suc cst_def) + also have aux2: "... = (\t = 0..cst. b ^ t) - ZLe ic p b cst l" + unfolding e_def ZLe_def using aux_sum[of "cst"] + by (auto simp: ZLe_def) + finally have aux_add_sub: + "(b ^(Suc q) * (1 - Z ic p l (Suc q)) + ((\t = 0..q. b^t) - ZLe ic p b q l)) + = (e - ZLe ic p b (Suc q) l)" + by (auto simp: cst_def e_def aux_sum) + + hence ih: "(\t = 0..q. b ^ t * ((1 - Z ic p l t) * S ic p k t)) + = (\t = 0..q. b^t) - ZLe ic p b q l && SKe ic p b q k" + using Suc[of "ic" "p" "l" "k"] by auto + + have "(\t = 0..Suc q. b ^ t * ((1 - Z ic p l t) * S ic p k t)) + = (\t = 0.. q. b ^ t * ((1 - Z ic p l t) * S ic p k t)) + + b^(Suc q) * ((1 - Z ic p l (Suc q)) * S ic p k (Suc q))" + by (auto cong: sum.cong) + + also have "... = ((\t = 0..q. b^t) - ZLe ic p b q l && SKe ic p b q k) + + b^(Suc q) * ((1 - Z ic p l (Suc q)) * S ic p k (Suc q))" + using ih by auto + + also have "... = ((\t = 0..q. b^t) - ZLe ic p b q l && SKe ic p b q k) + + b^(Suc q) * ((1 - Z ic p l (Suc q)) && S ic p k (Suc q))" + using bitAND_single_bit_mult_equiv by (simp add: S_def) + + also have "... = (b ^(Suc q) * (1 - Z ic p l (Suc q)) + ((\t = 0..q. b^t) - ZLe ic p b q l)) + && (b ^ (Suc q) * S ic p k (Suc q) + SKe ic p b q k)" + using bitAND_linear ske zle2 B_def b + by (smt add_ac(2) mult_ac(2) bitAND_linear power.simps(2) power_mult power_mult_distrib) + also have "... = (e - ZLe ic p b (Suc q) l && SKe ic p b (Suc q) k)" + using SKe_def aux_add_sub by (auto simp: add.commute) + + finally show ?case by (auto simp: e_def) +qed + +end diff --git a/thys/DPRM_Theorem/Register_Machine/RegisterMachineSpecification.thy b/thys/DPRM_Theorem/Register_Machine/RegisterMachineSpecification.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Register_Machine/RegisterMachineSpecification.thy @@ -0,0 +1,124 @@ +section \Register Machines\ + +subsection \Register Machine Specification\ +theory RegisterMachineSpecification + imports Main +begin + +subsubsection \Basic Datatype Definitions\ + +text \The following specification of register machines is inspired by @{cite "mech_turing"} (see + @{cite "mech_turing_afp"} for the corresponding AFP article).\ + +(* Type synonyms for registers (= register indices) the "tape" (sim. to a + * Turing machine) that contains a list of register values. + *) +type_synonym register = nat +type_synonym tape = "register list" + +(* The register machine understands "instructions" that operate on state(-id)s + * and modify register(-id)s. The machine stops at the HALT instruction. + *) +type_synonym state = nat +datatype instruction = + isadd: Add (modifies : register) (goes_to : state) | + issub: Sub (modifies : register) (goes_to : state) (goes_to_alt : state) | + ishalt: Halt +where + "modifies Halt = 0" | + "goes_to_alt (Add _ next) = next" + +(* A program, then, just becomes a list of these instructions *) +type_synonym program = "instruction list" + +(* A configuration of the (runtime of) a machine encodes information about the + * instruction and the state of the registers (i.e., the tape). We describe it + * here as a tuple. + *) +type_synonym configuration = "(state * tape)" + +subsubsection \Essential Functions to operate the Register Machine\ + +(* Given a tape of register values and some instruction(s) the register + * machine first reads the value of the register from the tape (by convention + * assume that the value "read" by the HALT state is zero). The machine then, + * fetches the next instruction from the program, and finally updates the + * tape to reflect changes by the last instruction. + *) + +definition read :: "tape \ program \ state \ nat" + where "read t p s = t ! (modifies (p!s))" + +definition fetch :: "state \ program \ nat \ state" where + "fetch s p v = (if issub (p!s) \ v = 0 then goes_to_alt (p!s) + else if ishalt (p!s) then s + else goes_to (p!s))" + +definition update :: "tape \ instruction \ tape" where + "update t i = (if ishalt i then t + else if isadd i then list_update t (modifies i) (t!(modifies i) + 1) + else list_update t (modifies i) (if t!(modifies i) = 0 then 0 else (t!(modifies i)) - 1) )" + +definition step :: "configuration \ program \ configuration" + where + "(step ic p) = (let nexts = fetch (fst ic) p (read (snd ic) p (fst ic)); + nextt = update (snd ic) (p!(fst ic)) + in (nexts, nextt))" + +fun steps :: "configuration \ program \ nat \ configuration" + where + steps_zero: "(steps c p 0) = c" + | steps_suc: "(steps c p (Suc n)) = (step (steps c p n) p)" + + +subsubsection \Validity Checks and Assumptions\ + +(* check bound for each type of instruction *) +(* take a m representing the upper bound for state number *) +fun instruction_state_check :: "nat \ instruction \ bool" + where isc_halt: "instruction_state_check _ Halt = True" + | isc_add: "instruction_state_check m (Add _ ns) = (ns < m)" + | isc_sub: "instruction_state_check m (Sub _ ns1 ns2) = ((ns1 < m) & (ns2 < m))" + +fun instruction_register_check :: "nat \ instruction \ bool" + where "instruction_register_check _ Halt = True" + | "instruction_register_check n (Add reg _) = (reg < n)" + | "instruction_register_check n (Sub reg _ _) = (reg < n)" + +(* passes function via currying into list_all *) +fun program_state_check :: "program \ bool" + where "program_state_check p = list_all (instruction_state_check (length p)) p" + +fun program_register_check :: "program \ nat \ bool" + where "program_register_check p n = list_all (instruction_register_check n) p" + +fun tape_check_initial :: "tape \ nat \ bool" + where "tape_check_initial t a = (t \ [] \ t!0 = a \ (\l>0. t ! l = 0))" + +fun program_includes_halt :: "program \ bool" + where "program_includes_halt p = (length p > 1 \ ishalt (p ! (length p -1)) \ (\k ishalt (p!k)))" + +text \Is Valid and Terminates\ + +definition is_valid + where "is_valid c p = (program_includes_halt p \ program_state_check p + \ (program_register_check p (length (snd c))))" + +definition is_valid_initial + where "is_valid_initial c p a = ((is_valid c p) + \ (tape_check_initial (snd c) a) + \ (fst c = 0))" + +definition correct_halt + where "correct_halt c p q = (ishalt (p ! (fst (steps c p q))) \ \halting\ + \ (\l<(length (snd c)). snd (steps c p q) ! l = 0))" + +definition terminates :: "configuration \ program \ nat \ bool" + where "terminates c p q = ((q>0) + \ (correct_halt c p q) + \ (\x ishalt (p ! (fst (steps c p x)))))" + +definition initial_config :: "nat \ nat \ configuration" where + "initial_config n a = (0, (a # replicate n 0))" + +end diff --git a/thys/DPRM_Theorem/Register_Machine/SingleStepRegister.thy b/thys/DPRM_Theorem/Register_Machine/SingleStepRegister.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Register_Machine/SingleStepRegister.thy @@ -0,0 +1,143 @@ +subsection \Single step relations\ + +subsubsection \Registers\ + +theory SingleStepRegister + imports RegisterMachineSimulation +begin + +lemma single_step_add: + fixes c :: configuration + and p :: program + and l :: register + and t a :: nat + +defines "cs \ fst (steps c p t)" + +assumes is_val: "is_valid_initial c p a" + and l: "l < length tape" + +shows "(\R+ p l (\k. S c p k t)) + = (if isadd (p!cs) \ l = modifies (p!cs) then 1 else 0)" +proof - + have ic: "c = (0, snd c)" + using is_val by (auto simp add: is_valid_initial_def) (metis prod.collapse) + + have add_if: "(\k = 0..length p-1. if isadd (p ! k) \ modifies (p ! cs) = modifies (p ! k) + then S c p k t else 0) + = (\k = 0..length p-1. if k=cs then + if isadd (p ! k) \ modifies (p ! cs) = modifies (p ! k) then S c p k t else 0 else 0)" + apply (rule sum.cong) + using S_unique cs_def by auto + + have bound: "fst (steps c p t) \ length p - 1" using is_val ic p_contains[of "c" "p" "a" "t"] + by (auto simp add: dual_order.strict_implies_order) + + thus ?thesis using S_unique add_if + apply (auto simp add: sum_radd.simps add_if S_def cs_def) + by (smt S_def sum.cong) +qed + +lemma single_step_sub: + fixes c :: configuration + and p :: program + and l :: register + and t a :: nat + +defines "cs \ fst (steps c p t)" + +assumes is_val: "is_valid_initial c p a" + +shows "(\R- p l (\k. Z c p l t * S c p k t)) + = (if issub (p!cs) \ l = modifies (p!cs) then Z c p l t else 0)" +proof - + have "fst c = 0" using is_val by (auto simp add: is_valid_initial_def) + hence ic: "c = (0, snd c)" by (metis prod.collapse) + + have bound: "cs \ length p - 1" using is_val ic p_contains[of "c" "p" "a" "t"] + by (auto simp add: dual_order.strict_implies_order cs_def) + + have sub_if: "(\k = 0..length p-1. if issub (p ! k) \ modifies (p ! cs) = modifies (p ! k) + then 1 * (if cs = k then (Suc 0) else 0) else 0) + =(\k = 0..length p-1. if k = cs then + (if issub (p ! k) \ modifies (p ! cs) = modifies (p ! k) + then (Suc 0) * (if cs = k then (Suc 0) else 0) + else 0) else 0)" + apply (rule sum.cong) using cs_def by auto + + show ?thesis using bound sub_if + apply (auto simp add: sum_rsub.simps cs_def Z_def S_def R_def) + by (metis One_nat_def cs_def) +qed + +lemma lm04_06_one_step_relation_register_old: + fixes l::register + and ic::configuration + and p::program + + defines "s \ fst ic" + and "tape \ snd ic" + + defines "m \ length p" + and "tape' \ snd (step ic p)" + + assumes is_val: "is_valid ic p" + and l: \l < length tape\ + + shows "(tape'!l) = (tape!l) + (if isadd (p!s) \ l = modifies (p!s) then 1 else 0) + - Z ic p l 0 * (if issub (p!s) \ l = modifies (p!s) then 1 else 0)" +proof - + show ?thesis + using l + apply (cases \p!s\) + apply (auto simp: assms(1-4) step_def update_def) + using nth_digit_0 by (auto simp add: Z_def R_def) +qed + +(* [T] 4.6 *) +lemma lm04_06_one_step_relation_register: + fixes l :: register + and c :: configuration + and p :: program + and t :: nat + and a :: nat + +defines "r \ R c p" +defines "s \ S c p" + +assumes is_val: "is_valid_initial c p a" + and l: "l < length (snd c)" + + shows "r l (Suc t) = r l t + (\R+ p l (\k. s k t)) + - (\R- p l (\k. (Z c p l t) * s k t))" +proof - + define cs where "cs \ fst (steps c p t)" + + have add: "(\R+ p l (\k. s k t)) + = (if isadd (p!cs) \ l = modifies (p!cs) then 1 else 0)" + using single_step_add[of "c" "p" "a" "l" "snd c" "t"] is_val l s_def cs_def by auto + + have sub: "(\R- p l (\k. Z c p l t * s k t)) + = (if issub (p!cs) \ l = modifies (p!cs) then Z c p l t else 0)" + using single_step_sub is_val l s_def cs_def Z_def R_def by auto + + have lhs: "r l (Suc t) = snd (steps c p (Suc t)) ! l" + by (simp add: r_def R_def del: steps.simps) + + have rhs: "r l t = snd (steps c p t) ! l" + by (simp add: r_def R_def del: steps.simps) + + have valid_time: "is_valid (steps c p t) p" using steps_is_valid_invar is_val + by (auto simp add: is_valid_initial_def) + + have l_time: "l < length (snd (steps c p t))" using l steps_tape_length_invar by auto + + from lhs rhs have "r l (Suc t) = r l t + (if isadd (p!cs) \ l = modifies (p!cs) then 1 else 0) + - (if issub (p!cs) \ l = modifies (p!cs) then Z c p l t else 0)" + using l_time valid_time lm04_06_one_step_relation_register_old steps.simps cs_def nth_digit_0 + Z_def R_def by auto + + thus ?thesis using add sub by simp +qed + +end diff --git a/thys/DPRM_Theorem/Register_Machine/SingleStepState.thy b/thys/DPRM_Theorem/Register_Machine/SingleStepState.thy new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/Register_Machine/SingleStepState.thy @@ -0,0 +1,79 @@ +subsubsection \States\ + +theory SingleStepState + imports RegisterMachineSimulation +begin + +lemma lm04_07_one_step_relation_state: + fixes d :: state + and c :: configuration + and p :: program + and t :: nat + and a :: nat + +defines "r \ R c p" +defines "s \ S c p" +defines "z \ Z c p" +defines "cs \ fst (steps c p t)" + +assumes is_val: "is_valid_initial c p a" + and "d < length p" + +shows "s d (Suc t) = (\S+ p d (\k. s k t)) + + (\S- p d (\k. z (modifies (p!k)) t * s k t)) + + (\S0 p d (\k. (1 - z (modifies (p!k)) t) * s k t)) + + (if ishalt (p!cs) \ d = cs then Suc 0 else 0)" +proof - + have ic: "c = (0, snd c)" + using is_val by (auto simp add: is_valid_initial_def) (metis prod.collapse) + have cs_bound: "cs < length p" using ic is_val p_contains[of "c" "p" "a" "t"] cs_def by auto + + have "(\k = 0..length p-1. + if isadd (p ! k) \ goes_to (p ! fst (steps c p t)) = goes_to (p ! k) + then if fst (steps c p t) = k + then Suc 0 else 0 else 0) + =(\k = 0..length p-1. + if fst (steps c p t) = k + then if isadd (p ! k) \ goes_to (p ! fst (steps c p t)) = goes_to (p ! k) + then Suc 0 else 0 else 0)" + apply (rule sum.cong) by auto + hence add: "(\S+ p d (\k. s k t)) = (if isadd (p!cs) \ d = goes_to (p!cs) then Suc 0 else 0)" + apply (auto simp add: sum_sadd.simps s_def S_def cs_def) + using cs_bound cs_def by auto + + have "(\k = 0..length p-1. + if issub (p ! k) \ goes_to (p ! fst (steps c p t)) = goes_to (p ! k) + then z (modifies (p ! k)) t * (if fst (steps c p t) = k then Suc 0 else 0) else 0) + = (\k = 0..length p-1. if k=cs then + if issub (p ! k) \ goes_to (p ! fst (steps c p t)) = goes_to (p ! k) + then z (modifies (p ! k)) t else 0 else 0)" + apply (rule sum.cong) + using z_def Z_def cs_def by auto + hence sub_zero: "(\S- p d (\k. z (modifies (p!k)) t * s k t)) + = (if issub (p!cs) \ d = goes_to (p!cs) then z (modifies (p!cs)) t else 0)" + apply (auto simp add: sum_ssub_nzero.simps s_def S_def cs_def) + using cs_bound cs_def by auto + + have "(\k = 0..length p-1. + if issub (p ! k) \ goes_to_alt (p ! fst (steps c p t)) = goes_to_alt (p ! k) + then (Suc 0 - z (modifies (p ! k)) t) * (if fst (steps c p t) = k then Suc 0 else 0) else 0) + = (\k = 0..length p-1. if k=cs then + if issub (p ! k) \ goes_to_alt (p ! fst (steps c p t)) = goes_to_alt (p ! k) + then (Suc 0 - z (modifies (p ! k)) t) else 0 else 0)" + apply (rule sum.cong) using z_def Z_def cs_def by auto + hence sub_nzero: "(\S0 p d (\k. (1 - z (modifies (p!k)) t) * s k t)) + = (if issub (p!cs) \ d = goes_to_alt (p!cs) then (1 - z (modifies (p!cs)) t) else 0)" + apply (auto simp: sum_ssub_zero.simps s_def S_def cs_def) + using cs_bound cs_def by auto + + have "s d (Suc t) = (if isadd (p!cs) \ d = goes_to (p!cs) then Suc 0 else 0) + + (if issub (p!cs) \ d = goes_to (p!cs) then z (modifies (p!cs)) t else 0) + + (if issub (p!cs) \ d = goes_to_alt (p!cs) then (1 - z (modifies (p!cs)) t) else 0) + + (if ishalt (p!cs) \ d = cs then Suc 0 else 0)" + apply (cases "p!cs") + by (auto simp: s_def S_def step_def fetch_def cs_def z_def Z_def Z_bounded R_def read_def) + + thus ?thesis using add sub_zero sub_nzero by auto +qed + +end diff --git a/thys/DPRM_Theorem/document/root.bib b/thys/DPRM_Theorem/document/root.bib new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/document/root.bib @@ -0,0 +1,123 @@ +@InProceedings{h10lecturenotes, + author = "Yuri Matiyasevich", + title = "On {H}ilbert's Tenth Problem", + publisher = "Pacific Institute for the Mathematical Sciences", + year = "2000", + editor = "Michael Lamoureux", + booktitle = "PIMS Distinguished Chair Lectures", + volume = "1" +} + +@InProceedings{cicm, + author="Bayer, Jonas + and David, Marco + and Pal, Abhik + and Stock, Benedikt", + editor="Kaliszyk, Cezary + and Brady, Edwin + and Kohlhase, Andrea + and Sacerdoti Coen, Claudio", + title="Beginners' Quest to Formalize Mathematics: A Feasibility Study in {I}sabelle", + booktitle="Intelligent Computer Mathematics", + year="2019", + publisher="Springer International Publishing", + address="Cham", + pages="16--27", + abstract="How difficult are interactive theorem provers to use? We respond by reviewing the formalization of Hilbert's tenth problem in Isabelle/HOL carried out by an undergraduate research group at Jacobs University Bremen. We argue that, as demonstrated by our example, proof assistants are feasible for beginners to formalize mathematics. With the aim to make the field more accessible, we also survey hurdles that arise when learning an interactive theorem prover. Broadly, we advocate for an increased adoption of interactive theorem provers in mathematical research and curricula.", + isbn="978-3-030-23250-4" +} + +@InProceedings{dprm_isabelle, + author = {Jonas Bayer and Marco David and Abhik Pal and Benedikt Stock and Dierk Schleicher}, + title = {{The DPRM Theorem in Isabelle (Short Paper)}}, + booktitle = {10th International Conference on Interactive Theorem Proving (ITP 2019)}, + pages = {33:1--33:7}, + series = {Leibniz International Proceedings in Informatics (LIPIcs)}, + ISBN = {978-3-95977-122-1}, + ISSN = {1868-8969}, + year = {2019}, + volume = {141}, + editor = {John Harrison and John O'Leary and Andrew Tolmach}, + publisher = {Schloss Dagstuhl--Leibniz-Zentrum fuer Informatik}, + address = {Dagstuhl, Germany}, + URL = {http://drops.dagstuhl.de/opus/volltexte/2019/11088}, + doi = {10.4230/LIPIcs.ITP.2019.33} +} + +@article{dprm_mizar1, + title = {The {M}atiyasevich Theorem. Preliminaries}, + volume = {25}, + url = {https://www.sciendo.com/article/10.1515/forma-2017-0029}, + doi = {10.1515/forma-2017-0029}, + pages = {315--322}, + number = {4}, + journal = {Formalized Mathematics}, + author = {Pak, Karol}, + urldate = {2021-06-20}, + date = {2018-03-28}, + year = {2018}, +} + +@article{dprm_mizar2, + title = {Diophantine sets. Preliminaries}, + volume = {26}, + url = {https://www.sciendo.com/article/10.2478/forma-2018-0007}, + doi = {10.2478/forma-2018-0007}, + pages = {81--90}, + number = {1}, + journal = {Formalized Mathematics}, + author = {Pak, Karol}, + year = {2018}, + urldate = {2021-06-20}, + date = {2018-07-28}, +} + +@InProceedings{dprm_coq, + author = {Dominique Larchey-Wendling and Yannick Forster}, + title = {{Hilbert's Tenth Problem in Coq}}, + booktitle = {4th International Conference on Formal Structures for Computation and Deduction (FSCD 2019)}, + pages = {27:1--27:20}, + series = {Leibniz International Proceedings in Informatics (LIPIcs)}, + ISBN = {978-3-95977-107-8}, + ISSN = {1868-8969}, + year = {2019}, + volume = {131}, + editor = {Herman Geuvers}, + publisher = {Schloss Dagstuhl--Leibniz-Zentrum fuer Informatik}, + address = {Dagstuhl, Germany}, + URL = {http://drops.dagstuhl.de/opus/volltexte/2019/10534}, + URN = {urn:nbn:de:0030-drops-105342}, + doi = {10.4230/LIPIcs.FSCD.2019.27}, + annote = {Keywords: Hilbert's tenth problem, Diophantine equations, undecidability, computability theory, reduction, Minsky machines, Fractran, Coq, type theory} +} + +@unpublished{dprm_lean, + author = {Carneiro, Mario}, + year = {2018}, + month = {02}, + title = {A {L}ean formalization of {M}atiyasevi\v{c}'s Theorem}, + note="\url{https://arxiv.org/abs/1802.01795v1}", +} + +@InProceedings{mech_turing, + author="Xu, Jian and Zhang, Xingyuan and Urban, Christian", + editor="Blazy, Sandrine and Paulin-Mohring, Christine and Pichardie, David", + title="Mechanising {T}uring Machines and Computability Theory in {I}sabelle/{HOL}", + booktitle="Interactive Theorem Proving. ITP 2013.", + series="Lecture Notes in Computer Science", + volume="7998", + year="2013", + publisher="Springer Berlin, Heidelberg", + pages="147--162" +} + +@article{mech_turing_afp, + author = {Jian Xu and Xingyuan Zhang and Christian Urban and Sebastiaan J. C. Joosten}, + title = {Universal {T}uring Machine}, + journal = {Archive of Formal Proofs}, + month = feb, + year = 2019, + note = {\url{https://isa-afp.org/entries/Universal_Turing_Machine.html}, + Formal proof development}, + ISSN = {2150-914x}, +} \ No newline at end of file diff --git a/thys/DPRM_Theorem/document/root.log b/thys/DPRM_Theorem/document/root.log new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/document/root.log @@ -0,0 +1,47 @@ +This is pdfTeX, Version 3.141592653-2.6-1.40.22 (MiKTeX 21.3) (preloaded format=pdflatex 2021.3.17) 22 DEC 2021 11:27 +entering extended mode +**./root.tex +(root.tex +LaTeX2e <2020-10-01> patch level 4 +L3 programming layer <2021-02-18> xparse <2020-03-03> +(C:\Users\jonas\AppData\Local\Programs\MiKTeX\tex/latex/base\article.cls +Document Class: article 2020/04/10 v1.4m Standard LaTeX document class +(C:\Users\jonas\AppData\Local\Programs\MiKTeX\tex/latex/base\size11.clo +File: size11.clo 2020/04/10 v1.4m Standard LaTeX file (size option) +) +\c@part=\count179 +\c@section=\count180 +\c@subsection=\count181 +\c@subsubsection=\count182 +\c@paragraph=\count183 +\c@subparagraph=\count184 +\c@figure=\count185 +\c@table=\count186 +\abovecaptionskip=\skip47 +\belowcaptionskip=\skip48 +\bibindent=\dimen138 +) + +! LaTeX Error: File `isabelle.sty' not found. + +Type X to quit or to proceed, +or enter new name. (Default extension: sty) + +Enter file name: +! Emergency stop. + + +l.3 + +*** (cannot \read from terminal in nonstop modes) + + +Here is how much of TeX's memory you used: + 220 strings out of 479334 + 2664 string characters out of 2857084 + 285137 words of memory out of 3000000 + 17766 multiletter control sequences out of 15000+200000 + 403730 words of font info for 28 fonts, out of 3000000 for 9000 + 1141 hyphenation exceptions out of 8191 + 44i,0n,50p,156b,36s stack positions out of 5000i,500n,10000p,200000b,50000s +! ==> Fatal error occurred, no output PDF file produced! diff --git a/thys/DPRM_Theorem/document/root.tex b/thys/DPRM_Theorem/document/root.tex new file mode 100644 --- /dev/null +++ b/thys/DPRM_Theorem/document/root.tex @@ -0,0 +1,105 @@ +\documentclass[11pt,a4paper]{article} +\usepackage{isabelle,isabellesym} +\usepackage{authblk} + +% further packages required for unusual symbols (see also +% isabellesym.sty), use only when needed + +%\usepackage{amssymb} + %for \, \, \, \, \, \, + %\, \, \, \, \, + %\, \, \ + +%\usepackage{eurosym} + %for \ + +%\usepackage[only,bigsqcap]{stmaryrd} + %for \ + +%\usepackage{eufrak} + %for \ ... \, \ ... \ (also included in amssymb) + +%\usepackage{textcomp} + %for \, \, \, \, \, + %\ + +% this should be the last package used +\usepackage{pdfsetup} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + +% for uniform font size +%\renewcommand{\isastyle}{\isastyleminor} + + +\begin{document} + +\title{Diophantine Equations and the DPRM Theorem} + +\author{Jonas Bayer\thanks{Equal contribution.} \hspace*{10ex} Marco David\textsuperscript{*} + \hspace*{10ex} Benedikt Stock\textsuperscript{*} \\ Abhik Pal + \hspace*{10ex} Yuri Matiyasevich\thanks{Contributed by supplying a detailed proof and an + initial introduction to Isabelle.} + \hspace*{10ex} Dierk Schleicher} + +\maketitle + +%\footnotetext[*]{Text} + +\begin{abstract} + We present a formalization of Matiyasevich's proof of the DPRM theorem, which states that every + recursively enumerable set of natural numbers is Diophantine. This result from 1970 yields a + negative solution to Hilbert's 10th problem over the integers. To represent recursively + enumerable sets in equations, we implement and arithmetize register machines. We formalize a + general theory of Diophantine sets and relations to reason about them abstractly. + Using several number-theoretic lemmas, we prove that exponentiation has a Diophantine + representation. +\end{abstract} + +\tableofcontents + +\newpage + +% sane default for proof documents +\parindent 0pt\parskip 0.5ex + +% Some preliminary text +\paragraph{Overview} +A previous short paper~\cite{dprm_isabelle} gives an overview +of the formalization. In particular, the challenges of implementing the notion of +diophantine predicates is discussed and a formal definition of register machines is described. +Another meta-publication~\cite{cicm} recounts our learning experience throughout this project. + +The present formalisation is based on Yuri Matiyasevich's monograph~\cite{h10lecturenotes} which +contains a full proof of the DPRM theorem. This result or parts of its proof have also been +formalized in other interactive theorem provers, notably in Coq~\cite{dprm_coq}, +Lean~\cite{dprm_lean} and Mizar~\cite{dprm_mizar1,dprm_mizar2}. + +\paragraph{Acknowledgements} +We want to thank everyone who participated in the formalization +during the early stages of this project: +Deepak Aryal, Bogdan Ciurezu, Yiping Deng, Prabhat Devkota, +Simon Dubischar, Malte Haßler, Yufei Liu and Maria Antonia Oprea. Moreover, we would like to +express our sincere gratitude to the entire welcoming and supportive Isabelle +community. In particular, we are indebted to Christoph Benzmüller for his +expertise, his advice and for connecting us with relevant experts in the field. +Among those, we specially want to thank Mathias Fleury for all his help with Isabelle. +Finally, we would like to thank the DFG for supporting our attendance at several events and conferences, allowing us to present the project to a broad audience. + +\newpage + +% generated text of all theories +\input{session} + +% optional bibliography +\bibliographystyle{abbrv} +\bibliography{root} + +\end{document} + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: t +%%% End: diff --git a/thys/ROOTS b/thys/ROOTS --- a/thys/ROOTS +++ b/thys/ROOTS @@ -1,684 +1,685 @@ ADS_Functor AI_Planning_Languages_Semantics AODV AVL-Trees AWN Abortable_Linearizable_Modules Abs_Int_ITP2012 Abstract-Hoare-Logics Abstract-Rewriting Abstract_Completeness Abstract_Soundness Ackermanns_not_PR Actuarial_Mathematics Adaptive_State_Counting Affine_Arithmetic Aggregation_Algebras Akra_Bazzi Algebraic_Numbers Algebraic_VCs Allen_Calculus Amicable_Numbers Amortized_Complexity AnselmGod Applicative_Lifting Approximation_Algorithms Architectural_Design_Patterns Aristotles_Assertoric_Syllogistic Arith_Prog_Rel_Primes ArrowImpossibilityGS Attack_Trees Auto2_HOL Auto2_Imperative_HOL AutoFocus-Stream Automated_Stateful_Protocol_Verification Automatic_Refinement AxiomaticCategoryTheory BDD BD_Security_Compositional BNF_CC BNF_Operations BTree Banach_Steinhaus Belief_Revision Bell_Numbers_Spivey BenOr_Kozen_Reif Berlekamp_Zassenhaus Bernoulli Bertrands_Postulate Bicategory BinarySearchTree Binding_Syntax_Theory Binomial-Heaps Binomial-Queues BirdKMP Blue_Eyes Bondy Boolean_Expression_Checkers Bounded_Deducibility_Security Buchi_Complementation Budan_Fourier Buffons_Needle Buildings BytecodeLogicJmlTypes C2KA_DistributedSystems CAVA_Automata CAVA_LTL_Modelchecker CCS CISC-Kernel CRDT CSP_RefTK CYK CZH_Elementary_Categories CZH_Foundations CZH_Universal_Constructions CakeML CakeML_Codegen Call_Arity Card_Equiv_Relations Card_Multisets Card_Number_Partitions Card_Partitions Cartan_FP Case_Labeling Catalan_Numbers Category Category2 Category3 Cauchy Cayley_Hamilton Certification_Monads Chandy_Lamport Chord_Segments Circus Clean Clique_and_Monotone_Circuits ClockSynchInst Closest_Pair_Points CoCon CoSMeDis CoSMed CofGroups Coinductive Coinductive_Languages Collections Combinable_Wands Combinatorics_Words Combinatorics_Words_Graph_Lemma Combinatorics_Words_Lyndon Comparison_Sort_Lower_Bound Compiling-Exceptions-Correctly Complete_Non_Orders Completeness Complex_Bounded_Operators Complex_Geometry Complx ComponentDependencies ConcurrentGC ConcurrentIMP Concurrent_Ref_Alg Concurrent_Revisions Conditional_Simplification Conditional_Transfer_Rule Consensus_Refined Constructive_Cryptography Constructive_Cryptography_CM Constructor_Funs Containers CoreC++ Core_DOM Core_SC_DOM Correctness_Algebras Cotangent_PFD_Formula Count_Complex_Roots CryptHOL CryptoBasedCompositionalProperties Cubic_Quartic_Equations DFS_Framework DOM_Components DPT-SAT-Solver DataRefinementIBP Datatype_Order_Generator Decl_Sem_Fun_PL Decreasing-Diagrams Decreasing-Diagrams-II Dedekind_Real Deep_Learning Delta_System_Lemma Density_Compiler Dependent_SIFUM_Refinement Dependent_SIFUM_Type_Systems Depth-First-Search Derangements Deriving Descartes_Sign_Rule Design_Theory Dict_Construction Differential_Dynamic_Logic Differential_Game_Logic Digit_Expansions Dijkstra_Shortest_Path Diophantine_Eqns_Lin_Hom Dirichlet_L Dirichlet_Series DiscretePricing Discrete_Summation DiskPaxos Dominance_CHK +DPRM_Theorem DynamicArchitectures Dynamic_Tables E_Transcendental Echelon_Form EdmondsKarp_Maxflow Efficient-Mergesort Elliptic_Curves_Group_Law Encodability_Process_Calculi Epistemic_Logic Equivalence_Relation_Enumeration Ergodic_Theory Error_Function Euler_MacLaurin Euler_Partition Eval_FO Example-Submission Extended_Finite_State_Machine_Inference Extended_Finite_State_Machines FFT FLP FOL-Fitting FOL_Axiomatic FOL_Harrison FOL_Seq_Calc1 FOL_Seq_Calc2 FOL_Seq_Calc3 Factor_Algebraic_Polynomial Factored_Transition_System_Bounding Falling_Factorial_Sum Farkas FeatherweightJava Featherweight_OCL Fermat3_4 FileRefinement FinFun Finger-Trees Finite-Map-Extras Finite_Automata_HF Finitely_Generated_Abelian_Groups First_Order_Terms First_Welfare_Theorem Fishburn_Impossibility Fisher_Yates Fishers_Inequality Flow_Networks Floyd_Warshall Flyspeck-Tame FocusStreamsCaseStudies Forcing Formal_Puiseux_Series Formal_SSA Formula_Derivatives Foundation_of_geometry Fourier FO_Theory_Rewriting Free-Boolean-Algebra Free-Groups Frequency_Moments Fresh_Identifiers FunWithFunctions FunWithTilings Functional-Automata Functional_Ordered_Resolution_Prover Furstenberg_Topology GPU_Kernel_PL Gabow_SCC GaleStewart_Games Gale_Shapley Game_Based_Crypto Gauss-Jordan-Elim-Fun Gauss_Jordan Gauss_Sums Gaussian_Integers GenClock General-Triangle Generalized_Counting_Sort Generic_Deriving Generic_Join GewirthPGCProof Girth_Chromatic GoedelGod Goedel_HFSet_Semantic Goedel_HFSet_Semanticless Goedel_Incompleteness Goodstein_Lambda GraphMarkingIBP Graph_Saturation Graph_Theory Green Groebner_Bases Groebner_Macaulay Gromov_Hyperbolicity Grothendieck_Schemes Group-Ring-Module HOL-CSP HOLCF-Prelude HRB-Slicing Hahn_Jordan_Decomposition Heard_Of Hello_World HereditarilyFinite Hermite Hermite_Lindemann Hidden_Markov_Models Higher_Order_Terms Hoare_Time Hood_Melville_Queue HotelKeyCards Huffman Hybrid_Logic Hybrid_Multi_Lane_Spatial_Logic Hybrid_Systems_VCs HyperCTL Hyperdual IEEE_Floating_Point IFC_Tracking IMAP-CRDT IMO2019 IMP2 IMP2_Binary_Heap IMP_Compiler IP_Addresses Imperative_Insertion_Sort Impossible_Geometry Incompleteness Incredible_Proof_Machine Independence_CH Inductive_Confidentiality Inductive_Inference InfPathElimination InformationFlowSlicing InformationFlowSlicing_Inter Integration Interpolation_Polynomials_HOL_Algebra Interpreter_Optimizations Interval_Arithmetic_Word32 Intro_Dest_Elim Iptables_Semantics Irrational_Series_Erdos_Straus Irrationality_J_Hancl Irrationals_From_THEBOOK IsaGeoCoq Isabelle_C Isabelle_Marries_Dirac Isabelle_Meta_Model Jacobson_Basic_Algebra Jinja JinjaDCI JinjaThreads JiveDataStoreModel Jordan_Hoelder Jordan_Normal_Form KAD KAT_and_DRA KBPs KD_Tree Key_Agreement_Strong_Adversaries Kleene_Algebra Knights_Tour Knot_Theory Knuth_Bendix_Order Knuth_Morris_Pratt Koenigsberg_Friendship Kruskal Kuratowski_Closure_Complement LLL_Basis_Reduction LLL_Factorization LOFT LTL LTL_Master_Theorem LTL_Normal_Form LTL_to_DRA LTL_to_GBA Lam-ml-Normalization LambdaAuth LambdaMu Lambda_Free_EPO Lambda_Free_KBOs Lambda_Free_RPOs Lambert_W Landau_Symbols Laplace_Transform Latin_Square LatticeProperties Launchbury Laws_of_Large_Numbers Lazy-Lists-II Lazy_Case Lehmer Lifting_Definition_Option Lifting_the_Exponent LightweightJava LinearQuantifierElim Linear_Inequalities Linear_Programming Linear_Recurrences Liouville_Numbers List-Index List-Infinite List_Interleaving List_Inversions List_Update LocalLexing Localization_Ring Locally-Nameless-Sigma Logging_Independent_Anonymity Lowe_Ontological_Argument Lower_Semicontinuous Lp LP_Duality Lucas_Theorem MDP-Algorithms MDP-Rewards MFMC_Countable MFODL_Monitor_Optimized MFOTL_Monitor MSO_Regex_Equivalence Markov_Models Marriage Mason_Stothers Matrices_for_ODEs Matrix Matrix_Tensor Matroids Max-Card-Matching Median_Method Median_Of_Medians_Selection Menger Mereology Mersenne_Primes Metalogic_ProofChecker MiniML MiniSail Minimal_SSA Minkowskis_Theorem Minsky_Machines Modal_Logics_for_NTS Modular_Assembly_Kit_Security Modular_arithmetic_LLL_and_HNF_algorithms Monad_Memo_DP Monad_Normalisation MonoBoolTranAlgebra MonoidalCategory Monomorphic_Monad MuchAdoAboutTwo Multiset_Ordering_NPC Multi_Party_Computation Multirelations Myhill-Nerode Name_Carrying_Type_Inference Nash_Williams Nat-Interval-Logic Native_Word Nested_Multisets_Ordinals Network_Security_Policy_Verification Neumann_Morgenstern_Utility No_FTL_observers Nominal2 Noninterference_CSP Noninterference_Concurrent_Composition Noninterference_Generic_Unwinding Noninterference_Inductive_Unwinding Noninterference_Ipurge_Unwinding Noninterference_Sequential_Composition NormByEval Nullstellensatz Octonions OpSets Open_Induction Optics Optimal_BST Orbit_Stabiliser Order_Lattice_Props Ordered_Resolution_Prover Ordinal Ordinal_Partitions Ordinals_and_Cardinals Ordinary_Differential_Equations PAC_Checker Package_logic PAL PCF PLM POPLmark-deBruijn PSemigroupsConvolution Padic_Ints Pairing_Heap Paraconsistency Parity_Game Partial_Function_MR Partial_Order_Reduction Password_Authentication_Protocol Pell Perfect-Number-Thm Perron_Frobenius Physical_Quantities Pi_Calculus Pi_Transcendental Planarity_Certificates Pluennecke_Ruzsa_Inequality Poincare_Bendixson Poincare_Disc Polynomial_Factorization Polynomial_Interpolation Polynomials Pop_Refinement Posix-Lexing Possibilistic_Noninterference Power_Sum_Polynomials Pratt_Certificate Prefix_Free_Code_Combinators Presburger-Automata Prim_Dijkstra_Simple Prime_Distribution_Elementary Prime_Harmonic_Series Prime_Number_Theorem Priority_Queue_Braun Priority_Search_Trees Probabilistic_Noninterference Probabilistic_Prime_Tests Probabilistic_System_Zoo Probabilistic_Timed_Automata Probabilistic_While Program-Conflict-Analysis Progress_Tracking Projective_Geometry Projective_Measurements Promela Proof_Strategy_Language PropResPI Propositional_Proof_Systems Prpu_Maxflow PseudoHoops Psi_Calculi Ptolemys_Theorem Public_Announcement_Logic QHLProver QR_Decomposition Quantales Quasi_Borel_Spaces Quaternions Quick_Sort_Cost RIPEMD-160-SPARK ROBDD RSAPSS Ramsey-Infinite Random_BSTs Random_Graph_Subgraph_Threshold Randomised_BSTs Randomised_Social_Choice Rank_Nullity_Theorem Real_Impl Real_Power Recursion-Addition Recursion-Theory-I Refine_Imperative_HOL Refine_Monadic RefinementReactive Regex_Equivalence Registers Regression_Test_Selection Regular-Sets Regular_Algebras Regular_Tree_Relations Relation_Algebra Relational-Incorrectness-Logic Relational_Disjoint_Set_Forests Relational_Forests Relational_Method Relational_Minimum_Spanning_Trees Relational_Paths Rep_Fin_Groups ResiduatedTransitionSystem Residuated_Lattices Resolution_FOL Rewrite_Properties_Reduction Rewriting_Z Ribbon_Proofs Robbins-Conjecture Robinson_Arithmetic Root_Balanced_Tree Roth_Arithmetic_Progressions Routing Roy_Floyd_Warshall SATSolverVerification SC_DOM_Components SDS_Impossibility SIFPL SIFUM_Type_Systems SPARCv8 Safe_Distance Safe_OCL Saturation_Framework Saturation_Framework_Extensions Schutz_Spacetime Secondary_Sylow Security_Protocol_Refinement Selection_Heap_Sort SenSocialChoice Separata Separation_Algebra Separation_Logic_Imperative_HOL SequentInvertibility Shadow_DOM Shadow_SC_DOM Shivers-CFA ShortestPath Show Sigma_Commit_Crypto Signature_Groebner Simpl Simple_Firewall Simplex Simplicial_complexes_and_boolean_functions SimplifiedOntologicalArgument Skew_Heap Skip_Lists Slicing Sliding_Window_Algorithm Smith_Normal_Form Smooth_Manifolds Sophomores_Dream Sort_Encodings Source_Coding_Theorem SpecCheck Special_Function_Bounds Splay_Tree Sqrt_Babylonian Stable_Matching Statecharts Stateful_Protocol_Composition_and_Typing Stellar_Quorums Stern_Brocot Stewart_Apollonius Stirling_Formula Stochastic_Matrices Stone_Algebras Stone_Kleene_Relation_Algebras Stone_Relation_Algebras Store_Buffer_Reduction Stream-Fusion Stream_Fusion_Code Strong_Security Sturm_Sequences Sturm_Tarski Stuttering_Equivalence Subresultants Subset_Boolean_Algebras SumSquares Sunflowers SuperCalc Surprise_Paradox Symmetric_Polynomials Syntax_Independent_Logic Szemeredi_Regularity Szpilrajn TESL_Language TLA Tail_Recursive_Functions Tarskis_Geometry Taylor_Models Three_Circles Timed_Automata Topological_Semantics Topology TortoiseHare Transcendence_Series_Hancl_Rucki Transformer_Semantics Transition_Systems_and_Automata Transitive-Closure Transitive-Closure-II Transitive_Models Treaps Tree-Automata Tree_Decomposition Triangle Trie Twelvefold_Way Tycon Types_Tableaus_and_Goedels_God Types_To_Sets_Extension UPF UPF_Firewall UTP Universal_Hash_Families Universal_Turing_Machine UpDown_Scheme Valuation Van_Emde_Boas_Trees Van_der_Waerden VectorSpace VeriComp Verified-Prover Verified_SAT_Based_AI_Planning VerifyThis2018 VerifyThis2019 Vickrey_Clarke_Groves Virtual_Substitution VolpanoSmith VYDRA_MDL WHATandWHERE_Security WOOT_Strong_Eventual_Consistency WebAssembly Weight_Balanced_Trees Weighted_Path_Order Well_Quasi_Orders Wetzels_Problem Winding_Number_Eval Word_Lib WorkerWrapper X86_Semantics XML Youngs_Inequality ZFC_in_HOL Zeta_3_Irrational Zeta_Function pGCL