diff --git a/thys/CVP_Hardness/BHLE.thy b/thys/CVP_Hardness/BHLE.thy --- a/thys/CVP_Hardness/BHLE.thy +++ b/thys/CVP_Hardness/BHLE.thy @@ -1,1706 +1,1705 @@ theory BHLE imports Main Additional_Lemmas begin section \Bounded Homogeneous Linear Equation Problem\ definition bhle :: "(int vec * int) set" where "bhle \ {(a,k). \x. a \ x = 0 \ dim_vec x = dim_vec a \ dim_vec a > 0 \ x \ 0\<^sub>v (dim_vec x) \ \x\\<^sub>\ \ k}" text \Reduction of bounded homogeneous linear equation to partition problem\ text \For the reduction function, one defines five-tuples for every element in a. The last tuple plays an important role. It consists only of four elements in order to add constraints to the other tuples. These values are formed in a way such that they form all constraints needed for the reductions. Note, that some indices have been changed with respect to \cite{EmBo81} to enable better indexing in the vectors/lists.\ definition b1 :: "nat \ int \ int \ int" where "b1 i M a = a + M * (5^(4*i-4) + 5^(4*i-3) + 5^(4*i-1))" definition b2 :: "nat \ int \ int" where "b2 i M = M * (5^(4*i-3) + 5^(4*i))" definition b2_last :: "nat \ int \ int" where "b2_last i M = M * (5^(4*i-3) + 1)" definition b3 :: "nat \ int \ int" where "b3 i M = M * (5^(4*i-4) + 5^(4*i-2))" definition b4 :: "nat \ int \ int \ int" where "b4 i M a = a + M * (5^(4*i-2) + 5^(4*i-1) + 5^(4*i))" definition b4_last :: "nat \ int \ int \ int" where "b4_last i M a = a + M * (5^(4*i-2) + 5^(4*i-1) + 1)" definition b5 :: "nat \ int \ int" where "b5 i M = M * (5^(4*i-1))" text \Change order of indices in \cite{EmBo81} such that $b3$ is in last place and can be omitted in the last entry. This ensures that the weight of the solution is $1$ or $-1$, essential for the proof of NP-hardnes.\ definition b_list :: "int list \ nat \ int \ int list" where "b_list as i M = [b1 (i+1) M (as!i), b2 (i+1) M, b4 (i+1) M (as!i), b5 (i+1) M, b3 (i+1) M]" definition b_list_last :: "int list \ nat \ int \ int list" where "b_list_last as n M = [b1 n M (last as), b2_last n M, b4_last n M (last as), b5 n M]" definition gen_bhle :: "int list \ int vec" where "gen_bhle as = (let M = 2*(\ias!i\)+1; n = length as in vec_of_list (concat (map (\i. b_list as i M) [0..0 then b_list_last as n M else [])))" text \The reduction function.\ definition reduce_bhle_partition:: "(int list) \ ((int vec) * int)" where "reduce_bhle_partition \ (\ a. (gen_bhle a, 1))" text \Lemmas for proof\ lemma dim_vec_gen_bhle: assumes "as\[]" shows "dim_vec (gen_bhle as) = 5 * (length as) - 1" using assms proof (induct as rule: list_nonempty_induct) case (single x) then show ?case unfolding gen_bhle_def Let_def b_list_def b_list_last_def by auto next case (cons x xs) define M where "M = (2 * (\i(x # xs) ! i\) + 1)" then show ?case using cons unfolding gen_bhle_def b_list_def b_list_last_def Let_def M_def[symmetric] by (subst dim_vec_of_list)+ (use length_concat_map_5[of "(\i. b1 (i + 1) M ((x#xs) ! i))" "(\i. b2 (i + 1) M )" "(\i. b4 (i + 1) M ((x#xs) ! i))" "(\i. b5 (i + 1) M )" "(\i. b3 (i + 1) M )"] in \simp\) qed lemma dim_vec_gen_bhle_empty: "dim_vec (gen_bhle []) = 0" unfolding gen_bhle_def by auto text \Lemmas about length\ lemma length_b_list: "length (b_list a i M) = 5" unfolding b_list_def by auto lemma length_b_list_last: "length (b_list_last a n M) = 4" unfolding b_list_last_def by auto lemma length_concat_map_b_list: "length (concat (map (\i. b_list as i M) [0..Last values of \gen_bhle\\ lemma gen_bhle_last0: assumes "length as > 0" shows "(gen_bhle as) $ ((length as -1) * 5) = b1 (length as) (2*(\ias!i\)+1) (last as)" unfolding gen_bhle_def Let_def proof (subst vec_of_list_append, subst index_append_vec(1), goal_cases) case 1 then show ?case using assms by (subst dim_vec_of_list)+ (split if_splits, subst length_b_list_last, subst length_concat_map_b_list, auto) next case 2 then show ?case using assms by (subst dim_vec_of_list, subst length_concat_map_b_list, subst vec_index_vec_of_list)+ (auto split: if_splits simp add: b_list_last_def) qed lemma gen_bhle_last1: assumes "length as > 0" shows "(gen_bhle as) $ ((length as -1) * 5 + 1) = b2_last (length as) (2*(\ias!i\)+1)" unfolding gen_bhle_def Let_def proof (subst vec_of_list_append, subst index_append_vec(1), goal_cases) case 1 then show ?case using assms by (subst dim_vec_of_list)+ (split if_split, subst length_b_list_last, subst length_concat_map_b_list, auto) next case 2 then show ?case using assms by (subst dim_vec_of_list, subst length_concat_map_b_list, subst vec_index_vec_of_list)+ (auto split: if_splits simp add: b_list_last_def) qed text \The third entry of the last tuple is omitted, thus we skip one lemma\ lemma gen_bhle_last3: assumes "length as > 0" shows "(gen_bhle as) $ ((length as -1) * 5 + 2) = b4_last (length as) (2*(\ias!i\)+1) (last as)" unfolding gen_bhle_def Let_def proof (subst vec_of_list_append, subst index_append_vec(1), goal_cases) case 1 then show ?case using assms by (subst dim_vec_of_list)+ (split if_split, subst length_b_list_last, subst length_concat_map_b_list, auto) next case 2 then show ?case using assms by (subst dim_vec_of_list, subst length_concat_map_b_list, subst vec_index_vec_of_list)+ (auto split: if_splits simp add: b_list_last_def) qed lemma gen_bhle_last4: assumes "length as > 0" shows "(gen_bhle as) $ ((length as-1) * 5 + 3) = b5 (length as) (2*(\ias!i\)+1)" unfolding gen_bhle_def Let_def proof (subst vec_of_list_append, subst index_append_vec(1), goal_cases) case 1 then show ?case using assms by (subst dim_vec_of_list)+ (split if_split, subst length_b_list_last, subst length_concat_map_b_list, auto) next case 2 then show ?case using assms by (subst dim_vec_of_list, subst length_concat_map_b_list, subst vec_index_vec_of_list)+ (auto split: if_splits simp add: b_list_last_def) qed text \Up to last values of \gen_bhle\\ lemma b_list_nth: assumes "ii. b_list as i M) [0..i. b_list as i M) [0..i. b_list as i M) [0..i. b_list as i M) [i..i. b_list as i M) [0..i. b_list as i M) [0..i. b_list as i M) [i..i. b_list as i M) [0..i. b_list as i M) [i..i. b_list as i M) [i..i. b_list as i M) [i+1..i. b_list as i M) [0..i. b_list as i M) [0..i. b_list as i M) [i+1..i. b_list as i M) [0..i. b_list as i M) [i+1..i. b_list as i M) [0..i. b_list as i M) [i + 1..i. b_list as i M) [i+1..i. b_list as i M) [i+1..i. b_list as i M) [0..ias!i\)+1) (as!i)" unfolding gen_bhle_def Let_def proof (subst vec_of_list_append, subst index_append_vec(1), goal_cases) case 1 then show ?case using assms by (subst dim_vec_of_list)+ (split if_split, subst length_b_list_last, subst length_concat_map_b_list, auto) next case 2 define M where "M = (2 * (\ias ! i\) + 1)" then show ?case unfolding M_def[symmetric] using assms by (subst dim_vec_of_list, subst length_concat_map_b_list, subst vec_index_vec_of_list)+ (subst b_list_nth0[OF assms, of M], auto split: if_splits, simp add: b_list_def) qed lemma gen_bhle_1: assumes "iias!i\)+1)" unfolding gen_bhle_def Let_def proof (subst vec_of_list_append, subst index_append_vec(1), goal_cases) case 1 then show ?case using assms by (subst dim_vec_of_list)+ (split if_split, subst length_b_list_last, subst length_concat_map_b_list, auto) next case 2 define M where "M = (2 * (\ias ! i\) + 1)" then show ?case unfolding M_def[symmetric] using assms by (subst dim_vec_of_list, subst length_concat_map_b_list, subst vec_index_vec_of_list)+ (subst b_list_nth[OF assms, of 1 M], auto split: if_splits, simp add: b_list_def) qed lemma gen_bhle_4: assumes "iias!i\)+1)" unfolding gen_bhle_def Let_def proof (subst vec_of_list_append, subst index_append_vec(1), goal_cases) case 1 then show ?case using assms by (subst dim_vec_of_list)+ (split if_split, subst length_b_list_last, subst length_concat_map_b_list, auto) next case 2 define M where "M = (2 * (\ias ! i\) + 1)" then show ?case unfolding M_def[symmetric] using assms by (subst dim_vec_of_list, subst length_concat_map_b_list, subst vec_index_vec_of_list)+ (subst b_list_nth[OF assms, of 4 M], auto split: if_splits, simp add: b_list_def) qed lemma gen_bhle_2: assumes "iias!i\)+1) (as!i)" unfolding gen_bhle_def Let_def proof (subst vec_of_list_append, subst index_append_vec(1), goal_cases) case 1 then show ?case using assms by (subst dim_vec_of_list)+ (split if_split, subst length_b_list_last, subst length_concat_map_b_list, auto) next case 2 define M where "M = (2 * (\ias ! i\) + 1)" then show ?case unfolding M_def[symmetric] using assms by (subst dim_vec_of_list, subst length_concat_map_b_list, subst vec_index_vec_of_list)+ (subst b_list_nth[OF assms, of 2 M], auto split: if_splits, simp add: b_list_def) qed lemma gen_bhle_3: assumes "iias!i\)+1)" unfolding gen_bhle_def Let_def proof (subst vec_of_list_append, subst index_append_vec(1), goal_cases) case 1 then show ?case using assms by (subst dim_vec_of_list)+ (split if_split, subst length_b_list_last, subst length_concat_map_b_list, auto) next case 2 define M where "M = (2 * (\ias ! i\) + 1)" then show ?case unfolding M_def[symmetric] using assms by (subst dim_vec_of_list, subst length_concat_map_b_list, subst vec_index_vec_of_list)+ (subst b_list_nth[OF assms, of 3 M], auto split: if_splits, simp add: b_list_def) qed text \Well-definedness of reduction function\ lemma well_defined_reduction_subset_sum: assumes "a \ partition_problem_nonzero" shows "reduce_bhle_partition a \ bhle" using assms unfolding partition_problem_nonzero_def reduce_bhle_partition_def bhle_def proof (safe, goal_cases) case (1 I) text \Given a subset $I$, we must construct a vector $x$ such that the properties of \bhle\ hold.\ have "finite I" using 1 by (meson subset_eq_atLeast0_lessThan_finite) have "length a > 0" using \a\[]\ by auto define n where "n = length a" define minus_x::"int list" where "minus_x = [0,0,-1,1,1]" define plus_x::"int list" where "plus_x = [1,-1,0,-1,0]" define plus_x_last::"int list" where "plus_x_last = [1,-1,0,-1]" define plus_minus::"int list" where "plus_minus = (if n-1\I then plus_x else minus_x)" define minus_plus::"int list" where "minus_plus = (if n-1\I then minus_x else plus_x)" define x::"int vec" where "x = vec_of_list(concat (map (\i. if i\I then plus_minus else minus_plus) [0..i. if i \ I then plus_minus else minus_plus) [0..length a > 0\ unfolding n_def[symmetric] by auto have "0 < dim_vec x" unfolding dimx_eq_5dima using \length a > 0\ by linarith define M where "M = 2*(\ia!i\)+1" text \Some conditional lemmas for the proof.\ have x_nth: "x $ (i*5+j) = (if i\I then plus_minus ! j else minus_plus ! j)" if "ii. if i \ I then plus_minus else minus_plus) [0..i. if i \ I then plus_minus else minus_plus) [0..i. if i\I then plus_minus else minus_plus) = (\i. [(if i\I then plus_minus!0 else minus_plus!0), (if i\I then plus_minus!1 else minus_plus!1), (if i\I then plus_minus!2 else minus_plus!2), (if i\I then plus_minus!3 else minus_plus!3), (if i\I then plus_minus!4 else minus_plus!4)])" unfolding plus_minus_def minus_plus_def plus_x_def minus_x_def by auto then show ?thesis unfolding if_rew length_concat_map_5[of "(\i. if i\I then plus_minus!0 else minus_plus!0)" "(\i. if i\I then plus_minus!1 else minus_plus!1)" "(\i. if i\I then plus_minus!2 else minus_plus!2)" "(\i. if i\I then plus_minus!3 else minus_plus!3)" "(\i. if i\I then plus_minus!4 else minus_plus!4)" "[0.. int list" using that(1) by (metis append_Nil map_append not_gr_zero upt_0 upt_append) have "concat (map (\i. if i \ I then plus_minus else minus_plus) [0..i. if i \ I then plus_minus else minus_plus) [i..i. if i \ I then plus_minus else minus_plus) [0.. = (if i \ I then plus_minus!j else minus_plus!j)" proof - have concat_rewr: "concat (map f [i.. int list" using that(1) upt_conv_Cons by force have length_if: "length (if i \ I then plus_minus else minus_plus) = 5" using length_plus_minus length_minus_plus by auto show ?thesis unfolding concat_rewr nth_append length_if using \j<5\ by auto qed finally show ?thesis unfolding x_def by (subst vec_index_vec_of_list) (subst nth_append, use lt in \auto\) qed have x_nth0: "x $ (i*5) = (if i\I then plus_minus ! 0 else minus_plus ! 0)" if "ii. if i \ I then plus_minus else minus_plus) [0..j=0..<5. (gen_bhle a) $ (i*5+j) * x $ (i*5+j)) = (b1 (i+1) M (a!i)) - (b2 (i+1) M) - (b5 (i+1) M)" if "i\I-{length a-1}" "n-1\I" for i proof - have "(\j=0..<5. (gen_bhle a) $ (i*5+j) * x $ (i*5+j)) = (gen_bhle a) $ (i*5) * x $ (i*5) + (gen_bhle a) $ (i*5+1) * x $ (i*5+1) + (gen_bhle a) $ (i*5+2) * x $ (i*5+2) + (gen_bhle a) $ (i*5+3) * x $ (i*5+3) + (gen_bhle a) $ (i*5+4) * x $ (i*5+4)" by (simp add: eval_nat_numeral) also have "\ = (b1 (i+1) M (a!i)) - (b2 (i+1) M) - (b5 (i+1) M)" using that 1 n_def \length a > 0\ apply (subst gen_bhle_0[of i a], fastforce) apply (subst gen_bhle_1[of i a], fastforce) apply (subst gen_bhle_2[of i a], fastforce) apply (subst gen_bhle_3[of i a], fastforce) apply (subst gen_bhle_4[of i a], fastforce) apply (subst x_nth[of i], fastforce, fastforce)+ apply (subst x_nth0[of i], fastforce) apply (unfold M_def plus_minus_def minus_plus_def plus_x_def minus_x_def) apply (simp add: eval_nat_numeral) done finally show ?thesis by auto qed have gen_bhle_in_I_minus: "(\j=0..<5. (gen_bhle a) $ (i*5+j) * x $ (i*5+j)) = (b3 (i+1) M) - (b4 (i+1) M (a!i)) + (b5 (i+1) M)" if "i\I-{length a-1}" "n-1\I" for i proof - have "(\j=0..<5. (gen_bhle a) $ (i*5+j) * x $ (i*5+j)) = (gen_bhle a) $ (i*5) * x $ (i*5) + (gen_bhle a) $ (i*5+1) * x $ (i*5+1) + (gen_bhle a) $ (i*5+2) * x $ (i*5+2) + (gen_bhle a) $ (i*5+3) * x $ (i*5+3) + (gen_bhle a) $ (i*5+4) * x $ (i*5+4)" by (simp add: eval_nat_numeral) also have "\ = (b3 (i+1) M) - (b4 (i+1) M (a!i)) + (b5 (i+1) M)" using that 1 n_def \length a > 0\ apply (subst gen_bhle_0[of i a], fastforce) apply (subst gen_bhle_1[of i a], fastforce) apply (subst gen_bhle_2[of i a], fastforce) apply (subst gen_bhle_3[of i a], fastforce) apply (subst gen_bhle_4[of i a], fastforce) apply (subst x_nth[of i], fastforce, fastforce)+ apply (subst x_nth0[of i], fastforce) apply (unfold M_def plus_minus_def minus_x_def) apply (simp add: eval_nat_numeral) done finally show ?thesis by auto qed have gen_bhle_not_in_I_plus: "(\j=0..<5. (gen_bhle a) $ (i*5+j) * x $ (i*5+j)) = (b3 (i+1) M) - (b4 (i+1) M (a!i)) + (b5 (i+1) M)" if "i\{0..I" for i proof - have "(\j=0..<5. (gen_bhle a) $ (i*5+j) * x $ (i*5+j)) = (gen_bhle a) $ (i*5) * x $ (i*5) + (gen_bhle a) $ (i*5+1) * x $ (i*5+1) + (gen_bhle a) $ (i*5+2) * x $ (i*5+2) + (gen_bhle a) $ (i*5+3) * x $ (i*5+3) + (gen_bhle a) $ (i*5+4) * x $ (i*5+4)" by (simp add: eval_nat_numeral) also have "\ = (b3 (i+1) M) - (b4 (i+1) M (a!i)) + (b5 (i+1) M)" using that 1 n_def \length a > 0\ apply (subst gen_bhle_0[of i a], fastforce) apply (subst gen_bhle_1[of i a], fastforce) apply (subst gen_bhle_2[of i a], fastforce) apply (subst gen_bhle_3[of i a], fastforce) apply (subst gen_bhle_4[of i a], fastforce) apply (subst x_nth[of i], fastforce, fastforce)+ apply (subst x_nth0[of i], fastforce) apply (unfold M_def minus_plus_def minus_x_def) apply (simp add: eval_nat_numeral) done finally show ?thesis by auto qed have gen_bhle_not_in_I_minus: "(\j=0..<5. (gen_bhle a) $ (i*5+j) * x $ (i*5+j)) = (b1 (i+1) M (a!i)) - (b2 (i+1) M) - (b5 (i+1) M)" if "i\{0..I" for i proof - have "(\j=0..<5. (gen_bhle a) $ (i*5+j) * x $ (i*5+j)) = (gen_bhle a) $ (i*5) * x $ (i*5) + (gen_bhle a) $ (i*5+1) * x $ (i*5+1) + (gen_bhle a) $ (i*5+2) * x $ (i*5+2) + (gen_bhle a) $ (i*5+3) * x $ (i*5+3) + (gen_bhle a) $ (i*5+4) * x $ (i*5+4)" by (simp add: eval_nat_numeral) also have "\ = (b1 (i+1) M (a!i)) - (b2 (i+1) M) - (b5 (i+1) M)" using that 1 n_def \length a > 0\ apply (subst gen_bhle_0[of i a], fastforce) apply (subst gen_bhle_1[of i a], fastforce) apply (subst gen_bhle_2[of i a], fastforce) apply (subst gen_bhle_3[of i a], fastforce) apply (subst gen_bhle_4[of i a], fastforce) apply (subst x_nth[of i], fastforce, fastforce)+ apply (subst x_nth0[of i], fastforce) apply (unfold M_def minus_plus_def plus_x_def) apply (simp add: eval_nat_numeral) done finally show ?thesis by auto qed have gen_bhle_last: "(\j=0..<4. (gen_bhle a) $ ((n-1)*5+j) * x $ ((n-1)*5+j)) = (b1 n M (a!(n-1))) - (b2_last n M) - (b5 n M)" proof - have "(\j=0..<4. (gen_bhle a) $ ((n-1)*5+j) * x $ ((n-1)*5+j)) = (gen_bhle a) $ ((n-1)*5) * x $ ((n-1)*5) + (gen_bhle a) $ ((n-1)*5+1) * x $ ((n-1)*5+1) + (gen_bhle a) $ ((n-1)*5+2) * x $ ((n-1)*5+2) + (gen_bhle a) $ ((n-1)*5+3) * x $ ((n-1)*5+3)" by (simp add: eval_nat_numeral) also have "\ = (b1 n M (a!(n-1))) - (b2_last n M) - (b5 n M)" using 1 n_def \length a > 0\ unfolding n_def apply (subst gen_bhle_last0[of a, OF \length a > 0\]) apply (subst gen_bhle_last1[of a, OF \length a > 0\]) apply (subst gen_bhle_last3[of a, OF \length a > 0\]) apply (subst gen_bhle_last4[of a, OF \length a > 0\]) apply (subst x_nth_last, simp)+ apply (subst x_nth0_last, simp add: n_def) apply (unfold M_def plus_x_last_def) apply (auto simp add: eval_nat_numeral last_conv_nth) done finally show ?thesis by auto qed text \The actual proof. \ have "(gen_bhle a) \ x = 0" proof - define f where "f = (\i. (\j = 0..<5. gen_bhle a $ (i*5+j) * x $ (i*5+j)))" have "(gen_bhle a) \ x = (\i = (\i<(n-1)*5. (gen_bhle a) $ i * x $ i) + (\i = (n-1)*5..<(n-1)*5 +4. (gen_bhle a) $ i * x $ i)" proof (subst split_sum_mid_less[of "(n-1)*5" "n*5-1"], goal_cases) case 1 then show ?case unfolding n_def using \0 < length a\ by linarith next case 2 have "n * 5 - 1 = (n-1) * 5 + 4" unfolding n_def using \0 < length a\ by linarith then show ?case by auto qed also have "\ = (\i = 0..j=0..<4. gen_bhle a $ ((n-1)*5+j) * x $ ((n-1)*5+j))" proof - have *: "(+) ((n - 1) * 5) ` {0..<4} = {(n-1)*5..<(n-1)*5+4}" by auto have "(\i = (n - 1) * 5..<(n - 1) * 5 + 4. gen_bhle a $ i * x $ i) = (\j = 0..<4. gen_bhle a $ ((n - 1) * 5 + j) * x $ ((n - 1) * 5 + j))" using sum.reindex[of "(\j. (n-1)*5+j)" "{0..<4}" "(\i. gen_bhle a $ i * x $ i)"] unfolding comp_def * by auto then show ?thesis unfolding f_def lessThan_atLeast0 by (subst sum_split_idx_prod[of "(\i. (gen_bhle a) $ i * x $ i)" "n-1" 5], auto) qed also have "\ = (\i\I-{n-1}. f i) + (\i\{0..j=0..<4. gen_bhle a $ ((n-1)*5+j) * x $ ((n-1)*5+j))" proof - have "I - {n - 1} \ (({0..finite I\) qed also have "\ = M * (\i\{0..I") case True have "sum f (I - {n - 1}) + sum f ({0..j = 0..<4. gen_bhle a $ ((n - 1) * 5 + j) * x $ ((n - 1) * 5 + j)) = (\i\I-{n-1}. (b1 (i+1) M (a!i)) - (b2 (i+1) M) - (b5 (i+1) M)) + (\i\{0..i\I-{n-1}. f i) = (\i\I-{n-1}. (b1 (i+1) M (a!i)) - (b2 (i+1) M) - (b5 (i+1) M)) " unfolding f_def using gen_bhle_in_I_plus[OF _ True] by (simp add: n_def) moreover have "(\i\{0..i\{0.. = (\i\I-{n-1}. (a!i) + (M * 5^(4*(i+1)-4) - M * 5^(4*(i+1)))) + (\i\{0..I-{n-1}" for i unfolding b1_def b2_def b5_def by (smt (verit, best) add_uminus_conv_diff right_diff_distrib) moreover have "b3 (i + 1) M - b4 (i + 1) M (a ! i) + b5 (i + 1) M = - (a!i) + (M * 5^(4*(i+1)-4) - M * 5^(4*(i+1)))" if "i\{0.. = (\i\I-{n-1}. (a!i)) + (\i\{0..i\{0.. I = I - {n -1}" using "1"(1) n_def by auto have sets2: "{0..i\I-{n-1}. M * 5^(4*i+4*1-4) - M * 5^(4*i+4*1)) + (\i\{0..i = 0..i. M * 5^(4*i+4*1-4) - M * 5^(4*i+4*1))" "I"] \finite I\ unfolding sets1 sets2 by auto show ?thesis apply (subst distrib_left)+ apply (subst sum.distrib)+ apply (subst sum_distrib_left) apply (subst right_diff_distrib)+ apply (subst subs[symmetric]) apply auto done qed also have "\ = (\i\I. (a!i)) + (\i\{0..i\{0.. {n-1})" using True by auto have "sum ((!) a) (I - {n-1}) + a ! (n-1) = sum ((!) a) I" by (subst sum.Int_Diff[of "I" _ "{n-1}"], unfold *, auto simp add: \finite I\) then show ?thesis using True by auto qed also have "\ = M * (\i\{0..I*) next case False have "sum f (I - {n - 1}) + sum f ({0..j = 0..<4. gen_bhle a $ ((n - 1) * 5 + j) * x $ ((n - 1) * 5 + j)) = (\i\{0..i\I-{n-1}. (b3 (i+1) M) - (b4 (i+1) M (a!i)) + (b5 (i+1) M)) + (b1 n M (a!(n-1))) - (b2_last n M) - (b5 n M)" proof - have "(\i\{0..i\{0..i\I-{n-1}. f i) = (\i\I-{n-1}. (b3 (i+1) M) - (b4 (i+1) M (a!i)) + (b5 (i+1) M)) " unfolding f_def using gen_bhle_in_I_minus[OF _ False] by (simp add: n_def) ultimately show ?thesis unfolding f_def using gen_bhle_last by auto qed also have "\ = (\i\{0..i\I-{n-1}. - (a!i) + (M * 5^(4*(i+1)-4) - M * 5^(4*(i+1)))) + (a!(n-1)) + M * 5^(4*n-4) - M*1" proof - have "b1 (i + 1) M (a ! i) - b2 (i + 1) M - b5 (i + 1) M = (a!i) + (M * 5^(4*(i+1)-4) - M * 5^(4*(i+1)))" if "i\{0..I-{n-1}" for i unfolding b3_def b4_def b5_def by (smt (verit, best) add_uminus_conv_diff right_diff_distrib) moreover have "b1 n M (a ! (n - 1)) - b2_last n M - b5 n M = (a!(n-1)) + M * 5^(4*n-4) - M" unfolding b1_def b2_last_def b5_def by (simp add: distrib_left) moreover have "- b4_last n M (a ! (n - 1)) + b5 n M = -(a!(n-1)) - M * 5^(4*n-2) - M" unfolding b4_last_def b5_def by (simp add: distrib_left) ultimately show ?thesis by auto qed also have "\ = (\i\{0..i\I-{n-1}. - (a!i)) + M * (\i\{0.. I = I - {n -1}" using "1"(1) n_def by auto have sets2: "{0..i\{0..i\I-{n-1}. M * 5^(4*i+4*1-4) - M * 5^(4*i+4*1)) = (\i = 0..i. M * 5^(4*i+4*1-4) - M * 5^(4*i+4*1))" "I"] \finite I\ unfolding sets1 sets2 by auto show ?thesis apply (subst distrib_left)+ apply (subst sum.distrib)+ apply (subst sum_distrib_left) apply (subst right_diff_distrib)+ apply (subst subs[symmetric]) apply auto done qed also have "\ = (\i\{0..i\I. - (a!i)) + M * (\i\{0.. {n-1} = {n-1}" using False n_def \length a > 0\ by auto then have **: "a ! (n-1) = sum ((!) a) (({0.. {n-1})" using False by auto have "sum ((!) a) (({0..finite I\) then show ?thesis using False by auto qed also have "\ = M * (\i\{0.. = M * ((\i\{1..i = Suc 0..i. 5 ^ (4*(i+1) - 4) - 5 ^ (4*(i+1)))) {0..i. 5^(4*i-4) - 5^(4*i))" 0 "n-1"] unfolding comp_def by auto show ?thesis by (subst distrib_left[symmetric], subst right_diff_distrib, subst mult_1_right) (subst sums[symmetric], use \0 < length a\ n_def in \force\) qed also have "\ = M * (((\i\{1..i\{1..i. 5^(4*i-4))" "(\i. (-1) * 5^(4*i))" "{1.. = M * ((\i\{1..i\{0..i. 5^(4*i-4))"] sum.atLeast_Suc_lessThan[of 0 n "(\i. 5^(4*i))"] by (smt (verit, best) One_nat_def Suc_eq_plus1 Suc_leI \0 < length a\ mult_eq_0_iff n_def power_0) also have "\ = M * ((\i\{0..i\{0..i. 5^(4*i-4))" 0 n] by auto also have "\ = 0" by auto finally show ?thesis by blast qed moreover have "dim_vec x = dim_vec (gen_bhle a)" using dim_vec_gen_bhle[OF \a\[]\] dimx_eq_5dima by simp moreover have "x \ 0\<^sub>v (dim_vec x)" proof (rule ccontr) assume "\ x \ 0\<^sub>v (dim_vec x)" then have "x = 0\<^sub>v (dim_vec x)" by auto have "dim_vec x > 3" using \0 < length a\ dimx_eq_5dima by linarith have "(n - Suc 0) * 5 + 3 < dim_vec x" unfolding dimx_eq_5dima n_def using \length a > 0\ by linarith then have "x $ ((n-1)*5 + 3) = 0" using \dim_vec x > 3\ by (subst \x=0\<^sub>v (dim_vec x)\, subst index_zero_vec(1)[of "(n-1)*5+3"], auto) moreover have "x $ ((n-1)*5+3) \ 0" proof - have "\ ((n - 1) * 5 + 3 < (n - 1) * 5)" by auto then show ?thesis unfolding x_def vec_of_list_index nth_append length_concat_map plus_x_last_def by auto qed ultimately show False by auto qed moreover have "\x\\<^sub>\ \ 1" proof - let ?x_list = "(concat (map (\i. if i \ I then plus_minus else minus_plus) [0.. {-1,0,1::int}" using \length a > 0\ 1 unfolding n_def plus_minus_def minus_plus_def plus_x_def minus_x_def by (subst set_concat, subst set_map)(auto simp add: atLeast0LessThan) have "?x_list ! i \ {-1,0,1::int}" if "i< length ?x_list" for i using nth_mem[OF that] set by auto then have *:"?x_list ! i \ {-1,0,1::int}" if "i< (n - 1) * 5" for i using that unfolding length_concat_map by auto have **: "plus_x_last ! (i - (n - 1) * 5)\{-1,0,1::int}" if "\ (i<(n-1)*5)" "ilength a > 0\ unfolding dimx_eq_5dima n_def by auto then show ?thesis unfolding plus_x_last_def by (smt (z3) add.assoc add_diff_cancel_right' diff_is_0_eq insertCI less_Suc_eq not_le not_less_iff_gr_or_eq nth_Cons' numeral_1_eq_Suc_0 numeral_Bit0 plus_1_eq_Suc) qed have "x$i\{-1,0,1::int}" if "ix$i\\1" if "ii. dim_vec x > i)" 0] \dim_vec x > 0\) qed ultimately show ?case by (subst exI[of _ x], auto) qed text \NP-hardness of reduction function\ lemma NP_hardness_reduction_subset_sum: assumes "reduce_bhle_partition a \ bhle" shows "a \ partition_problem_nonzero" using assms unfolding reduce_bhle_partition_def bhle_def partition_problem_nonzero_def proof (safe, goal_cases) case (1 x) text \Given a vector $x$ from \bhle\ find the subset $I$ such that it has the same sum as its complement.\ have "length a > 0" using 1(3) dim_vec_gen_bhle dim_vec_gen_bhle_empty by auto define I where "I = {i\{0..0}" define n where "n = length a" then have "n > 0" using \length a>0\ by auto have dim_vec_x_5n: "dim_vec x = 5 * n - 1" unfolding n_def using 1 by (metis dim_vec_gen_bhle dim_vec_gen_bhle_empty less_not_refl2) have "I\{0..This is the trickiest part in the proof. One first has to generate equations from $x$ which form conditions on the entries of $x$. To do so, we consider the formula \gen_bhle a \ x = 0\ and rewrite it in basis $5$. Every digit then gives us an arithmetic expression in entries of $x$ equals to zero. From these equations, we can deduce the wanted properties.\ moreover have "sum ((!) a) I = sum ((!) a) ({0..ia ! i\) + 1" text \Rewriting the first formula in a huge sum.\ define a0 where "a0 = (\i. if i mod (5::nat) \ {0,2} then a!(i div 5) else 0)" define a1 where "a1 = (\i. if i mod (5::nat) \ {0,4} then 1 else 0::int)" define a1_last where "a1_last = (\i. if i mod (5::nat) \ {0} then 1 else 0::int)" define a2 where "a2 = (\i. if i mod (5::nat) \ {0,1} then 1 else 0::int)" define a3 where "a3 = (\i. if i mod (5::nat) \ {4,2} then 1 else 0::int)" define a3_last where "a3_last = (\i. if i mod (5::nat) \ {2} then 1 else 0::int)" define a4 where "a4 = (\i. if i mod (5::nat) \ {0,2,3} then 1 else 0::int)" define a5 where "a5 = (\i. if i mod (5::nat) \ {1,2} then (if i div 5 < n-1 then 5^(4*(i div 5 +1)) else 1) else 0::int)" define a0_rest' where "a0_rest' = (\i. a1 i * 5 ^ (4 * (i div 5)) + a2 i * 5 ^ (4 * (i div 5) + 1) + a3 i * 5 ^ (4 * (i div 5) + 2) + a4 i * 5 ^ (4 * (i div 5) + 3) + a5 i)" define a0_last where "a0_last = (\i. a1_last i * 5 ^ (4 * (i div 5)) + a2 i * 5 ^ (4 * (i div 5) + 1) + a3_last i * 5 ^ (4 * (i div 5) + 2) + a4 i * 5 ^ (4 * (i div 5) + 3) + a5 i)" define a0_rest where "a0_rest = (\i. if i div 5 < n-1 then a0_rest' i else a0_last i)" let ?P0 = "(\i. b1 (i div 5 + 1) M (a!(i div 5)))" let ?P1 = "(\i. (if i div 5 < n-1 then b2 (i div 5 + 1) M else b2_last (i div 5 + 1) M))" let ?P4 = "(\i. (if i div 5 < n-1 then b3 (i div 5 + 1) M else 0))" let ?P2 = "(\i. (if i div 5 < n-1 then b4 (i div 5 + 1) M (a!(i div 5)) else b4_last (i div 5 + 1) M (a!(i div 5))))" let ?P3 = "(\i. b5 (i div 5 + 1) M)" have cases_a0: "(a0 i + M * (a0_rest i)) = (if i mod 5 = 0 then ?P0 i else (if i mod 5 = 1 then ?P1 i else (if i mod 5 = 2 then ?P2 i else (if i mod 5 = 3 then ?P3 i else ?P4 i))))" if "i<5*n" for i proof (cases "i div 5 < n-1") case True then show ?thesis by (subst mod_5_if_split[of i "a0 i + M * (a0_rest i)" ?P0 ?P1 ?P2 ?P3 ?P4]) (auto simp add: a0_def a0_rest_def a0_rest'_def a1_def a2_def a3_def a4_def a5_def b1_def b2_def b3_def b4_def b5_def) next case False then have "i div 5 = n-1" using that by auto then show ?thesis by (subst mod_5_if_split[of i "a0 i + M * (a0_rest i)" ?P0 ?P1 ?P2 ?P3 ?P4]) (auto simp add: False a0_def a0_rest_def a0_last_def a1_def a1_last_def a2_def a3_def a3_last_def a4_def a5_def b1_def b2_def b2_last_def b3_def b4_def b4_last_def b5_def) qed text \Splitting of the first part of the sum containing the $a_i$.\ have gen_bhle_nth: "gen_bhle a $ i = a0 i + M * (a0_rest i)" if "ilength a >0\ dim_vec_gen_bhle by auto have gen_bhle_subst: "gen_bhle a $ i = (concat (map (\i. b_list a i M) [0..length a > 0\ by auto have len_concat_map: "length ?concat_map = 5 * (n-1)"using length_concat_map_b_list . show ?thesis proof (cases "i div 5 < n-1") case True then have "i<5*(n-1)" by auto then have j_th: "(?concat_map @ ?last)! i = ?concat_map ! i" by (subst nth_append, subst len_concat_map, auto) have "concat (map (\i. b_list a i M) [0..i. b_list a i M) [0.. = (concat (map (\i. b_list a i M) [0..i. b_list a i M) [i div 5.. = concat (map (\i. b_list a i M) [i div 5..i. b_list a i M) [0..i. b_list a i M) [i div 5.. = (b_list a (i div 5) M @ concat (map (\i. b_list a i M) [i div 5 +1.. = b_list a (i div 5) M ! (i mod 5)" unfolding nth_append b_list_def by auto finally have i_th: "?concat_map ! i = b_list a (i div 5) M ! (i mod 5)" by auto show ?thesis apply(subst gen_bhle_subst, subst j_th, subst i_th, subst cases_a0) subgoal apply (use that dim_vec_gen_bhle n_def \length a > 0\ in \auto\) done subgoal apply (subst mod_5_if_split[of i "b_list a (i div 5) M ! (i mod 5)" ?P0 ?P1 ?P2 ?P3 ?P4]) apply (use True in \auto simp add: b_list_def that\) done done next case False then have "i \ 5*(n-1)" by auto then obtain j where "i = 5*(n-1) + j" and "j<4" using \i \length a > 0\ unfolding dim_gen_bhle n_def by (subst split_lower_plus_diff[of i "5*(length a-1)" "5*(length a)"], auto) have j_th: "(?concat_map @ ?last)! i = ?last ! j" using \i nth_append_length_plus[of ?concat_map ?last j] unfolding len_concat_map by (auto simp add: \i = 5*(n-1) + j\ \j<4\) have "j = i mod 5" using \i = 5*(n-1) + j\ \j<4\ by auto have "n = i div 5 + 1" using \i = 5*(n-1) + j\ \j<4\ \length a > 0\ n_def by auto then have "i div 5 = n-1" by auto have "last a = a ! (i div 5)" unfolding \i div 5 = n-1\ by (subst last_conv_nth[of a]) (use \length a > 0\ n_def in \auto\) have *: "i mod 5 = 4 \ [] ! 0 = 0" by (use \j < 4\ \j = i mod 5\ in \presburger\) show ?thesis apply(subst gen_bhle_subst, subst j_th, subst cases_a0) subgoal apply (use that dim_vec_gen_bhle n_def \length a > 0\ in \auto\) done subgoal apply (subst mod_5_if_split[of i "b_list_last a n M ! j" ?P0 ?P1 ?P2 ?P3 ?P4]) apply (auto simp add: b_list_last_def that \j = i mod 5\ \n = i div 5 + 1\ \last a = a ! (i div 5)\ *) done done qed qed have "gen_bhle a \ x = (\i<5*n-1. x $ i * (a0 i + M * a0_rest i))" using 1(1) gen_bhle_nth unfolding scalar_prod_def Let_def dim_vec_x_5n 1(2)[symmetric] by (smt (verit, best) lessThan_atLeast0 lessThan_iff mult.commute sum.cong) then have sum_gen_bhle: "(\i<5 * n-1. x $ i * (a0 i + M * a0_rest i)) = 0" using 1(1) by simp text \The first equation containing the $a_i$\ have eq_0: "(\ii<5*n-1. x$i * (a0_rest i)) = 0" proof - have *: "(\i<5*n-1. \a0 i\) < M" proof - have "5 * n - 1 = Suc (5 * n - 2)" using \n > 0\ by auto have "5 * n - 2 = Suc (5 * n - 3)" using \n > 0\ by auto have "5 * n - 3 = Suc (5 * n - 4)" using \n > 0\ by auto have "5 * n - 4 = Suc (5 * n - 5)" using \n > 0\ by auto have "(\i<5*n-1. \a0 i\) = (\i<5*(n-1). \a0 i\) + (\j<4.\a0 ((n-1)*5+j)\)" proof - have "5*n-5 = 5*(n-1)" by auto have "(\i<5*n-1. \a0 i\) = (\i<5*n-5. \a0 i\) + \a0 (5*(n-1))\ + \a0 (5*(n-1)+1)\ + \a0 (5*(n-1)+2)\ + \a0 (5*(n-1)+3)\" unfolding \5 * n - 1 = Suc (5 * n - 2)\ sum.lessThan_Suc[of "(\i.\a0 i\)" "5 * n - 2"] unfolding \5 * n - 2 = Suc (5 * n - 3)\ sum.lessThan_Suc[of "(\i.\a0 i\)" "5 * n - 3"] unfolding \5 * n - 3 = Suc (5 * n - 4)\ sum.lessThan_Suc[of "(\i.\a0 i\)" "5 * n - 4"] unfolding \5 * n - 4 = Suc (5 * n - 5)\ sum.lessThan_Suc[of "(\i.\a0 i\)" "5 * n - 5"] unfolding \5*n-5 = 5*(n-1)\ by (auto simp add: Suc3_eq_add_3 add.commute) moreover have "\a0 (5*(n-1))\ + \a0 (5*(n-1)+1)\ + \a0 (5*(n-1)+2)\ + \a0 (5*(n-1)+3)\ = (\j<4.\a0 ((n-1)*5+j)\)" by (simp add:eval_nat_numeral) ultimately show ?thesis using \5*n-5 = 5*(n-1)\ by auto qed also have "\ = (\ij<5. \a0 (i*5+j)\)) + (\j<4.\a0 ((n-1)*5+j)\)" using sum_split_idx_prod[of "(\i. \a0 i\)" "n-1" 5] by (simp add: lessThan_atLeast0 mult.commute) also have "\ = (\ij<5::nat. \if j \ {0, 2} then a ! i else 0\))" proof - have "(5::nat) = Suc 4" by auto have "(\ij<5::nat. \a0 (i*5+j)\)) = (\ij<5::nat. \if j \ {0, 2} then a ! i else 0\))" by(rule sum.cong[OF refl], rule sum.cong[OF refl], auto simp add: a0_def div_mult_self3[of 5]) moreover have "(\j<4::nat.\a0 ((n-1)*5+j)\) = (\j<4::nat.\if j \ {0, 2} then a ! (n-1) else 0\)" by(rule sum.cong[OF refl], auto simp add: a0_def div_mult_self3[of 5]) moreover have "(\j<4::nat.\if j \ {0, 2} then a ! (n-1) else 0\) = (\j<5::nat.\if j \ {0, 2} then a ! (n-1) else 0\)" unfolding \5=Suc 4\ using sum.lessThan_Suc[of "(\j.\if j \ {0, 2} then a ! (n-1) else 0\)" "4::nat"] by auto ultimately have *: "(\ij<5. \a0 (i*5+j)\)) + (\j<4.\a0 ((n-1)*5+j)\) = (\ij<5::nat. \if j \ {0, 2} then a ! i else 0\)) + (\j<5::nat. \if j \ {0, 2} then a ! (n-1) else 0\)" by auto show ?thesis by (subst *, subst sum.lessThan_Suc[symmetric], auto simp add: \n>0\) qed also have "\ = (\ia ! i\)" by (auto simp add: eval_nat_numeral) also have "\ = 2 * (\ia ! i\)" by (simp add: sum_distrib_left) finally show ?thesis unfolding M_def n_def by linarith qed have **: "\i<5*n-1. \x $ i\ \ 1" using 1(5) by (metis dim_vec_x_5n order_trans vec_index_le_linf_norm) have "(\i<5*n-1. x $ i * a0 i) = 0 \ (\i<5*n-1. x$i * (a0_rest i)) = 0" using split_eq_system[OF * ** sum_gen_bhle] by auto moreover have "(\i<5*n-1. x $ i * a0 i) = (\ij = i*5..j = i*5.. = (\j<5. x$(i*5 + j) * (if j \ {0, 2} then a!i else 0))" using mod_mult_self3[of i 5] div_rule by (subst sum.reindex[of "(\j. i*5+j)" "{0..<5}"], simp, unfold comp_def) (metis (no_types, lifting) One_nat_def lessThan_atLeast0 lessThan_iff nat_mod_lem not_less_eq not_numeral_less_one sum.cong) also have "\ = x$(i*5 + 0) * a!i + x$(i*5 + 2) * a!i" by (auto simp add: eval_nat_numeral split: if_splits) finally show ?thesis by (simp add: distrib_left mult.commute) qed then have *: "(\i<(n-1)*5. x $ i * a0 i) = (\ij = 5*(n-1)..<5*(n-1)+4. x $ j * a0 j) = ?h (n-1)" proof - have div_rule: "((n-1) * 5 + xa) div 5 = (n-1)" if "xa <4" for xa using that by auto have "(\j = (n-1)*5..<(n-1)*5+4. ?g j) = sum ?g ((+) ((n-1) * 5) ` {0..<4})" by (simp add: add.commute) also have "\ = (\j<4. x$((n-1)*5 + j) * (if j \ {0, 2} then a!(n-1) else 0))" proof - have *: "(\xa = 0..<4.?g ((n - 1) * 5 + xa)) = (\j = 0..<4. x $ ((n - 1) * 5 + j) * (if j \ {0, 2} then a ! (n - 1) else 0))" (is "(\xa = 0..<4. ?g' xa) = (\j = 0..<4. ?h' j)") by (rule sum.cong) auto show ?thesis using mod_mult_self3[of "n-1" 4] div_rule by (subst sum.reindex[of "(\j. (n-1)*5+j)" "{0..<4}"], simp, unfold comp_def lessThan_atLeast0) (use * in \auto\) qed also have "\ = x$((n-1)*5 + 0) * a!(n-1) + x$((n-1)*5 + 2) * a!(n-1)" by (auto simp add: eval_nat_numeral split: if_splits) finally show ?thesis unfolding a0_def by (simp add: distrib_left mult.commute) qed have "5 * (n - 1) < 5 * n - 1" using \n> 0\ by auto then have ***: "(\i<5*n-1. x $ i * a0 i) = (\i<5*(n-1). x $ i * a0 i) + (\i=5*(n-1)..<5*n-1. x $ i * a0 i)" by (subst split_sum_mid_less[of "5*(n-1)" "5*n-1"], auto) have ****: "5 * n - 1 = 5 * (n - 1) + 4" using \n> 0\ by auto show ?thesis by (unfold ***, subst mult.commute[of 5 "n-1"], unfold * **** **) (subst sum.lessThan_Suc[of ?h "n-1", symmetric], auto simp add: \n>0\) qed ultimately show "(\ii<5*n-1. x$i * (a0_rest i)) = 0" by auto qed let ?eq_0'_left = "(\i<5*n-1. x$i * (a0_rest i))" interpret digits 5 by (simp add: digits_def) have digit_a0_rest: "digit ?eq_0'_left k = 0" for k using eq_0' by (simp add: eq_0' digit_altdef) text \Define the digits in basis 5.\ define d1 where "d1 = (\i. x$(i*5) + (if ii. x$(i*5) + x$(i*5+1))" define d3 where "d3 = (\i. (if ii. x$(i*5) + x$(i*5+2) + x$(i*5+3))" define d5 where "d5 = (\i. x$(5*i+1) + x$(5*i+2))" define d where "d = (\k. (if k mod 4 = 0 then (if k = 0 then d5 (n-1) + d1 0 else (d1 (k div 4) + d5 (k div 4 -1))) else (if k mod 4 = 1 then d2 (k div 4) else (if k mod 4 = 2 then d3 (k div 4) else d4 (k div 4)))))" text \Rewrite the sum in basis 5.\ have rewrite_digits: "(\i<5*n-1. x$i * (a0_rest i)) = (\k<4*n. d k * 5^k)" proof - define f1::"nat \ nat \ int" where "f1 = (\i j. ((if i {0, 4} then 1 else 0) else (if j\{0} then 1 else 0)) * 5 ^ (4 * i) + (if j \ {0, 1} then 1 else 0) * 5 ^ (4 * i + 1) + (if i {4, 2} then 1 else 0) else (if j\{2} then 1 else 0)) * 5 ^ (4 * i + 2) + (if j \ {0, 2, 3} then 1 else 0) * 5 ^ (4 * i + 3) + (if j \ {1, 2} then if i < n - 1 then 5 ^ (4 * (i + 1)) else 1 else 0)))" have "f1 (n-1) 4 = 0" unfolding f1_def by auto define f2 where "f2 = (\i. x $ (i * 5) * (5 ^ (4 * i) + 5 ^ (4 * i + 1) + 5 ^ (4 * i + 3)) + x $ (i * 5 + 1) * (5 ^ (4 * i + 1) + (if i < n - 1 then 5 ^ (4 * (i + 1)) else 1)) + x $ (i * 5 + 4) * (if ii. d1 i * (5 ^ (4 * i)) + d5 i * (if i < n - 1 then 5 ^ (4 * (i + 1)) else 1) + d2 i * 5 ^ (4 * i + 1) + d3 i * 5 ^ (4 * i + 2) + d4 i * 5 ^ (4 * i + 3))" have f2_f3: "f2 i = f3 i" if "i = f3 i" unfolding f3_def d1_def d2_def d3_def d4_def d5_def using True by (subst distrib_right)+ (simp add: mult.commute[of 5 i]) ultimately show ?thesis by auto next case False then have "f2 i = x $ (i * 5) * 5 ^ (4 * i) + x $ (i * 5) * 5 ^ (4 * i + 1) + x $ (i * 5) * 5 ^ (4 * i + 3) + x $ (i * 5 + 1) * 5 ^ (4 * i + 1) + x $ (i * 5 + 1) + x $ (i * 5 + 2) * 5 ^ (4 * i + 2) + x $ (i * 5 + 2) * 5 ^ (4 * i + 3) + x $ (i * 5 + 2) + x $ (i * 5 + 3) * 5 ^ (4 * i + 3)" unfolding f2_def by (subst ring_distribs)+ simp also have "\ = f3 i" unfolding f3_def d1_def d2_def d3_def d4_def d5_def using False by(simp add: algebra_simps) ultimately show ?thesis by auto qed define x_pad where "x_pad = (\i. if i<5*n-1 then x$i else 1)" have pad: "(\i<5*n-1. x$i * (a0_rest i)) = (\i<5*n. x_pad i * (a0_rest i))" proof - have "Suc (5 * n - 1) = 5*n" using \n>0\ by auto have "(\i<5 * n - 1. x_pad i * a0_rest i) = (\i<5 * n - 1. x$i * a0_rest i)" by(rule sum.cong)(auto simp: x_pad_def) moreover have "x_pad (5 * n - 1) * a0_rest (5 * n - 1) = 0" proof - have "\((5 * n - 1) div 5 < n - 1)" using \n>0\ by auto moreover have "(5 * n - 1) mod 5 = 4" proof - have "5 * n - 1 = 5*(n-1)+4" using \n>0\ by auto show ?thesis unfolding \5 * n - 1 = 5*(n-1)+4\ by auto qed ultimately show ?thesis unfolding a0_rest_def a0_last_def a1_last_def a2_def a3_last_def a4_def a5_def by auto qed ultimately show ?thesis using sum.lessThan_Suc[of "(\i. x_pad i * (a0_rest i))" "5 * n - 1"] unfolding \Suc (5 * n - 1) = 5*n\ by auto qed have *: "(\i<5*n. x_pad i * (a0_rest i)) = (\ij<5. x_pad (i*5+j) * (a0_rest (i*5+j))))" (is "\ = (\ij<5. ?f0 i j))") using sum_split_idx_prod[of "(\i. x_pad i * a0_rest i)" n 5] unfolding mult.commute[of n 5] using atLeast0LessThan by auto also have "\ = (\ij<5. x_pad (i * 5 + j) * f1 i j)" proof (subst sum.cong[of _ _ "(\i. (\j<(5::nat). ?f0 i j))" "(\i. (\j<(5::nat). x_pad (i * 5 + j) * f1 i j))", symmetric], simp, goal_cases) case (1 i) have **: "a0_rest (i * 5 + j) = f1 i j" if "j<5" for j using that lt_5_split[of j] 1 unfolding f1_def a0_rest_def a0_rest'_def a0_last_def a1_def a1_last_def a2_def a3_def a3_last_def a4_def a5_def by auto show ?case by (rule sum.cong) (use ** 1 in auto) next case 2 then show ?case using * by auto qed also have "\ = (\ij<5. x_pad (i * 5 + j) * f1 i j) = x_pad ((n-1) * 5 + 0) * f1 (n-1) 0 + x_pad ((n-1) * 5 + 1) * f1 (n-1) 1 + x_pad ((n-1) * 5 + 2) * f1 (n-1) 2 + x_pad ((n-1) * 5 + 3) * f1 (n-1) 3 + x_pad ((n-1) * 5 + 4) * f1 (n-1) 4" by (simp add: eval_nat_numeral) also have "\ = x_pad ((n-1) * 5 + 0) * f1 (n-1) 0 + x_pad ((n-1) * 5 + 1) * f1 (n-1) 1 + x_pad ((n-1) * 5 + 2) * f1 (n-1) 2 + x_pad ((n-1) * 5 + 3) * f1 (n-1) 3" using \f1 (n-1) 4 = 0\ by auto also have "\ = f2 i" proof - have "x_pad ((n-1) * 5 + 0) * f1 (n-1) 0 = x $ ((n-1)*5) * (5^(4 * (n - 1)) + 5 ^ (4 * (n - 1) + 1) + 5 ^ (4 * (n - 1) + 3))" unfolding f1_def x_pad_def using \n>0\ by auto moreover have "x_pad ((n-1) * 5 + 1) * f1 (n-1) 1 = x $ ((n-1)*5 + 1) * (5^(4 * (n - 1) + 1) + 1)" unfolding f1_def x_pad_def using \n>0\ by auto moreover have "x_pad ((n-1) * 5 + 2) * f1 (n-1) 2 = x $ ((n-1)*5 + 2) * (5^(4 * (n - 1) + 2) + 5 ^ (4 * (n - 1) + 3) + 1)" unfolding f1_def x_pad_def using \n>0\ by auto moreover have "x_pad ((n-1) * 5 + 3) * f1 (n-1) 3 = x $ ((n-1)*5 + 3) * 5^(4 * (n - 1) + 3)" unfolding f1_def x_pad_def using \n>0\ by auto ultimately show ?thesis unfolding f2_def using \i=n-1\ by auto qed finally show ?thesis by auto qed qed also have "\ = (\i0 < length a\ n_def in \presburger\) also have "\ = (\ilength a > 0\ n_def in \auto\) also have "\ = (\i = (\iiiii = (\iiiii = ?f4") using sum.lessThan_Suc[of "(\i. d1 i * 5 ^ (4 * i))" "n-1"] using sum.lessThan_Suc[of "(\i. d2 i * 5 ^ (4 * i + 1))" "n-1"] using sum.lessThan_Suc[of "(\i. d3 i * 5 ^ (4 * i + 2))" "n-1"] using sum.lessThan_Suc[of "(\i. d4 i * 5 ^ (4 * i + 3))" "n-1"] using \0 < length a\ n_def by force finally have "(\i<5*n. x_pad i * (a0_rest i)) = ?f4" using * by auto then have "(\i<5*n-1. x$i * (a0_rest i)) = ?f4" using pad by auto moreover have "(\iii. d1 i * 5 ^ (4 * i))" "n-1"] using \0 < length a\ n_def by force moreover have "(\iiii. d1 (i + 1) * 5 ^ (4 * (i + 1)))" "(\i. d5 i * 5 ^ (4 * (i + 1)))" "{..ii\{1..i. i - 1) {1..0 < length a\ add_diff_cancel_left' atLeastLessThan_iff image_eqI n_def plus_1_eq_Suc zero_less_Suc) then show ?thesis by (subst sum.reindex_bij_betw[of "(\i. i-1)" "{1..i. (d1 (i + 1) + d5 i) * 5 ^ (4 * (i + 1)))", symmetric]) auto qed moreover have "(\i\{1..i0, of "(\i. (if i = 0 then d5 (n - 1) + d1 0 else d1 i + d5 (i - 1)) * 5 ^ (4 * i))"] unfolding atLeast0LessThan n_def by (auto) ultimately have "(\i<5*n-1. x$i * (a0_rest i)) = (\ii<5*n-1. x$i * (a0_rest i)) = (\i = (\ij<4. d (i*4+j) * 5^(i*4+j)))" proof (rule sum.cong, goal_cases) case (2 i) have d_rew: "(\j<4. d (i * 4 + j) * 5 ^ (i * 4 + j)) = d (i * 4) * 5 ^ (i * 4) + d (i * 4 + 1) * 5 ^ (i * 4 + 1) + d (i * 4 + 2) * 5 ^ (i * 4 + 2) + d (i * 4 + 3) * 5 ^ (i * 4 + 3)" by (simp add: eval_nat_numeral) have d1_rew: "d (i * 4) = (if i = 0 then d5 (n - 1) + d1 0 else d1 i + d5 (i - 1))" unfolding d_def by auto have d2_rew: "d (i*4+1) = d2 i" unfolding d_def by (metis add.commute add.right_neutral div_mult_self3 mod_Suc mod_div_trivial mod_mult_self2_is_0 one_eq_numeral_iff plus_1_eq_Suc semiring_norm(85) zero_neq_numeral) have d3_rew: "d (i*4+2) = d3 i" unfolding d_def by (metis add.commute add_2_eq_Suc' add_cancel_right_left div_less div_mult_self1 less_Suc_eq mod_mult_self1 nat_mod_lem numeral_Bit0 one_add_one zero_neq_numeral) have d4_rew: "d (i*4+3) = d4 i" unfolding d_def by auto show ?case by (subst d_rew, subst d1_rew, subst d2_rew, subst d3_rew, subst d4_rew) (auto simp add: mult.commute) qed auto moreover have "\ = (\k<4*n. d k * 5^k)" using sum_split_idx_prod[of "(\k. d k * 5^k)" n 4, symmetric] by (simp add: lessThan_atLeast0 mult.commute) ultimately show ?thesis by auto qed text \Some helping lemmas to get equations.\ have xi_le_1: "\x$i\\1" if "i< dim_vec x" for i using 1(5) that unfolding linf_norm_vec_Max by auto have xs_le_2: "\x$i + x$j\\2" if "i< dim_vec x" "j< dim_vec x" for i j proof - have "\x$i + x$j\ \ \x$i\ + \x$j\" by (auto simp add: abs_triangle_ineq) then show ?thesis using xi_le_1[OF that(1)] xi_le_1[OF that(2)] by auto qed have if_xi_le_1: "\(if i < n - 1 then x $ (i * 5 + 4) else 0)\ \ 1" if "i< n" for i using that xi_le_1[of "i*5+4"] unfolding dim_vec_x_5n by auto have prec_0: "i * 5 < dim_vec x" if "id1 i\\2" if "id2 i\\2" if "id3 i\\2" if "id4 i\\3" if "id5 i\\2" if "id5 (n - Suc 0) + d1 0\ < 5" using abs_d5[of "n-1"] abs_d1[of 0] \0 < length a\ n_def by fastforce moreover have "\d1 (i div 4) + d5 (i div 4 - Suc 0)\ < 5" if "0 < i" and "i<4*n" and "i mod 4 = 0" for i using that abs_d1[of "i div 4"] abs_d5[of "i div 4 - 1"] \0 < length a\ n_def by fastforce moreover have "\d2 (i div 4)\ < 5" if "i<4*n" and "i mod 4 = 1" for i using that abs_d2[of "i div 4"] \0 < length a\ n_def by fastforce moreover have "\d3 (i div 4)\ < 5" if "i<4*n" and "i mod 4 = 2" for i using that abs_d3[of "i div 4"] \0 < length a\ n_def by fastforce moreover have "\d4 (i div 4)\ < 5" if "i<4*n" and "i mod 4 = 3" for i using that abs_d4[of "i div 4"] \0 < length a\ n_def by fastforce ultimately have d_lt_5: "\d i\ < 5" if "i < 4 * n" for i using that by (subst mod_4_choices[of i], unfold d_def, auto) have sum_zero: "(\k<4*n. d k * 5^k) = 0" using eq_0' rewrite_digits by auto then have d_eq_0: "d k = 0" if "k<4*n" for k using respresentation_in_basis_eq_zero[OF sum_zero _ _ that] d_lt_5 by auto text \These are the main equations.\ have eq_1: "x$(i*5) + (if i {1..0 < length a\ add_cancel_right_left bits_mod_0 bot_nat_0.not_eq_extremum - mult.commute mult_is_0 n_def zero_neq_numeral) + using d_eq_0 [of 0] \0 < n\ unfolding d_def d1_def d5_def + by (cases \n = 1\) (simp_all add: ac_simps) have eq_3: "x$(i*5) + x$(i*5+1) = 0" if "i\{0..{0..0 < length a\ add.right_neutral add_self_div_2 add_self_mod_2 atLeastLessThan_iff bits_one_mod_two_eq_one div_mult2_eq div_mult_self4 mod_mult2_eq mod_mult_self3 mult.commute mult_2 mult_is_0 n_def nat_1_add_1 nat_mod_lem neq0_conv numeral_Bit0 one_div_two_eq_zero) qed have eq_5: "x$(i*5) + x$(i*5+2) + x$(i*5+3) = 0" if "i\{0..{0..{0..{1.. {0..{0..This defines the weight of the solution, since $x_{i,0} + x_{i,4}$ does not depend on the index i. We take $x_{n-1,0} + x_{n-1,4}$, since we omitted the last element (thus $x_{n-1,4} = 0$) to ensure that the weight has absolute value at most $1$.\ define w where "w = x$((n-1)*5)" have w_eq_02: "w = x$(i*5) + (if i{0..n-1" using that by auto then show ?thesis proof (induct rule: Nat.inc_induct) case (step m) then show ?case unfolding w_def using eq_1'[of "Suc m"] by auto qed (unfold w_def, auto) qed have "\w\ \ 1" using xi_le_1[of "(n-1)*5"] \n > 0\ unfolding w_def dim_vec_x_5n by auto text \Rule out the all zero solution.\ moreover have "w\0" proof (rule ccontr) assume "\ w \ 0" then have "w = 0" by blast then have "x$((n-1)*5) = 0" unfolding w_def by auto have zero_eq_min2: "x$(i*5) = - (if i{0..w=0\ by auto have two0_plus_4: "2 * x$(i*5) + x$(i*5+3) = 0" if "i \ {0..x$((n-1)*5) = 0\ by auto next case False then have "i \ {0.. 0" then have "\2 * x $ (i * 5) + x $ (i * 5 + 3)\\1" using xi_le_1[OF prec_0[where i=i]] xi_le_1[OF prec_i[where i=i and j=3]] by (auto simp add: \i) then show False using two0_plus_4[OF \i \ {0..] by auto qed qed have "x$j = 0" if "j<5*n" "j mod 5 = 0" for j proof - from \j mod 5 = 0\ \j < 5 * n\ obtain i where "i*5 = j" "ii*5 = j\[symmetric] using p_zero[OF \i] by auto qed moreover have p_one: "x$(i*5+1) = 0" if "ij<5*n\ \j mod 5 = 1\ by (metis add.commute less_mult_imp_div_less mod_mult_div_eq mult.commute) then show ?thesis unfolding \i*5+1 = j\[symmetric] using p_one[OF \i] by auto qed moreover have p_four: "x$(i*5+4) = 0" if "iw=0\ by auto then have "x$j = 0" if "j<5*n-1" "j mod 5 = 4" for j proof - obtain i where "i*5+4 = j" "ij<5*n-1\ \j mod 5 = 4\ by (metis add.assoc add.commute less_Suc_eq less_diff_conv less_mult_imp_div_less mod_mult_div_eq mult.commute mult_Suc_right not_less_eq numeral_nat(3) plus_1_eq_Suc) then show ?thesis unfolding \i*5+4 = j\[symmetric] using p_four by auto qed moreover have p_two: "x$(i*5+2) = 0" if "ii] eq_4' by auto next case False then have "i=n-1" using that by auto show ?thesis using eq_4[of "n-1"] unfolding \i=n-1\ using \n>0\ by auto qed then have "x$j = 0" if "j<5*n" "j mod 5 = 2" for j proof - obtain i where "i*5+2 = j" "ij<5*n\ \j mod 5 = 2\ by (metis add.commute less_mult_imp_div_less mod_mult_div_eq mult.commute) then show ?thesis unfolding \i*5+2 = j\[symmetric] using p_two[OF \i] by auto qed moreover have p_three: "x$(i*5+3) = 0" if "ij<5*n\ \j mod 5 = 3\ by (metis add.commute less_mult_imp_div_less mod_mult_div_eq mult.commute) then show ?thesis unfolding \i*5+3 = j\[symmetric] using p_three[OF \i] by auto qed ultimately have "x$j = 0" if "j<5*n-1" for j using mod_5_choices[of j "(\j. x $ j = 0)"] that by auto then have "x = 0\<^sub>v (dim_vec x)" unfolding dim_vec_x_5n[symmetric] by auto then show False using 1(4) by auto qed text \Then we can deduce the wanted property for both cases $w=1$ and $w=-1$. The only differences between the two cases is the switch of signs.\ ultimately have "w=1 \ w = -1" by auto then consider (pos) "w=1" | (neg) "w=-1" by blast then show ?thesis proof cases case pos have 01:"(x$(i*5) = 0 \ x$(i*5+4) = 1) \ (x$(i*5) = 1 \ x$(i*5+4) = 0)" if "in>0\ unfolding dim_vec_x_5n by auto then have "x$(i*5) \ {-1,0,1}" using xi_le_1[of "i*5"] by auto then consider (minus) "x$(i*5) = -1" | (zero) "x$(i*5) = 0" | (plus) "x$(i*5) = 1" by blast then show ?thesis proof (cases) case minus then have "2 = x $ (i * 5 + 4)" using w_eq_02[of i] that \w=1\ by auto then have False using xi_le_1[of "i*5+4"] that unfolding dim_vec_x_5n by linarith then show ?thesis by auto next case zero then have "x $ (i * 5 + 4) = 1" using w_eq_02[of i] that unfolding \w=1\ by auto then show ?thesis using zero by auto next case plus then have "x $ (i * 5 + 4) = 0" using w_eq_02[of i] that unfolding \w=1\ by auto then show ?thesis using plus by auto qed qed have "n-1\I" using w_def unfolding \w=1\ I_def n_def[symmetric] by (auto simp add: \n>0\ mult.commute) have I_0: "x$(i*5) = 1" if "i\I" for i proof (cases "iw=1\ by auto qed have I_2: "x$(i*5+2) = 0" if "i\I" for i proof (cases "i0 < n\) moreover have "1 + x $ (i * 5 + 1) + x $ (i * 5 + 2) = 0" using eq_2 w_eq_02[of 0] unfolding \i=n-1\ by (metis \0 < n\ add_cancel_right_left atLeast0LessThan lessThan_iff mult_zero_left pos) ultimately show ?thesis by auto qed have notI_0: "x$(i*5) = 0" if "i\{0..n-1\I\ that by auto qed have notI_2: "x$(i*5+2) = -1" if "i\{0..n-1\I\ that by auto qed have "0 = (\i\I. (x $ (i * 5) + x $ (i * 5 + 2)) * a ! i) + (\i\{0..I" for i using I_0 I_2 that by auto moreover have "(x $ (i * 5) + x $ (i * 5 + 2)) = -1" if "i\{0..i\I. a ! i) - (\i\{0.. x$(i*5+4) = -1) \ (x$(i*5) = -1 \ x$(i*5+4) = 0)" if "in>0\ unfolding dim_vec_x_5n by auto then have "x$(i*5) \ {-1,0,1}" using xi_le_1[of "i*5"] by auto then consider (minus) "x$(i*5) = -1" | (zero) "x$(i*5) = 0" | (plus) "x$(i*5) = 1" by blast then show ?thesis proof (cases) case plus then have "-2 = x $ (i * 5 + 4)" using w_eq_02[of i] that \w=-1\ by force then have False using xi_le_1[of "i*5+4"] that unfolding dim_vec_x_5n by linarith then show ?thesis by auto next case zero then have "x $ (i * 5 + 4) = -1" using w_eq_02[of i] that unfolding \w=-1\ by auto then show ?thesis using zero by auto next case minus then have "x $ (i * 5 + 4) = 0" using w_eq_02[of i] that unfolding \w=-1\ by auto then show ?thesis using minus by auto qed qed have "n-1\I" using w_def unfolding \w=-1\ I_def n_def[symmetric] by (auto simp add: \n>0\ mult.commute) have I_0: "x$(i*5) = -1" if "i\I" for i proof (cases "iw=-1\ by auto qed have I_2: "x$(i*5+2) = 0" if "i\I" for i proof (cases "i0 < n\) moreover have "-1 + x $ (i * 5 + 1) + x $ (i * 5 + 2) = 0" using eq_2 w_eq_02[of 0] unfolding \i=n-1\ by (metis \0 < n\ add_cancel_right_left lessThan_atLeast0 lessThan_iff mult_zero_left neg) ultimately show ?thesis by auto qed have notI_0: "x$(i*5) = 0" if "i\{0..n-1\I\ that by auto qed have notI_2: "x$(i*5+2) = 1" if "i\{0..n-1\I\ that by auto qed have "0 = (\i\I. (x $ (i * 5) + x $ (i * 5 + 2)) * a ! i) + (\i\{0..I" for i using I_0 I_2 that by auto moreover have "(x $ (i * 5) + x $ (i * 5 + 2)) = 1" if "i\{0..i\I. a ! i) - (\i\{0..length a > 0\ by auto qed text \The Gap-SVP is NP-hard.\ lemma "is_reduction reduce_bhle_partition partition_problem_nonzero bhle" unfolding is_reduction_def proof (safe, goal_cases) case (1 a) then show ?case using well_defined_reduction_subset_sum by auto next case (2 a) then show ?case using NP_hardness_reduction_subset_sum by auto qed end \ No newline at end of file diff --git a/thys/Crypto_Standards/PKCS1v2_2.thy b/thys/Crypto_Standards/PKCS1v2_2.thy --- a/thys/Crypto_Standards/PKCS1v2_2.thy +++ b/thys/Crypto_Standards/PKCS1v2_2.thy @@ -1,3298 +1,3300 @@ theory PKCS1v2_2 imports Words begin text \The PKCS #1 standard defines primitives and schemes related to the RSA cryptosystem. It is currently in version 2.2. https://www.rfc-editor.org/rfc/rfc8017 http://mpqs.free.fr/h11300-pkcs-1v2-2-rsa-cryptography-standard-wp_EMC_Corporation_Public-Key_Cryptography_Standards_(PKCS).pdf In this theory, we translate the functions defined in PKCS #1 v2.2 to Isabelle. We adhere as closely to the written standard as possible, including function and variable names. It should be painfully obvious, even to non-Isabelle users, that the definitions here match exactly those in the standard. Notes: - We will only consider the case of n = p*q, and ignore the "multi-prime" RSA (see section 3 of PKCS #1). That is, in the notation of the standard, we only consider u = 2. - The abbreviation CRT will refer to the Chinese Remainder Theorem. - The standard includes some error checking on input. We aim to keep the definitions of the cryptographic primitives to mathematical functions. In the case that the standard specifies an error case, we define a "foo_inputValid" Isabelle function to capture that requirement separate from the foo function definition. - PKCS #1 defines an octet as an 8-bit byte. We use the same terminology, and use "octets" to mean a string (list) of octets. Words.thy implements the conversions from non-negative integers to octet strings and back again, as defined in PCKS #1 v 2.2. The structure of this theory mimics exactly the structure of the standard. Definitions found in Section 3 (for example) of the standard are found in the section titled Section 3 in this theory. In addition to those definitions, each section contains a subsection that proves all the lemmas that one might care to know about them. For example, you may find there lemmas about converting between private keys for use in the CRT decryption method and the corresponding key for use without the CRT method. There are also many less lofty but practical lemmas. Note that Section 9 must precede Section 8 because of dependencies among their definitions. \ text \Some of the schemes describe in PKCS #1 assume the existence of a hash function or a mask generating function (mgf). For this translation, we only need to know the type of such functions. For certain lemmas we may need to assume things, such as the output of a hash function is always of a given length. Those assumptions will be made as they are needed. \ type_synonym mgf_type = "octets \ nat \ octets" type_synonym hash_type = "octets \ octets" section \Section 3: RSA Key Types\ named_theorems ValidKeyDefs subsection \Section 3.1: RSA Public Key\ text \The current standard uses Carmichael's totient, l = lcm((p-1),(q-1)). This is more efficient than using Euler's totient, (p-1)*(q-1).\ definition PKCS1_validRSApublicKey :: "nat \ nat \ nat \ nat \ bool" where [ValidKeyDefs]:"PKCS1_validRSApublicKey n e p q \ ( let l = (lcm (p-1) (q-1)) in (prime p) \ (2 < p) \ (prime q) \ (2 < q) \ (p\q) \ (n = p*q) \ (2 < e) \ (e < n) \ (gcd e l = 1) )" definition PKCS1_validRSApublicKeyPair :: "nat \ nat \ bool" where "PKCS1_validRSApublicKeyPair n e \ (\p q. PKCS1_validRSApublicKey n e p q)" subsection \Section 3.2: RSA Private Key\ definition PKCS1_validRSAprivateKey :: "nat \ nat \ nat \ nat \ nat \ bool" where [ValidKeyDefs]:"PKCS1_validRSAprivateKey n d p q e \ ( let l = (lcm (p-1) (q-1)) in (PKCS1_validRSApublicKey n e p q) \ (0 < d) \ (d < n) \ (e*d mod l = 1) )" definition PKCS1_validRSAprivateKey_CRT :: "nat \ nat \ nat \ nat \ nat \ nat \ bool" where [ValidKeyDefs]: "PKCS1_validRSAprivateKey_CRT p q dP dQ qInv e \ (PKCS1_validRSApublicKey (p*q) e p q) \ (0 < dP) \ (dP < p) \ (e*dP mod (p-1) = 1) \ (0 < dQ) \ (dQ < q) \ (e*dQ mod (q-1) = 1) \ (q*qInv mod p = 1)" subsection \Valid Key Lemmas\ named_theorems ValidKeyThms lemma validPublicKey_pq_symm: assumes "PKCS1_validRSApublicKey n e p q" shows "PKCS1_validRSApublicKey n e q p" by (metis assms PKCS1_validRSApublicKey_def lcm.commute mult.commute) lemma validPrivateKey_pq_symm: assumes "PKCS1_validRSAprivateKey n d p q e" shows "PKCS1_validRSAprivateKey n d q p e" by (metis assms ValidKeyDefs(1,2) lcm.commute mult.commute) lemma validPrivateKey_ed_symm: assumes "PKCS1_validRSAprivateKey n d p q e" "2 < d" shows "PKCS1_validRSAprivateKey n e p q d" proof - let ?l = "lcm (p-1) (q-1)" have n: "(prime p) \ (2 < p) \ (prime q) \ (2 < q) \ (p\q) \ (n = p*q)" using assms(1) ValidKeyDefs(1,2) by algebra have e1: "(2 < e) \ (e < n) \ (gcd e ?l = 1)" using assms(1) ValidKeyDefs(1,2) by metis have e0: "0 < e" using e1 by fastforce have d1: "(2 < d) \ (d < n) \ (e*d mod ?l = 1)" using assms ValidKeyDefs(2) by algebra have d2: "gcd d ?l = 1" by (metis d1 gcd.bottom_right_bottom gcd.left_commute gcd_0_nat gcd_red_nat mod_mult_self2_is_0) show ?thesis by (metis n e1 e0 d1 d2 ValidKeyDefs(1,2) mult.commute) qed lemma n_prime_factorization: assumes "prime (p::nat)" "prime q" "n = p*q" "p \ q" shows "prime_factors n = {p, q} \ multiplicity p n = 1 \ multiplicity q n = 1" proof - have 2: "0 < p \ 0 < q" using assms prime_gt_0_nat by presburger have 3: "multiplicity p n = 1 \ multiplicity q n = 1" by (metis assms 2 One_nat_def multiplicity_times_same neq0_conv not_prime_unit prime_multiplicity_other mult.commute) have 4: "(count (prime_factorization n) p = 1) \ (count (prime_factorization n) q = 1)" by (metis assms(1,2) 3 prime_factorization.rep_eq) have 5: "(p \ prime_factors n) \ (q \ prime_factors n)" by (metis 4 count_greater_zero_iff less_one) have 6: "prime_factors n \ {p,q}" by (metis assms in_prime_factors_iff insert_iff prime_dvd_mult_nat primes_dvd_imp_eq subsetI) have "prime_factors n = {p,q}" using 5 6 by blast then show ?thesis using 3 by blast qed lemma n_positive: assumes "prime (p::nat)" "prime q" "n = p*q" shows "0 < n" using assms prime_gt_0_nat by simp lemma n_positive2: assumes "PKCS1_validRSApublicKey n e p q" shows "0 < n" using assms PKCS1_validRSApublicKey_def n_positive by algebra lemma p_minus_1_gr_1: assumes "PKCS1_validRSApublicKey n e p q" shows "1 < p - 1" using assms PKCS1_validRSApublicKey_def by force lemma q_minus_1_gr_1: assumes "PKCS1_validRSApublicKey n e p q" shows "1 < q - 1" using assms PKCS1_validRSApublicKey_def by force lemma n_totient: assumes "prime (p::nat)" "prime q" "n = p*q" "p \ q" shows "totient n = (p-1)*(q-1)" by (metis assms nat_less_le prime_imp_coprime prime_nat_iff residues_prime.intro residues_prime.prime_totient_eq totient_mult_coprime) lemma n_odd: assumes "PKCS1_validRSApublicKey n e p q" shows "n mod 2 = 1" by (metis PKCS1_validRSApublicKey_def assms even_mult_iff parity_cases prime_odd_nat) lemma n_odd2: assumes "PKCS1_validRSApublicKeyPair n e" shows "n mod 2 = 1" using PKCS1_validRSApublicKeyPair_def assms n_odd by blast lemma l_dvd_totient_n: assumes "prime (p::nat)" "prime q" "n = p*q" "p \ q" "l = lcm (p-1) (q-1)" shows "l dvd totient n" using assms n_totient by force lemma l_less_n: assumes "prime (p::nat)" "prime q" "n = p*q" "p \ q" "l = lcm (p-1) (q-1)" shows "l < n" by (metis assms dual_order.strict_trans l_dvd_totient_n less_1_mult linorder_neqE_nat n_positive nat_dvd_not_less prime_gt_1_nat totient_gt_0_iff totient_less) lemma dP_from_d: assumes "PKCS1_validRSAprivateKey n d p q e" "dP = d mod (p-1)" shows "e*dP mod (p-1) = 1 \ 0 < dP \ dP < p" proof - have "2 < p" using assms(1) ValidKeyDefs(1,2) by algebra then have p1: "1 < p - 1" by linarith let ?l = "lcm (p-1) (q-1)" have "e*d mod ?l = 1" using assms(1) PKCS1_validRSAprivateKey_def by algebra then have a: "e*dP mod (p-1) = 1" by (metis assms(2) dvd_lcm1 mod_if mod_mod_cancel mod_mult_right_eq p1) have b: "dP < p" by (metis assms(2) p1 add_diff_inverse_nat bot_nat_0.not_eq_extremum le_simps(1) mod_less_divisor nat_diff_split not_one_le_zero trans_less_add2) - have c: "0 < dP" by (metis a bits_mod_0 mult_eq_0_iff neq0_conv zero_neq_one) + from a have c: "0 < dP" + by (cases \dP = 0\) simp_all show ?thesis using a b c by blast qed lemma dQ_from_d: assumes "PKCS1_validRSAprivateKey n d p q e" "dQ = d mod (q-1)" shows "(e*dQ mod (q-1) = 1) \ 0 < dQ \ dQ < q" using assms dP_from_d validPrivateKey_pq_symm by meson text\If you start with a valid private RSA key and you know the inverse of q (mod p), then it's easy to write down a valid private RSA key for decrypting with the Chinese Remainder Theorem.\ lemma validPrivateKey_to_validPrivateCRTkey [ValidKeyThms]: assumes "PKCS1_validRSAprivateKey n d p q e" "(q*qInv mod p = 1)" "dP = d mod (p-1)" "dQ = d mod (q-1)" shows "PKCS1_validRSAprivateKey_CRT p q dP dQ qInv e" by (metis assms dP_from_d dQ_from_d ValidKeyDefs) text\It is possible to go from a CRT private key to a "plain" private key, but the computation is more involved. The "bezw" function here is for "Bezout's theorem witness", meaning bezw(a,b) will return two integers (x,y) so that a*x + b*y = gcd a b.\ definition d_from_dP_dQ :: "nat \ nat \ nat \ nat \ nat" where [ValidKeyDefs]: "d_from_dP_dQ dP dQ p q = ( let pm1 = p-1; qm1 = q-1; g = gcd pm1 qm1; l = lcm pm1 qm1; u = fst (bezw pm1 qm1); v = snd (bezw pm1 qm1) in nat ( ((dP*v*qm1 + dQ*u*pm1) div g) mod l) )" lemma d_from_dP_dQ_h1: fixes pm1 qm1 :: nat and u v :: int assumes "g = gcd pm1 qm1" shows "g dvd (dP*v*qm1 + dQ*u*pm1)" by (simp add: assms) lemma d_from_dP_dQ_h2: assumes "d = d_from_dP_dQ dP dQ p q" "n = p*q" "PKCS1_validRSAprivateKey_CRT p q dP dQ qInv e" shows "d < n" by (smt (z3) pos_mod_bound ValidKeyDefs(1,3) assms d_from_dP_dQ_def int_nat_eq l_less_n lcm_int_int_eq lcm_pos_nat less_asym of_nat_0_less_iff of_nat_less_imp_less prime_gt_1_nat zero_less_diff) lemma d_from_dP_dQ_h3: assumes "PKCS1_validRSAprivateKey_CRT p q dP dQ qInv e" "d = d_from_dP_dQ dP dQ p q" "l = lcm (p-1) (q-1)" shows "e*d mod l = 1" proof - let ?pm1 = "p-1" let ?qm1 = "q-1" let ?g = "gcd ?pm1 ?qm1" let ?u = "fst (bezw ?pm1 ?qm1)" let ?v = "snd (bezw ?pm1 ?qm1)" have gl: "?pm1*?qm1 = ?g*l" using assms(3) prod_gcd_lcm_nat by presburger have g: "?u*?pm1 + ?v*?qm1 = ?g" using bezw_aux by presburger have e: "e*dP mod ?pm1 = 1" using assms(1) ValidKeyDefs(3) by algebra then obtain x where x: "e*dP = 1 + x*?pm1" by (metis add.commute div_mod_decomp) have "e*dQ mod ?qm1 = 1" using assms(1) ValidKeyDefs(3) by algebra then obtain y where y: "e*dQ = 1 + y*?qm1" by (metis add.commute div_mod_decomp) have m1: "1 < ?pm1 \ 1 < ?qm1" using assms(1) ValidKeyDefs(1,3) by auto have m2: "0 < ?pm1*?qm1" using m1 by force have m3: "0 < l" using assms(3) lcm_pos_nat m2 nat_0_less_mult_iff by algebra let ?dint = "( (dP *?v * ?qm1 + dQ * ?u * ?pm1) div ?g) mod l" have a0: "0 \ ?dint" using m3 Euclidean_Rings.pos_mod_sign by simp have a1: "d = nat ?dint" by (metis assms(2,3) d_from_dP_dQ_def gcd_int_int_eq lcm_int_int_eq) let ?eint = "int e" let ?eiml = "?eint mod l" have a2: "?eiml*?dint mod l = ?eint*((dP *?v * ?qm1 + dQ * ?u * ?pm1) div ?g) mod l" by (meson mod_mult_eq) have a3: "?eiml*?dint mod l = (?eint*(dP *?v * ?qm1 + dQ * ?u * ?pm1) div ?g) mod l" by (smt (verit, ccfv_threshold) a2 d_from_dP_dQ_h1 div_0 div_mult1_eq div_mult_mod_eq dvd_div_mult_self of_nat_0 right_diff_distrib) have a4: "?eiml*?dint mod l = ((?eint*dP *?v * ?qm1 + ?eint*dQ * ?u * ?pm1) div ?g) mod l" by (metis (no_types, opaque_lifting) a3 distrib_left mult.assoc) have a5: "?eiml*?dint mod l = (((1 + x*?pm1)*?v * ?qm1 + (1 + y*?qm1) * ?u * ?pm1) div ?g) mod l" by (metis (no_types, lifting) a4 x y of_nat_mult) have a6: "((1 + x*?pm1)*?v * ?qm1 + (1 + y*?qm1) * ?u * ?pm1) div ?g = (?v*?qm1 + x*?pm1*?v*?qm1 + ?u*?pm1 + y*?qm1*?u*?pm1) div ?g" by (simp add: comm_semiring_class.distrib) have a7: "((1 + x*?pm1)*?v * ?qm1 + (1 + y*?qm1) * ?u * ?pm1) div ?g = ((?v*?qm1 + ?u*?pm1) + (x*?pm1*?v*?qm1 + y*?qm1*?u*?pm1)) div ?g" using a6 by presburger have a8: "((1 + x*?pm1)*?v * ?qm1 + (1 + y*?qm1) * ?u * ?pm1) div ?g = ((?g) + (x*?pm1*?v*?qm1 + y*?qm1*?u*?pm1)) div ?g" by (metis a7 g add.commute) have a9: "((1 + x*?pm1)*?v * ?qm1 + (1 + y*?qm1) * ?u * ?pm1) div ?g = (?g + (x*?v*?pm1*?qm1 + y*?u*?pm1*?qm1)) div ?g" by (smt (verit, ccfv_threshold) a8 mult.left_commute mult_of_nat_commute of_nat_mult) have a10: "((1 + x*?pm1)*?v * ?qm1 + (1 + y*?qm1) * ?u * ?pm1) div ?g = (?g + (x*?v + y*?u)*?pm1*?qm1) div ?g" by (smt (verit) a9 distrib_left mult.assoc mult.commute of_nat_mult x y) have a11: "((1 + x*?pm1)*?v * ?qm1 + (1 + y*?qm1) * ?u * ?pm1) div ?g = (?g + (x*?v + y*?u)*?g*l) div ?g" by (metis (no_types, opaque_lifting) a10 gl mult.assoc of_nat_mult) have a12: "((1 + x*?pm1)*?v * ?qm1 + (1 + y*?qm1) * ?u * ?pm1) div ?g = (1*?g + (x*?v + y*?u)*l*?g) div ?g" using a11 mult.commute mult.left_commute by auto have a13: "(1*?g + (x*?v + y*?u)*l*?g) = (1 + (x*?v + y*?u)*l)*?g" by (smt (verit) semiring_normalization_rules(1) mult_cancel_right1 nat_mult_1) have a14: "(1*?g + (x*?v + y*?u)*l*?g) div ?g = 1 + (x*?v + y*?u)*l" by (metis a13 gl m2 mult.commute mult_zero_left nat_neq_iff nonzero_mult_div_cancel_left of_nat_eq_0_iff) have a15: "?eiml*?dint mod l = (1 + (x*?v + y*?u)*l) mod l" using a12 a14 a5 by presburger have a16: "?eiml*?dint mod l = 1 mod l" by (metis a15 mod_mult_self2 mult.commute of_nat_1 of_nat_mod) have a17: "1 mod l = 1" by (metis m1 assms(3) dvd_eq_mod_eq_0 dvd_lcm1 lcm_0_iff_nat less_one linorder_neqE_nat mod_less) show ?thesis by (smt (verit) a0 a1 a16 a17 int_nat_eq mod_mult_left_eq of_nat_eq_iff of_nat_mod of_nat_mult) qed lemma d_from_dP_dQ_h4: assumes "d = d_from_dP_dQ dP dQ p q" "PKCS1_validRSAprivateKey_CRT p q dP dQ qInv e" shows "0 < d" - by (metis assms d_from_dP_dQ_h3 bits_mod_0 mult_0_right neq0_conv zero_neq_one) + using d_from_dP_dQ_h3 [OF assms(2) assms(1)] + by (cases \d = 0\) simp_all text\As said above, we can convert a valid RSA private key for CRT decryption into a valid RSA private key for "plain" decryption.\ lemma validPrivateCRTKey_to_validPrivateKey [ValidKeyThms]: assumes "PKCS1_validRSAprivateKey_CRT p q dP dQ qInv e" "d = d_from_dP_dQ dP dQ p q" "n=p*q" shows "PKCS1_validRSAprivateKey n d p q e" by (metis ValidKeyDefs(2,3) assms d_from_dP_dQ_h2 d_from_dP_dQ_h3 d_from_dP_dQ_h4) text\The standard only insists that the decryption exponent d is < n. But it's only d's value modulo l that matters. If the standard insisted that d < l, then there would be a unique decryption exponent d for a fixed RSA public key (n,e,p,q). \ lemma d_unique_mod_l: assumes "PKCS1_validRSAprivateKey n d1 p q e" "PKCS1_validRSAprivateKey n d2 p q e" "l = (lcm (p-1) (q-1))" shows "[d1 = d2] (mod l)" proof - have d1: "[e*d1 = 1] (mod l)" by (metis assms(1,3) ValidKeyDefs(2) cong_def mod_mod_trivial) have d2: "[e*d2 = 1] (mod l)" by (metis assms(2,3) ValidKeyDefs(2) cong_def mod_mod_trivial) show ?thesis by (metis (mono_tags, opaque_lifting) d1 d2 cong_def mod_mult_right_eq mult.left_commute mult.right_neutral) qed lemma d_mod_l_valid: assumes "PKCS1_validRSAprivateKey n d1 p q e" "0 < d2" "d2 < n" "[d1=d2] (mod l)" "l = (lcm (p-1) (q-1))" shows "PKCS1_validRSAprivateKey n d2 p q e" by (metis (mono_tags, lifting) ValidKeyDefs(2) assms cong_def mod_mult_right_eq) text\In contrast, the standard requires that dP < p and dQ < q. Given that, we know that their values are unique for a fixed RSA public key (n,e,p,q)\ lemma dP_dQ_unique: assumes "PKCS1_validRSAprivateKey_CRT p q dP1 dQ1 qInv e" "PKCS1_validRSAprivateKey_CRT p q dP2 dQ2 qInv e" shows "dP1 = dP2 \ dQ1 = dQ2" proof - let ?l = "lcm (p-1) (q-1)" have e1: "gcd e ?l = 1" using assms(1) ValidKeyDefs(1,3) by algebra have ep1: "gcd e (p-1) = 1" by (metis e1 dvd_lcm1 gcd_lcm_distrib is_unit_gcd_iff) have ep2: "coprime e (p-1)" using ep1 by blast have eq1: "gcd e (q-1) = 1" by (metis e1 dvd_lcm1 gcd_lcm_distrib is_unit_gcd_iff lcm.commute) have eq2: "coprime e (q-1)" using eq1 by blast have p0: "0 < p-1" using assms(1) ValidKeyDefs(1,3) by force have p1: "e*dP1 mod (p-1) = 1" using assms(1) ValidKeyDefs(3) by blast have p2: "e*dP2 mod (p-1) = 1" using assms(2) ValidKeyDefs(3) by blast have p3: "e*dP1 mod (p-1) = e*dP2 mod (p-1)" using p1 p2 by presburger have p4: "dP1 mod (p-1) = dP2 mod (p-1)" by (meson p3 ep2 cong_mult_lcancel_nat cong_def) have p5: "dP1 < p \ dP2 < p" using assms(1,2) ValidKeyDefs(3) by blast have p6: "dP1 < p-1 \ dP2 < p-1" by (metis p1 p2 p5 Suc_leI diff_Suc_1 less_imp_Suc_add linorder_neqE_nat mod_mult_self2_is_0 not_le zero_neq_one) have p7: "dP1 mod (p-1) = dP1 \ dP2 mod (p-1) = dP2" using p0 p6 by force have p: "dP1 = dP2" using p4 p7 by argo have q0: "0 < q-1" using assms(1) ValidKeyDefs(1,3) by force have q1: "e*dQ1 mod (q-1) = 1" using assms(1) ValidKeyDefs(3) by blast have q2: "e*dQ2 mod (q-1) = 1" using assms(2) ValidKeyDefs(3) by blast have q3: "e*dQ1 mod (q-1) = e*dQ2 mod (q-1)" using q1 q2 by presburger have q4: "dQ1 mod (q-1) = dQ2 mod (q-1)" by (meson q3 eq2 cong_mult_lcancel_nat cong_def) have q5: "dQ1 < q \ dQ2 < q" using assms(1,2) ValidKeyDefs(3) by blast have q6: "dQ1 < q-1 \ dQ2 < q-1" by (metis q1 q2 q5 Suc_leI diff_Suc_1 less_imp_Suc_add linorder_neqE_nat mod_mult_self2_is_0 not_le zero_neq_one) have q7: "dQ1 mod (q-1) = dQ1 \ dQ2 mod (q-1) = dQ2" using q0 q6 by force have q: "dQ1 = dQ2" using q4 q7 by argo show ?thesis using p q by blast qed text \Finally, if you have a valid public key, there is a decryption exponent d that completes a valid private key. Of course in order to compute d, one must know the factorization of n = p*q.\ lemma ValidPublicKey_to_ValidPrivateKey: assumes "PKCS1_validRSApublicKey n e p q" shows "\d. PKCS1_validRSAprivateKey n d p q e" proof - let ?l = "lcm (p-1) (q-1)" have 1: "gcd e ?l = 1" using assms(1) ValidKeyDefs(1) by algebra let ?ei = "int e" let ?li = "int ?l" obtain x y where 2: "x*?ei + y*?li = 1" by (metis 1 bezw_aux of_nat_1) have 3: "x*?ei mod ?li = 1" by (smt (verit, ccfv_SIG) 2 assms int_ops(2) lcm_0_iff_nat lcm_eq_1_iff less_one q_minus_1_gr_1 linorder_neqE_nat mod_less mod_mult_self3 nat_dvd_not_less of_nat_mod p_minus_1_gr_1) let ?d = "nat (x mod ?li)" have 4: "?d * e mod ?l = 1" by (smt (z3) 1 2 3 Euclidean_Rings.pos_mod_sign gcd_0_nat int_nat_eq mod_mult_left_eq mult_cancel_right1 mult_eq_0_iff nat_int nat_one_as_int nat_times_as_int of_nat_le_0_iff of_nat_mod) have 5: "0 < ?d" by (metis 4 mod_0 mult_is_0 not_gr_zero one_neq_zero) have 6: "?d < n" by (smt (verit, ccfv_threshold) 4 5 Euclidean_Rings.pos_mod_bound ValidKeyDefs(1) assms gcd_0_nat l_less_n less_imp_of_nat_less mod_by_0 mult.right_neutral nat_less_iff of_nat_le_0_iff split_nat) show ?thesis by (metis 4 5 6 assms ValidKeyDefs(2) mult.commute) qed lemma ValidPublicKeyPair_to_ValidPrivateKey: assumes "PKCS1_validRSApublicKeyPair n e" shows "\p q d. PKCS1_validRSAprivateKey n d p q e" using ValidPublicKey_to_ValidPrivateKey assms PKCS1_validRSApublicKeyPair_def by fast section \Section 4: Data Conversion Primitives\ text \"Two data conversion primitives are employed in the schemes defined in this document: •I2OSP – Integer-to-Octet-String primitive •OS2IP – Octet-String-to-Integer primitive For the purposes of this document, and consistent with ASN.1 syntax, an octet string is an ordered sequence of octets (eight-bit bytes). The sequence is indexed from first (conventionally, leftmost) to last (rightmost). For purposes of conversion to and from integers, the first octet is considered the most significant in the following conversion primitives. Write the integer x in its unique xLen-digit representation in base 256: x = x_{xLen–1} 256^{xLen–1} + x_{xLen–2} 256^{xLen–2} + ... + x_1 256 + x_0, where 0 \ xi < 256 (note that one or more leading digits will be zero if x is less than 256^{xLen–1}). Let the octet X_i have the integer value x{xLen–i} for 1 \ i \xLen. Output the octet string X = X_1 X_2 ... X_{xLen}." Octets.thy is a separate theory that contains all the foundational lemmas related to these data conversion primitives. Update: Octets.thy has been generalized to Words.thy, to handle conversions from natural numbers to a list of w-bit values and back again. \ subsection \Section 4.1: Integer to Octet String\ definition PKCS1_I2OSP_inputValid :: "nat \ nat \ bool" where "PKCS1_I2OSP_inputValid x xLen = (x < 256^xLen)" definition PKCS1_I2OSP :: "nat \ nat \ octets" where "PKCS1_I2OSP x xLen = nat_to_octets_len x xLen" subsection \Section 4.2: Octet String to Integer\ definition PKCS1_OS2IP :: "octets \ nat" where "PKCS1_OS2IP X = octets_to_nat X" subsection \Data Conversion Lemmas\ lemma octet_len_I2OSP_inputValid1: assumes "(octet_length x \ xLen)" shows "(PKCS1_I2OSP_inputValid x xLen)" by (metis PKCS1_I2OSP_inputValid_def assms leD leI nat_lowbnd_word_len2 Twoto8 zero_less_numeral) lemma octet_len_I2OSP_inputValid2: assumes "(PKCS1_I2OSP_inputValid x xLen)" shows "(octet_length x \ xLen)" using PKCS1_I2OSP_inputValid_def assms nat_bnd_word_len2 Twoto8 zero_less_numeral by presburger lemma I2OSP_valid_len: assumes "PKCS1_I2OSP_inputValid x xLen" shows "length (PKCS1_I2OSP x xLen) = xLen" using assms PKCS1_I2OSP_inputValid_def PKCS1_I2OSP_def nat_to_words_len_upbnd Twoto8 zero_less_numeral by algebra lemma I2OSP_OS2IP: "PKCS1_OS2IP (PKCS1_I2OSP x xLen) = x" using PKCS1_I2OSP_inputValid_def PKCS1_I2OSP_def PKCS1_OS2IP_def nat_to_words_len_to_nat zero_less_numeral by force lemma I2OSP_octets_valid: "octets_valid (PKCS1_I2OSP x xLen)" using words_valid_def PKCS1_I2OSP_def nat_to_words_len_valid by presburger lemma OS2IP_I2OSP: assumes "octets_valid os" shows "PKCS1_I2OSP (PKCS1_OS2IP os) (length os) = os" by (metis assms PKCS1_I2OSP_def PKCS1_OS2IP_def words_valid_def words_to_nat_to_words_len) section \Section 5: Cryptographic Primitives\ subsection \Section 5.1: Encryption and Decryption Primitives\ named_theorems RSAprims subsubsection \Section 5.1.1: Encryption Primitives\ definition PKCS1_RSAEP_messageValid :: "nat \ nat \ bool" where [RSAprims]: "PKCS1_RSAEP_messageValid n m = (m < n)" definition PKCS1_RSAEP :: "nat \ nat \ nat \ nat" where [RSAprims]: "PKCS1_RSAEP n e m = (m ^ e) mod n" subsubsection \Section 5.1.2: Decryption Primitives\ abbreviation PKCS1_RSADP_ciphertextValid :: "nat \ nat \ bool" where "PKCS1_RSADP_ciphertextValid n c \ PKCS1_RSAEP_messageValid n c" abbreviation PKCS1_RSADP :: "nat \ nat \ nat \ nat" where "PKCS1_RSADP n d c \ PKCS1_RSAEP n d c" abbreviation PKCS1_RSADP_CRT_ciphertextValid :: "nat \ nat \ nat \ bool" where "PKCS1_RSADP_CRT_ciphertextValid p q c \ PKCS1_RSADP_ciphertextValid (p*q) c" definition PKCS1_RSADP_CRT :: "nat \ nat \ nat \ nat \ nat \ nat \ nat" where [RSAprims]: "PKCS1_RSADP_CRT p q dP dQ qInv c = (let m1 = int ((c ^ dP) mod p); m2 = int ((c ^ dQ) mod q); h = (m1 - m2)*qInv mod p in nat ( m2 + (q*h) ) )" subsubsection \RSA Encryption/Decryption Lemmas\ named_theorems RSAprimsThms lemma encryptValidCiphertext: assumes "0 < n" shows "PKCS1_RSADP_ciphertextValid n (PKCS1_RSAEP n e m)" using assms RSAprims(1,2) by auto lemma encryptValidCiphertext2: assumes "PKCS1_validRSApublicKey n e p q" shows "PKCS1_RSADP_ciphertextValid n (PKCS1_RSAEP n e m)" using assms RSAprims(1,2) n_positive2 by auto lemma encryptValidI2OSP: assumes "k = octet_length n" "c = PKCS1_RSAEP n e m" "0 < n" shows "PKCS1_I2OSP_inputValid c k" using PKCS1_RSAEP_messageValid_def assms encryptValidCiphertext less_imp_le_nat octet_len_I2OSP_inputValid1 word_len_mono by presburger lemma encryptValidI2OSP2: assumes "k = octet_length n" "c = PKCS1_RSAEP n e m" "PKCS1_validRSApublicKey n e p q" shows "PKCS1_I2OSP_inputValid c k" using assms RSAprims(1,2) n_positive2 encryptValidI2OSP by auto lemma decryptCRTmessageValid_h: fixes n p q :: nat and m1 m2 h m :: int assumes "0 < n" "n = p*q" "m1 = int ((c ^ dP) mod p)" "m2 = int ((c ^ dQ) mod q)" "h = (m1 - m2)*qInv mod p" "m = m2 + (q*h)" shows "m < n \ 0 \ m" proof - have p0: "0 < p" using assms(1,2) by simp have q0: "0 < q" using assms(1,2) by simp - have 1: "m2 \ q-1" by (simp add: assms(4) q0 pos_mod_bound Suc_leI of_nat_diff) + have 1: "m2 \ q-1" by (simp add: assms(4) q0 Suc_leI of_nat_diff) have 2: "h < p" by (metis assms(5) p0 pos_mod_bound of_nat_0_less_iff) have 20: "h \ p-1" using 2 by auto have 3: "m \ (q-1) + q*h" using 1 assms(6) by linarith have 4: "q*h \ q*(p-1)" using q0 20 by auto have 5: "m \ (q-1) + q*(p-1)" using 3 4 by linarith have 6: "q*(p-1) = q*p - q" by (simp add: diff_mult_distrib2) have 7: "m \ (q - 1) + (q*p - q)" using 5 6 by argo have 8: "m \ (q - q) + (q*p - 1)" using 7 by (metis (no_types, opaque_lifting) Nat.add_diff_assoc2 One_nat_def Suc_leI add.commute assms(1,2) mult.right_neutral mult_le_mono2 nat_0_less_mult_iff) have 9: "m \ 0 + q*p - 1" using 8 by fastforce have "m \ n - 1" using assms(2) 9 by (metis mult.commute plus_nat.add_0) then have a1: "m < n" using assms(1) by linarith have a2: "0 \ m" using p0 assms(4,5,6) by simp show ?thesis using a1 a2 by simp qed lemma decryptCRTmessageValid: assumes "0 < n" "n = p*q" shows "PKCS1_RSAEP_messageValid n (PKCS1_RSADP_CRT p q dP dQ qInv c)" proof - let ?m1 = "int ((c ^ dP) mod p)" let ?m2 = "int ((c ^ dQ) mod q)" let ?h = "(?m1 - ?m2)*qInv mod p" let ?m = "?m2 + (q*?h)" have a1: "?m < n \ 0 \ ?m" using assms decryptCRTmessageValid_h by blast have a2: "(nat ?m) < n" using a1 assms(1) by linarith have a3: "nat ?m = PKCS1_RSADP_CRT p q dP dQ qInv c" using PKCS1_RSADP_CRT_def by metis show ?thesis using a2 a3 PKCS1_RSAEP_messageValid_def by algebra qed lemma decryptCRTmessageValid2: assumes "PKCS1_validRSAprivateKey_CRT p q dP dQ qInv e" "n = p*q" shows "PKCS1_RSAEP_messageValid n (PKCS1_RSADP_CRT p q dP dQ qInv c)" by (metis assms decryptCRTmessageValid ValidKeyDefs(1,3) bot_nat_0.not_eq_extremum less_nat_zero_code) lemma RSAEP_RSADP_onePrime_p: assumes "PKCS1_validRSAprivateKey n d p q e" shows "[m^(e*d) = m] (mod p)" proof - let ?l = "lcm (p-1) (q-1)" obtain lp where lp: "lp*(p-1) = ?l" using dvd_div_mult_self by blast obtain lq where lq: "lq*(q-1) = ?l" using dvd_div_mult_self by blast have 1: "e*d mod ?l = 1" using assms PKCS1_validRSAprivateKey_def by algebra obtain c where c0: "e*d = 1 + c*?l" by (metis 1 Nat.add_0_right Suc_eq_plus1_left add_Suc_right div_mod_decomp) let ?cp = "c*lp" let ?cq = "c*lq" have c1: "?cp*(p-1) = c*?l" using lp by force have c2: "?cq*(q-1) = c*?l" using lq by force have p1: "e*d = 1 + ?cp*(p-1)" using c0 c1 by presburger show ?thesis proof (cases) assume "m mod p = 0" then show ?thesis by (metis cong_def add_gr_0 c0 dvd_imp_mod_0 dvd_power dvd_trans mod_0_imp_dvd zero_less_one) next assume "m mod p \ 0" then have p201: "coprime m p" by (meson assms ValidKeyDefs(1,2) coprime_commute dvd_imp_mod_0 prime_imp_coprime) have p202: "totient p = (p-1)" using ValidKeyDefs(1,2) assms totient_prime by algebra have p203: "[m^(p-1) = 1] (mod p)" by (metis p201 p202 euler_theorem) have "m^(e*d) = m^(1 + ?cp*(p-1))" using p1 by presburger then have "m^(e*d) = m * m ^(?cp*(p-1))" by force then have "m^(e*d) = m * (m^(p-1))^?cp" by (metis mult.commute power_mult) then have "[m^(e*d) = m * (1)^?cp] (mod p)" by (metis p203 cong_pow cong_scalar_left) then show ?thesis by force qed qed lemma RSAEP_RSADP_onePrime_q: assumes "PKCS1_validRSAprivateKey n d p q e" shows "[m^(e*d) = m] (mod q)" using assms validPrivateKey_pq_symm RSAEP_RSADP_onePrime_p by fast lemma simple_CRT: assumes "prime (p::nat)" "prime q" "p \ q" "[a = b] (mod p)" "[a = b] (mod q)" shows "[a = b] (mod p*q)" by (simp add: assms coprime_cong_mult_nat primes_coprime) lemma simple_CRT_int: assumes "prime (p::nat)" "prime q" "p \ q" "[(a::int) = (b::nat)] (mod p)" "[a = b] (mod q)" "0 \ a" shows "[a = b] (mod p*q)" by (metis assms cong_int_iff simple_CRT zero_le_imp_eq_int) text\The first of 4 "inverse" theorems. Here, if you start with a valid message and a valid key, if you encrypt the message and ("plainly") decrypt the resulting ciphertext, you will get back the original message.\ lemma RSAEP_RSADP [RSAprimsThms]: assumes "PKCS1_RSAEP_messageValid n m" "PKCS1_validRSAprivateKey n d p q e" shows "PKCS1_RSADP n d (PKCS1_RSAEP n e m) = m" proof - have p: "[m^(e*d) = m] (mod p)" using assms(2) RSAEP_RSADP_onePrime_p by auto have q: "[m^(e*d) = m] (mod q)" using assms(2) RSAEP_RSADP_onePrime_q by auto have n: "[m^(e*d) = m] (mod n)" by (metis p q assms(2) ValidKeyDefs(1,2) simple_CRT) show ?thesis by (metis n assms(1) RSAprims(1,2) cong_def mod_less power_mod power_mult) qed text\This is the second of the 4 inverse theorems. Here we start with a valid ciphertext and a valid key. If we decrypt the ciphertext and encrypt the resulting plaintext, we will get back the original ciphertext. This is a direct result of the 1st inverse theorem because for RSA, the plain decryption function is equal to the encryption function.\ lemma RSADP_RSAEP [RSAprimsThms]: assumes "PKCS1_RSADP_ciphertextValid n c" "PKCS1_validRSAprivateKey n d p q e" shows "PKCS1_RSAEP n e (PKCS1_RSADP n d c) = c" by (metis assms RSAEP_RSADP PKCS1_RSAEP_def mult.commute power_mod power_mult) text\As noted above, the value of the decryption exponent d is only unqiue modulo l. In the case that we have two valid RSA keys where only the ds are different, using either key will result in the same values when decrypting.\ lemma RSADP_equiv_d [RSAprimsThms]: assumes "PKCS1_validRSAprivateKey n d1 p q e" "PKCS1_validRSAprivateKey n d2 p q e" shows "PKCS1_RSADP n d1 c = PKCS1_RSADP n d2 c" by (smt (verit, ccfv_SIG) assms RSAprims(1,2) ValidKeyDefs(2) RSADP_RSAEP RSAEP_RSADP le_eq_less_or_eq less_le_trans mod_less_divisor power_mod) lemma fermat2: assumes "prime p" "\ p dvd a" shows "a^b mod p = a^(b mod (p-1)) mod p" proof - let ?pm1 = "p-1" let ?x = "(b div ?pm1)" let ?y = "(b mod ?pm1)" have 1: "b = ?x*?pm1 + ?y" by presburger have 2: "a^?pm1 mod p = 1" by (metis assms fermat_theorem cong_def mod_less prime_nat_iff) have 3: "a^b = ((a^?pm1)^?x)*a^?y" by (metis 1 mult.commute power_add power_mult) have 4: "a^b mod p = ((a^?pm1)^?x)*a^?y mod p" using 3 by presburger have 5: "a^b mod p = ((1)^?x)*a^?y mod p" by (metis 2 4 mod_mult_cong power_mod) have 6: "a^b mod p = (1)*a^?y mod p" using 5 by force then show ?thesis by fastforce qed text\In section 3 we saw how to convert a valid RSA key for "plain" decryption into a valid RSA key for CRT decryption. Here we show that those keys are equivalent (result in the same plaintext) when using the appropriate decryption primitive.\ lemma RSADP_CRT_equiv1 [RSAprimsThms]: assumes "PKCS1_validRSAprivateKey n d p q e" "(q*qInv mod p = 1)" "dP = d mod (p-1)" "dQ = d mod (q-1)" shows "PKCS1_RSADP n d c = PKCS1_RSADP_CRT p q dP dQ qInv c" proof - let ?m1 = "int ((c ^ dP) mod p)" let ?m2 = "int ((c ^ dQ) mod q)" let ?h = "((?m1 - ?m2)*qInv) mod p" let ?m = "?m2 + (q*?h)" have n0: "0 < n" using n_positive2 assms(1) ValidKeyDefs(1,2) by meson have n1: "n = p*q" using assms(1) ValidKeyDefs(1,2) by algebra have r0: "nat ?m = PKCS1_RSADP_CRT p q dP dQ qInv c" using PKCS1_RSADP_CRT_def by metis have r1: "0 \ ?m" using decryptCRTmessageValid_h n0 n1 by blast have r2: "?m < n" using decryptCRTmessageValid_h n0 n1 by blast have r3: "(nat ?m) < n" using r1 r2 by linarith have r4: "(nat ?m) mod n = nat ?m" using r1 r3 by force let ?a = "c^d mod n" have a0: "PKCS1_RSADP n d c = ?a" using PKCS1_RSAEP_def by meson have q1: "[?m = ?m2] (mod q)" using cong_def mod_mult_self2 by blast have q2: "[?m2 = ?a] (mod q)" proof - have 1: "q dvd n" using assms(1) ValidKeyDefs(1,2) by simp have 2: "?a mod q = c^d mod q" using 1 mod_mod_cancel by blast have 3: "prime q" using assms(1) ValidKeyDefs(1,2) by algebra have 4: "c^d mod q = c^(d mod (q-1)) mod q" by (metis 3 PKCS1_validRSAprivateKey_def assms(1) dQ_from_d dvd_imp_mod_0 dvd_power dvd_trans fermat2) have 5: "c^d mod q = c^dQ mod q" using 4 assms(4) by blast then show ?thesis by (metis (no_types, lifting) 2 5 cong_def cong_mod_left of_nat_mod) qed have q3: "[?m = ?a] (mod q)" using q1 q2 cong_trans by blast have p1: "[?m = ?m1] (mod p)" proof - have "[?m = ?m2 + q*((?m1 - ?m2)*qInv)] (mod p)" by (meson cong_add_lcancel cong_def mod_mult_right_eq) then have "[?m = ?m2 + (?m1-?m2)*q*qInv] (mod p)" by (simp add: mult.assoc mult.left_commute) then have 1: "[?m = ?m2 + (?m1-?m2)*(q*qInv)] (mod p)" by (simp add: mult.assoc) have "[(?m1-?m2)*(q*qInv) = (?m1-?m2)*1] (mod p)" by (metis assms(2) cong_mod_right cong_refl cong_scalar_right mult.commute of_nat_1 of_nat_mod) then have "[(?m1-?m2)*(q*qInv) = ?m1-?m2] (mod p)" by fastforce then have "[?m = ?m2 + ?m1 - ?m2] (mod p)" by (smt (verit) 1 cong_add_lcancel cong_def mod_mult_right_eq of_nat_1 of_nat_mod add_diff_cancel_left' diff_add_cancel) then show ?thesis by simp qed have p2: "[?m1 = ?a] (mod p)" by (smt (verit) ValidKeyDefs(1,2) assms(1,3) cong_def dP_from_d dvd_imp_mod_0 dvd_power dvd_trans dvd_triv_right fermat2 mod_mod_cancel mod_mod_trivial mult.commute of_nat_mod) have p3: "[?m = ?a] (mod p)" using p1 p2 cong_trans by blast have a2: "[?m = ?a] (mod n)" using r1 q3 p3 assms(1) ValidKeyDefs(1,2) simple_CRT_int by algebra have a3: "?a < n" by (metis ValidKeyDefs(2) assms(1) gr_implies_not_zero mod_less_divisor neq0_conv) have a4: "?a mod n = ?a" using a3 by auto have a5: "?a mod n = (nat ?m) mod n" by (metis r1 a2 cong_def nat_int nat_mod_distrib of_nat_0_le_iff r1) have a6: "?a = nat ?m" using a4 a5 r4 by argo show ?thesis using a6 r0 a0 by argo qed text\In section 3 we saw how to convert a valid RSA key for CRT decryption into a valid RSA key for "plain" decryption. Here we show that those keys are equivalent (result in the same plaintext) when using the appropriate decryption primitive.\ lemma RSADP_CRT_equiv2 [RSAprimsThms]: assumes "PKCS1_validRSAprivateKey_CRT p q dP dQ qInv e" "d = d_from_dP_dQ dP dQ p q" "n=p*q" shows "PKCS1_RSADP n d c = PKCS1_RSADP_CRT p q dP dQ qInv c" by (metis (no_types, opaque_lifting) PKCS1_validRSAprivateKey_CRT_def RSADP_CRT_equiv1 assms dP_dQ_unique validPrivateCRTKey_to_validPrivateKey validPrivateKey_to_validPrivateCRTkey) text\We end this section with the final 2 of the 4 inverse theorems. These show that the RSA decryption operation using the Chinese Remainder Theorem is an inverse function of the RSA encryption (in either order.)\ lemma RSAEP_RSADP_CRT [RSAprimsThms]: assumes "PKCS1_RSAEP_messageValid n m" "PKCS1_validRSAprivateKey_CRT p q dP dQ qInv e" "n=p*q" shows "PKCS1_RSADP_CRT p q dP dQ qInv (PKCS1_RSAEP n e m) = m" by (metis RSADP_CRT_equiv2 RSAEP_RSADP assms validPrivateCRTKey_to_validPrivateKey) lemma RSADP_CRT_RSAEP [RSAprimsThms]: assumes "PKCS1_RSADP_ciphertextValid n c" "PKCS1_validRSAprivateKey_CRT p q dP dQ qInv e" "n=p*q" shows "PKCS1_RSAEP n e (PKCS1_RSADP_CRT p q dP dQ qInv c) = c" by (metis RSADP_CRT_equiv2 RSADP_RSAEP assms validPrivateCRTKey_to_validPrivateKey) subsection \Section 5.2: RSA Signature/Verification Primitives\ text \"The main mathematical operation in each primitive is exponentiation, as in the encryption and decryption primitives of Section 5.1. RSASP1 and RSAVP1 are the same as RSADP and RSAEP except for the names of their input and output arguments; they are distinguished as they are intended for different purposes." Because the signature primitives are simply aliases of the encryption primitives, there are no lemmas in this subsection.\ subsubsection \Section 5.2.1: Signature Primitives\ abbreviation PKCS1_RSASP1 :: "nat \ nat \ nat \ nat" where "PKCS1_RSASP1 n d m \ PKCS1_RSADP n d m" abbreviation PKCS1_RSASP1_CRT :: "nat \ nat \ nat \ nat \ nat \ nat \ nat" where "PKCS1_RSASP1_CRT p q dP dQ qInv m \ PKCS1_RSADP_CRT p q dP dQ qInv m" subsubsection \Section 5.2.2: Verification Primitive\ abbreviation PKCS1_RSAVP1 :: "nat \ nat \ nat \ nat" where "PKCS1_RSAVP1 n e s \ PKCS1_RSAEP n e s" section \Section 7: Encryption Schemes\ text \"For the purposes of this document, an encryption scheme consists of an encryption operation and a decryption operation, where the encryption operation produces a ciphertext from a message with a recipient's RSA public key, and the decryption operation recovers the message from the ciphertext with the recipient's corresponding RSA private key. An encryption scheme can be employed in a variety of applications. A typical application is a key establishment protocol, where the message contains key material to be delivered confidentially from one party to another. For instance, PKCS #7 employs such a protocol to deliver a content- encryption key from a sender to a recipient; the encryption schemes defined here would be suitable key-encryption algorithms in that context. Two encryption schemes are specified in this document: RSAES-OAEP and RSAES-PKCS1-v1_5. RSAES-OAEP is required to be supported for new applications; RSAES-PKCS1-v1_5 is included only for compatibility with existing applications. The encryption schemes given here follow a general model similar to that employed in IEEE 1363-2000 combining encryption and decryption primitives with an encoding method for encryption. The encryption operations apply a message encoding operation to a message to produce an encoded message which is then converted to an integer message representative. An encryption primitive is applied to the message representative to produce the ciphertext. Reversing this, the decryption operations apply a decryption primitive to the ciphertext to recover a message representative, which is then converted to an octet string encoded message. A message decoding operation is applied to the encoded message to recover the message and verify the correctness of the decryption." \ subsection \Section 7.1.1: RSAES-OAEP\ text \In this locale, we fix a mask generating function MGF. We also fix a hash function, with with output length hLen. The only place that the hash function is used is to hash the (optional) label L to get lHash. When a label is not provided, a default lHash value is used. These default lHash values may be found below. Moreover, we abstract away the underlying RSA primitives for encryption and decryption. We do this to elide over the two options for decryption. We only need to know that the output of either encryption or decryption is less than the modulus n and that in either order, these two functions are inverses of each other. We have proven these facts above for both "plain" RSA decryption and decryption using the Chinese Remainder Theorem. \ locale OAEP = fixes MGF :: mgf_type and Hash :: hash_type and hLen :: nat and RSAENCRYPT RSADECRYPT :: "nat \ nat" and n :: nat assumes mgf_valid: "\x y. octets_valid (MGF x y)" and mgf_len: "\x y. length (MGF x y) = y" and hash_valid: "\x. octets_valid (Hash x)" and hash_len : "\x. length (Hash x) = hLen" and RSA_npos: "0 < n" and RSA_bnd: "\m. (RSAENCRYPT m) < n" "\c. (RSADECRYPT c) < n" and RSA_inv: "\m. (m < n) \ RSADECRYPT (RSAENCRYPT m) = m" "\c. (c < n) \ RSAENCRYPT (RSADECRYPT c) = c" begin named_theorems OAEPdefs named_theorems OAEPthms definition k where "k = octet_length n" subsubsection \Section 7.1.1: Encryption Operation\ text\ RSAES-OAEP-ENCRYPT ((n, e), M, L) "Options: Hash hash function (hLen denotes the length in octets of the hash function output) MGF mask generation function Input: (n, e) recipient's RSA public key (k denotes the length in octets of the RSA modulus n) M message to be encrypted, an octet string of length mLen, where mLen\k – 2hLen – 2 L optional label to be associated with the message; the default value for L, if L is not provided, is the empty string Output: C ciphertext, an octet string of length k" Notes: - The only place that Hash is used in RSAES_OAEP_Encrypt is to hash the label L to get lHash. When no label L is given there are default values for lHash. - Not listed as an input in the standard, the RSA-OAEP-ENCRYPT operation requires a random nonce called seed. From step d of the operations, the standard requires you to: "Generate a random octet string seed of length hLen." The seed could be the output of Hash, but the standard does not insist on that. In any case, we include seed in the list of inputs to RSAES-OAEP-ENCRYPT. \ definition PKCS1_OAEP_PS :: "nat \ octets" where "PKCS1_OAEP_PS mLen = replicate (k - mLen - (2*hLen) - 2) 0" definition PKCS1_OAEP_DB :: "octets \ octets \ octets \ octets" where "PKCS1_OAEP_DB lHash PS M = lHash @ PS @ [1] @ M" definition PKCS1_OAEP_dbMask :: "octets \ octets" where "PKCS1_OAEP_dbMask seed = MGF seed (k - hLen - 1)" definition PKCS1_OAEP_maskedDB :: "octets \ octets \ octets" where "PKCS1_OAEP_maskedDB DB dbMask = xor_octets DB dbMask" definition PKCS1_OAEP_seedMask :: "octets \ octets" where "PKCS1_OAEP_seedMask maskedDB = MGF maskedDB hLen" definition PKCS1_OAEP_maskedSeed :: "octets \ octets \ octets" where "PKCS1_OAEP_maskedSeed seed seedMask = xor_octets seed seedMask" definition PKCS1_OAEP_EM :: "octets \ octets \ octets" where "PKCS1_OAEP_EM maskedSeed maskedDB = [0] @ maskedSeed @ maskedDB" definition PKCS1_RSAES_OAEP_Encrypt_lengthValid :: "nat \ bool" where [OAEPdefs]: "PKCS1_RSAES_OAEP_Encrypt_lengthValid mLen \ (mLen + 2*hLen + 2 \ k)" definition PKCS1_RSAES_OAEP_Encrypt :: "octets \ octets \ octets \ octets" where [OAEPdefs]: "PKCS1_RSAES_OAEP_Encrypt M L seed = ( let mLen = length M; lHash = Hash L; PS = PKCS1_OAEP_PS mLen; DB = PKCS1_OAEP_DB lHash PS M; dbMask = PKCS1_OAEP_dbMask seed; maskedDB = PKCS1_OAEP_maskedDB DB dbMask; seedMask = PKCS1_OAEP_seedMask maskedDB; maskedSeed = PKCS1_OAEP_maskedSeed seed seedMask; EM = PKCS1_OAEP_EM maskedSeed maskedDB; m = PKCS1_OS2IP EM; c = RSAENCRYPT m in PKCS1_I2OSP c k )" text \We write down simple lemmas for each of the intermediate values in the OAEP encoding.\ lemma OAEP_PS_len: "length (PKCS1_OAEP_PS mLen) = k - mLen - 2*hLen - 2" using PKCS1_OAEP_PS_def by force lemma OAEP_PS_octetsValid: "octets_valid (PKCS1_OAEP_PS mLen)" using words_valid_def PKCS1_OAEP_PS_def by force lemma OAEP_DB_len: assumes "PKCS1_RSAES_OAEP_Encrypt_lengthValid mLen" "DB = PKCS1_OAEP_DB lHash PS M" "length lHash = hLen" "length PS = k - mLen - 2*hLen - 2" "length M = mLen" shows "length DB = k - hLen - 1" proof - have 1: "mLen + 2*hLen + 2 \ k" using assms(1) OAEPdefs(1) by fast have 2: "length DB = hLen + (k - mLen - 2*hLen - 2) + 1 + mLen" using PKCS1_OAEP_DB_def assms(2,3,4,5) by simp show ?thesis using 1 2 by force qed lemma OAEP_DB_octetsValid: assumes "DB = PKCS1_OAEP_DB lHash PS M" "octets_valid lHash" "octets_valid PS" "octets_valid M" shows "octets_valid DB" using assms PKCS1_OAEP_DB_def words_valid_concat words_valid_cons by fastforce lemma OAEP_dbMask_len: "length (PKCS1_OAEP_dbMask seed) = k - hLen - 1" using PKCS1_OAEP_dbMask_def mgf_len by simp lemma OAEP_dbMask_octetsValid: "octets_valid (PKCS1_OAEP_dbMask seed)" using PKCS1_OAEP_dbMask_def mgf_valid by presburger lemma OAEP_maskedDB_len: assumes "maskedDB = PKCS1_OAEP_maskedDB DB dbMask" "length DB = x" "length dbMask = x" shows "length maskedDB = x" by (simp add: PKCS1_OAEP_maskedDB_def assms xor_words_length) lemma OAEP_maskedDB_octetsValid: assumes "maskedDB = PKCS1_OAEP_maskedDB DB dbMask" "octets_valid DB" "octets_valid dbMask" shows "octets_valid maskedDB" by (simp add: assms xor_valid_words PKCS1_OAEP_maskedDB_def) lemma OAEP_seedMask_len: assumes "seedMask = PKCS1_OAEP_seedMask maskedDB" shows "length seedMask = hLen" using assms PKCS1_OAEP_seedMask_def mgf_len by auto lemma OAEP_seedMask_octetsValid: assumes "seedMask = PKCS1_OAEP_seedMask maskedDB" shows "octets_valid seedMask" using assms PKCS1_OAEP_seedMask_def mgf_valid by presburger lemma OAEP_maskedSeed_len: assumes "maskedSeed = PKCS1_OAEP_maskedSeed seed seedMask" "length seedMask = hLen" "length seed = hLen" shows "length maskedSeed = hLen" by (simp add: PKCS1_OAEP_maskedSeed_def assms xor_words_length) lemma OAEP_maskedSeed_octetsValid: assumes "maskedSeed = PKCS1_OAEP_maskedSeed seed seedMask" "octets_valid seedMask" "octets_valid seed" shows "octets_valid maskedSeed" using assms PKCS1_OAEP_maskedSeed_def xor_valid_words by presburger lemma OAEP_EM_octetsValid: assumes "EM = PKCS1_OAEP_EM maskedSeed maskedDB" "octets_valid maskedSeed" "octets_valid maskedDB" shows "octets_valid EM" using PKCS1_OAEP_EM_def assms words_valid_def by fastforce lemma OAEP_EM_len: assumes "PKCS1_RSAES_OAEP_Encrypt_lengthValid mLen" "EM = PKCS1_OAEP_EM maskedSeed maskedDB" "length maskedSeed = hLen" "length maskedDB = k - hLen - 1" shows "length EM = k" proof - have "mLen + 2*hLen + 2 \ k" using assms(1) OAEPdefs(1) by auto then have "hLen + 1 \ k" by linarith then show ?thesis using assms(2,3,4) PKCS1_OAEP_EM_def by force qed lemma OAEP_EM_strip0_len: assumes "EM = PKCS1_OAEP_EM maskedSeed maskedDB" "length EM = k" shows "length (words_strip_zero EM) \ k - 1" by (metis PKCS1_OAEP_EM_def append.simps(2) assms length_tl list.sel(1) words_strip_zero.simps(2) strip_zero_length) lemma OAEP_m_less_n: assumes "length (words_strip_zero EM) \ k - 1" "octets_valid EM" "m = PKCS1_OS2IP EM" shows "m < n" proof - have 1: "256^(k-1) \ n" by (metis RSA_npos k_def word_length_eq3 Twoto8 zero_less_numeral) have 2: "octets_to_nat EM < 256^(k-1)" by (metis assms(1,2) words_strip0_to_nat_len_bnd2 words_valid_def Twoto8) show ?thesis using 1 2 PKCS1_OS2IP_def assms(3) by force qed subsubsection \Section 7.1.2: Decryption Operation\ text\ "RSAES-OAEP-DECRYPT (K, C, L) Options: Hash hash function (hLen denotes the length in octets of the hash function output) MGF mask generation function Input: K recipient's RSA private key (k denotes the length in octets of the RSA modulus n) C ciphertext to be decrypted, an octet string of length k, where k\ 2hLen + 2 L optional label whose association with the message is to be verified; the default value for L, if L is not provided, is the empty string Output: M message, an octet string of length mLen, where mLen\k – 2hLen – 2" \ definition PKCS1_OAEP_decode_Y :: "octets \ nat" where "PKCS1_OAEP_decode_Y EM = hd EM" definition PKCS1_OAEP_decode_maskedSeed :: "octets \ octets" where "PKCS1_OAEP_decode_maskedSeed EM = take hLen (drop 1 EM)" definition PKCS1_OAEP_decode_maskedDB :: "octets \ octets" where "PKCS1_OAEP_decode_maskedDB EM = drop (1 + hLen) EM" definition PKCS1_OAEP_decode_seedMask :: "octets \ octets" where "PKCS1_OAEP_decode_seedMask maskedDB = MGF maskedDB hLen" definition PKCS1_OAEP_decode_seed :: "octets \ octets \ octets" where "PKCS1_OAEP_decode_seed maskedSeed seedMask = xor_octets maskedSeed seedMask" definition PKCS1_OAEP_decode_dbMask :: "octets \ octets" where "PKCS1_OAEP_decode_dbMask seed = MGF seed (k - hLen - 1)" definition PKCS1_OAEP_decode_DB :: "octets \ octets \ octets" where "PKCS1_OAEP_decode_DB maskedDB dbMask = xor_octets maskedDB dbMask" text \"Separate DB into an octet string lHash' of length hLen, a (possibly empty) padding string PS consisting of octets with hexadecimal value 0x00, and a message M as DB = lHash' || PS || 0x01 || M . If there is no octet with hexadecimal value 0x01 to separate PS from M, if lHash does not equal lHash', or if Y is nonzero, output ``decryption error'' and stop."\ definition PKCS1_OAEP_decode_lHash :: "octets \ octets" where "PKCS1_OAEP_decode_lHash DB = take hLen DB" definition PKCS1_OAEP_decode_validPS :: "octets \ bool" where "PKCS1_OAEP_decode_validPS DB \ (let PS_0x01_M = drop hLen DB; Ox01_M = words_strip_zero PS_0x01_M in (Ox01_M \ []) \ (Ox01_M ! 0 = 1) )" definition PKCS1_OAEP_decode_PSlen :: "octets \ nat" where "PKCS1_OAEP_decode_PSlen DB = (let PS_0x01_M = drop hLen DB; Ox01_M = words_strip_zero PS_0x01_M; tl = length PS_0x01_M; sl = length Ox01_M in tl - sl )" definition PKCS1_OAEP_decode_M :: "octets \ octets" where "PKCS1_OAEP_decode_M DB = drop 1 (words_strip_zero (drop hLen DB))" definition PKCS1_RSAES_OAEP_Decrypt_validInput :: "octets \ octets \ bool" where [OAEPdefs]: "PKCS1_RSAES_OAEP_Decrypt_validInput C lHash \ ( let c = PKCS1_OS2IP C; m = RSADECRYPT c; EM = PKCS1_I2OSP m k; Y = PKCS1_OAEP_decode_Y EM; maskedSeed = PKCS1_OAEP_decode_maskedSeed EM; maskedDB = PKCS1_OAEP_decode_maskedDB EM; seedMask = PKCS1_OAEP_decode_seedMask maskedDB; seed = PKCS1_OAEP_decode_seed maskedSeed seedMask; dbMask = PKCS1_OAEP_decode_dbMask seed; DB = PKCS1_OAEP_decode_DB maskedDB dbMask; lHash' = PKCS1_OAEP_decode_lHash DB in (length C = k) \ (2*hLen + 2 \ k) \ (c < n) \ (Y = 0) \ (lHash' = lHash) \ (PKCS1_OAEP_decode_validPS DB) )" definition PKCS1_RSAES_OAEP_Decrypt :: "octets \ octets" where [OAEPdefs]: "PKCS1_RSAES_OAEP_Decrypt C = ( let c = PKCS1_OS2IP C; m = RSADECRYPT c; EM = PKCS1_I2OSP m k; maskedSeed = PKCS1_OAEP_decode_maskedSeed EM; maskedDB = PKCS1_OAEP_decode_maskedDB EM; seedMask = PKCS1_OAEP_decode_seedMask maskedDB; seed = PKCS1_OAEP_decode_seed maskedSeed seedMask; dbMask = PKCS1_OAEP_decode_dbMask seed; DB = PKCS1_OAEP_decode_DB maskedDB dbMask in PKCS1_OAEP_decode_M DB )" text \At times it is convenient to have direct access to the seed and lHash decoded from a ciphertext. Also it is convenient to have direct access to the routine to check the padding string.\ definition PKCS1_RSAES_OAEP_Decrypt_seed :: "octets \ octets" where [OAEPdefs]: "PKCS1_RSAES_OAEP_Decrypt_seed C = ( let c = PKCS1_OS2IP C; m = RSADECRYPT c; EM = PKCS1_I2OSP m k; maskedSeed = PKCS1_OAEP_decode_maskedSeed EM; maskedDB = PKCS1_OAEP_decode_maskedDB EM; seedMask = PKCS1_OAEP_decode_seedMask maskedDB in PKCS1_OAEP_decode_seed maskedSeed seedMask )" definition PKCS1_RSAES_OAEP_Decrypt_lHash :: "octets \ octets" where [OAEPdefs]: "PKCS1_RSAES_OAEP_Decrypt_lHash C = ( let c = PKCS1_OS2IP C; m = RSADECRYPT c; EM = PKCS1_I2OSP m k; Y = PKCS1_OAEP_decode_Y EM; maskedSeed = PKCS1_OAEP_decode_maskedSeed EM; maskedDB = PKCS1_OAEP_decode_maskedDB EM; seedMask = PKCS1_OAEP_decode_seedMask maskedDB; seed = PKCS1_OAEP_decode_seed maskedSeed seedMask; dbMask = PKCS1_OAEP_decode_dbMask seed; DB = PKCS1_OAEP_decode_DB maskedDB dbMask in PKCS1_OAEP_decode_lHash DB )" definition PKCS1_RSAES_OAEP_Decrypt_validPS :: "octets \ bool" where [OAEPdefs]: "PKCS1_RSAES_OAEP_Decrypt_validPS C \ ( let c = PKCS1_OS2IP C; m = RSADECRYPT c; EM = PKCS1_I2OSP m k; maskedSeed = PKCS1_OAEP_decode_maskedSeed EM; maskedDB = PKCS1_OAEP_decode_maskedDB EM; seedMask = PKCS1_OAEP_decode_seedMask maskedDB; seed = PKCS1_OAEP_decode_seed maskedSeed seedMask; dbMask = PKCS1_OAEP_decode_dbMask seed; DB = PKCS1_OAEP_decode_DB maskedDB dbMask in (PKCS1_OAEP_decode_validPS DB) )" text \We write down simple lemmas for each of the intermediate values in the OAEP decoding.\ lemma OAEP_decode_EM_octetsValid: assumes "EM = PKCS1_I2OSP m k" shows "octets_valid EM" by (simp add: I2OSP_octets_valid assms) lemma OAEP_decode_EM_len: assumes "EM = PKCS1_I2OSP m k" "m < n" shows "length EM = k" by (metis I2OSP_valid_len PKCS1_I2OSP_inputValid_def k_def assms less_trans word_length_eq2 Twoto8 zero_less_numeral) lemma OAEP_decode_Y_valid: assumes "octets_valid EM" "Y = PKCS1_OAEP_decode_Y EM" "0 < length EM" shows "Y < 256" using assms words_valid_def PKCS1_OAEP_decode_Y_def by simp lemma OAEP_decode_maskedSeed_octetsValid: assumes "octets_valid EM" "maskedSeed = PKCS1_OAEP_decode_maskedSeed EM" shows "octets_valid maskedSeed" by (metis assms PKCS1_OAEP_decode_maskedSeed_def words_valid_drop words_valid_take) lemma OAEP_decode_maskedSeed_len: assumes "hLen + 1 \ length EM" "maskedSeed = PKCS1_OAEP_decode_maskedSeed EM" shows "length maskedSeed = hLen" using assms PKCS1_OAEP_decode_maskedSeed_def by auto lemma OAEP_decode_maskedDB_octetsValid: assumes "octets_valid EM" "maskedDB = PKCS1_OAEP_decode_maskedDB EM" shows "octets_valid maskedDB" by (metis assms PKCS1_OAEP_decode_maskedDB_def words_valid_drop) lemma OAEP_decode_maskedDB_len: assumes "maskedDB = PKCS1_OAEP_decode_maskedDB EM" shows "length maskedDB = (length EM) - 1 - hLen" using assms PKCS1_OAEP_decode_maskedDB_def by force lemma OAEP_decode_seedMask_octetsValid: assumes "seedMask = PKCS1_OAEP_decode_seedMask maskedDB" shows "octets_valid seedMask" using assms PKCS1_OAEP_decode_seedMask_def mgf_valid by simp lemma OAEP_decode_seedMask_len: assumes "seedMask = PKCS1_OAEP_decode_seedMask maskedDB" shows "length seedMask = hLen" using assms PKCS1_OAEP_decode_seedMask_def mgf_len by presburger lemma OAEP_decode_seed_octetsValid: assumes "seed = PKCS1_OAEP_decode_seed maskedSeed seedMask" "octets_valid maskedSeed" "octets_valid seedMask" shows "octets_valid seed" using assms PKCS1_OAEP_decode_seed_def xor_valid_words by simp lemma OAEP_decode_seed_len: assumes "seed = PKCS1_OAEP_decode_seed maskedSeed seedMask" "length maskedSeed = hLen" "length seedMask = hLen" shows "length seed = hLen" using assms PKCS1_OAEP_decode_seed_def xor_words_length by simp lemma OAEP_decode_dbMask_octetsValid: assumes "dbMask = PKCS1_OAEP_decode_dbMask seed" shows "octets_valid dbMask" using assms PKCS1_OAEP_decode_dbMask_def mgf_valid by presburger lemma OAEP_decode_dbMask_len: assumes "dbMask = PKCS1_OAEP_decode_dbMask seed" shows "length dbMask = k - hLen - 1" using assms PKCS1_OAEP_decode_dbMask_def mgf_len by presburger lemma OAEP_decode_DB_octetsValid: assumes "DB = PKCS1_OAEP_decode_DB maskedDB dbMask" "octets_valid maskedDB" "octets_valid dbMask" shows "octets_valid DB" using assms PKCS1_OAEP_decode_DB_def xor_valid_words by auto lemma OAEP_decode_DB_len: assumes "DB = PKCS1_OAEP_decode_DB maskedDB dbMask" "length maskedDB = x" "length dbMask = x" shows "length DB = x" using assms PKCS1_OAEP_decode_DB_def xor_words_length by auto lemma OAEP_decode_lHash_octetsValid: assumes "lHash' = PKCS1_OAEP_decode_lHash DB" "octets_valid DB" shows "octets_valid lHash'" by (metis assms PKCS1_OAEP_decode_lHash_def in_set_takeD words_valid_def) lemma OAEP_decode_lHash_len: assumes "lHash' = PKCS1_OAEP_decode_lHash DB" "hLen \ length DB" shows "length lHash' = hLen" using assms PKCS1_OAEP_decode_lHash_def by auto lemma OAEP_decode_PSlen_pos: assumes "PSlen = PKCS1_OAEP_decode_PSlen DB" shows "0 \ PSlen" using assms PKCS1_OAEP_decode_PSlen_def by blast lemma OAEP_decode_validPS_DBlen: assumes "PKCS1_OAEP_decode_validPS DB" "PSlen = PKCS1_OAEP_decode_PSlen DB" shows "hLen + PSlen + 1 \ length DB" by (metis assms PKCS1_OAEP_decode_PSlen_def PKCS1_OAEP_decode_validPS_def Suc_eq_plus1 Suc_leI add.commute diff_is_0_eq' drop_drop length_0_conv length_drop linorder_not_less strip_zero_drop) lemma OAEP_decode_validPS_DB_PS: assumes "PKCS1_OAEP_decode_validPS DB" "PSlen = PKCS1_OAEP_decode_PSlen DB" shows "take PSlen (drop hLen DB) = replicate PSlen 0" by (metis PKCS1_OAEP_decode_PSlen_def append_eq_conv_conj assms(2) length_replicate strip_zero_concat) lemma OAEP_decode_inputValid_DB: assumes "PKCS1_RSAES_OAEP_Decrypt_validInput C lHash" "c = PKCS1_OS2IP C" "m = RSADECRYPT c" "EM = PKCS1_I2OSP m k" "maskedSeed = PKCS1_OAEP_decode_maskedSeed EM" "maskedDB = PKCS1_OAEP_decode_maskedDB EM" "seedMask = PKCS1_OAEP_decode_seedMask maskedDB" "seed = PKCS1_OAEP_decode_seed maskedSeed seedMask" "dbMask = PKCS1_OAEP_decode_dbMask seed" "DB = PKCS1_OAEP_decode_DB maskedDB dbMask" "lHash' = PKCS1_OAEP_decode_lHash DB" "M = PKCS1_OAEP_decode_M DB" "PSlen = PKCS1_OAEP_decode_PSlen DB" "PS = replicate PSlen 0" shows "DB = lHash @ PS @ [1] @ M" proof - have 0: "lHash = lHash'" by (metis OAEPdefs(3) assms(1,2,3,4,5,6,7,8,9,10,11)) have 1: "take hLen DB = lHash" using 0 PKCS1_OAEP_decode_lHash_def assms(11) by presburger have 2: "(drop (hLen+PSlen) DB) ! 0 = 1" by (metis assms(1,2,3,4,5,6,7,8,9,10,13) drop_drop strip_zero_drop PKCS1_OAEP_decode_PSlen_def PKCS1_OAEP_decode_validPS_def OAEPdefs(3) add.commute) have 3: "drop (hLen+PSlen+1) DB = M" by (metis PKCS1_OAEP_decode_M_def PKCS1_OAEP_decode_PSlen_def add.commute assms(12,13) drop_drop strip_zero_drop) have 4: "take PSlen (drop hLen DB) = PS" by (metis PKCS1_OAEP_decode_PSlen_def append_eq_conv_conj assms(13,14) length_replicate strip_zero_concat) show ?thesis by (metis Cons_nth_drop_Suc One_nat_def PKCS1_OAEP_decode_PSlen_def append_Cons append_Nil PKCS1_OAEP_decode_validPS_def OAEPdefs(3) add.commute assms(1,2,3,4,5,6,7,8,9,10,13) append_take_drop_id drop0 drop_drop length_greater_0_conv strip_zero_drop 1 2 3 4) qed lemma OAEP_decode_M_octetsValid: assumes "M = PKCS1_OAEP_decode_M DB" "octets_valid DB" shows "octets_valid M" by (metis assms PKCS1_OAEP_decode_M_def words_valid_def in_set_dropD words_to_nat_to_strip_words words_upper_bound) lemma OAEP_decode_M_len: assumes "M = PKCS1_OAEP_decode_M DB" "PSlen = PKCS1_OAEP_decode_PSlen DB" shows "length M = (length DB) - hLen - PSlen - 1" by (metis assms PKCS1_OAEP_decode_M_def PKCS1_OAEP_decode_PSlen_def length_drop strip_zero_drop) lemma OAEP_decode_M_drop: assumes "M = PKCS1_OAEP_decode_M DB" "PSlen = PKCS1_OAEP_decode_PSlen DB" shows "M = drop (hLen + PSlen + 1) DB" by (metis (no_types) PKCS1_OAEP_decode_M_def PKCS1_OAEP_decode_PSlen_def add.commute assms(1,2) drop_drop strip_zero_drop) subsubsection \RSAES-OAEP Lemmas\ text \In the following lemma, we put together all the facts that we know about the intermediate values when encrypting a message M with RSA-OAEP.\ lemma OAEP_Encrypt_IntVals: assumes "PKCS1_RSAES_OAEP_Encrypt_lengthValid mLen" "lHash = Hash L" "PS = PKCS1_OAEP_PS mLen" "DB = PKCS1_OAEP_DB lHash PS M" "dbMask = PKCS1_OAEP_dbMask seed" "maskedDB = PKCS1_OAEP_maskedDB DB dbMask" "seedMask = PKCS1_OAEP_seedMask maskedDB" "maskedSeed = PKCS1_OAEP_maskedSeed seed seedMask" "EM = PKCS1_OAEP_EM maskedSeed maskedDB" "m = PKCS1_OS2IP EM" "c = RSAENCRYPT m" "C = PKCS1_I2OSP c k" "length M = mLen" "octets_valid M" "length seed = hLen" "octets_valid seed" shows "length lHash = hLen \ octets_valid lHash \ length PS = k - mLen - 2*hLen - 2 \ octets_valid PS \ length DB = k - hLen - 1 \ octets_valid DB \ length dbMask = k - hLen - 1 \ octets_valid dbMask \ length maskedDB = k - hLen - 1 \ octets_valid maskedDB \ length seedMask = hLen \ octets_valid seedMask \ length maskedSeed = hLen \ octets_valid maskedSeed \ length EM = k \ octets_valid EM \ m < n \ c < n \ length C = k \ octets_valid C \ C = PKCS1_RSAES_OAEP_Encrypt M L seed \ PKCS1_OAEP_decode_Y EM = 0 \ PKCS1_OAEP_decode_M DB = M \ lHash = PKCS1_OAEP_decode_lHash DB \ PKCS1_OAEP_decode_validPS DB" proof - have lH1: "length lHash = hLen" using assms(2) hash_len by blast have lH2: "octets_valid lHash" using assms(2) hash_valid by fast have ps1: "length PS = k - mLen - 2*hLen - 2" using assms(3) OAEP_PS_len by blast have ps2: "octets_valid PS" using assms(3) OAEP_PS_octetsValid by fast have db1: "length DB = k - hLen - 1" using assms(1,4,13) lH1 ps1 OAEP_DB_len by algebra have db2: "octets_valid DB" by (simp add: OAEP_DB_octetsValid assms(4,14) lH2 ps2) have db3: "lHash = PKCS1_OAEP_decode_lHash DB" by (simp add: PKCS1_OAEP_DB_def PKCS1_OAEP_decode_lHash_def assms(4) lH1) have db40: "drop hLen DB = PS@[1::nat]@M" by (simp add: PKCS1_OAEP_DB_def assms(4) lH1) have db41: "words_strip_zero (drop hLen DB) = [1]@M" by (simp add: assms(3) db40 PKCS1_OAEP_PS_def strip_prepend_zeros) have db4: "PKCS1_OAEP_decode_validPS DB" using db41 PKCS1_OAEP_decode_validPS_def by force have db5: "PKCS1_OAEP_decode_M DB = M" by (simp add: PKCS1_OAEP_decode_M_def db41) have dbm1: "length dbMask = k - hLen - 1" using assms(5) OAEP_dbMask_len by fast have dbm2: "octets_valid dbMask" using assms(5) OAEP_dbMask_octetsValid by simp have mdb1: "length maskedDB = k - hLen - 1" using assms(6) OAEP_maskedDB_len db1 dbm1 by blast have mdb2: "octets_valid maskedDB" using assms(6) OAEP_maskedDB_octetsValid db2 dbm2 by blast have sm1: "length seedMask = hLen" using assms(7) OAEP_seedMask_len by blast have sm2: "octets_valid seedMask" using assms(7) OAEP_seedMask_octetsValid by blast have ms1: "length maskedSeed = hLen" using assms(8,15) sm1 OAEP_maskedSeed_len by blast have ms2: "octets_valid maskedSeed" using assms(8,16) sm2 OAEP_maskedSeed_octetsValid by blast have em1: "length EM = k" using assms(1,9) ms1 mdb1 OAEP_EM_len by blast have em2: "octets_valid EM" using assms(9) ms2 mdb2 OAEP_EM_octetsValid by simp have em3: "length (words_strip_zero EM) \ k - 1" using assms(9) em1 OAEP_EM_strip0_len by fast have em4: "PKCS1_OAEP_decode_Y EM = 0" by (simp add: assms(9) PKCS1_OAEP_EM_def PKCS1_OAEP_decode_Y_def) have m1: "m < n" using em3 em2 assms(10,14,15) OAEP_m_less_n PKCS1_RSAEP_messageValid_def by algebra have c1: "c < n" using RSA_bnd(1) assms(11) by simp have C1: "length C = k" using OAEP_decode_EM_len k_def assms(12) c1 by blast have C2: "octets_valid C" by (simp add: I2OSP_octets_valid assms(12)) have C3: "C = PKCS1_RSAES_OAEP_Encrypt M L seed" using OAEPdefs(2) assms(2,3,4,5,6,7,8,9,10,11,12,13,14) by metis show ?thesis using lH1 lH2 ps1 ps2 db1 db2 db3 db4 db5 dbm1 dbm2 mdb1 mdb2 sm1 sm2 ms1 ms2 em1 em2 em3 em4 m1 c1 C1 C2 C3 by blast qed text \Assuming you have valid inputs, this lemma shows that the intermediate values computed while encrypting a message with RSA-OAEP match the corresponding intermediate values computed while decrypting that ciphertext.\ lemma OAEP_EncryptDecrypt_IntValsMatch: assumes "PKCS1_RSAES_OAEP_Encrypt_lengthValid mLen" "octets_valid M" "length seed = hLen" "octets_valid seed" "mLen = length M" "lHash = Hash L" "PS = PKCS1_OAEP_PS mLen" "DB = PKCS1_OAEP_DB lHash PS M" "dbMask = PKCS1_OAEP_dbMask seed" "maskedDB = PKCS1_OAEP_maskedDB DB dbMask" "seedMask = PKCS1_OAEP_seedMask maskedDB" "maskedSeed = PKCS1_OAEP_maskedSeed seed seedMask" "EM = PKCS1_OAEP_EM maskedSeed maskedDB" "m = PKCS1_OS2IP EM" "c = RSAENCRYPT m" "C = PKCS1_I2OSP c k" "c' = PKCS1_OS2IP C" "m' = RSADECRYPT c'" "EM' = PKCS1_I2OSP m' k" "Y' = PKCS1_OAEP_decode_Y EM'" "maskedSeed' = PKCS1_OAEP_decode_maskedSeed EM'" "maskedDB' = PKCS1_OAEP_decode_maskedDB EM'" "seedMask' = PKCS1_OAEP_decode_seedMask maskedDB'" "seed' = PKCS1_OAEP_decode_seed maskedSeed' seedMask'" "dbMask' = PKCS1_OAEP_decode_dbMask seed'" "DB' = PKCS1_OAEP_decode_DB maskedDB' dbMask'" "M' = PKCS1_OAEP_decode_M DB'" "lHash' = PKCS1_OAEP_decode_lHash DB'" shows "DB = DB' \ dbMask = dbMask' \ maskedDB = maskedDB' \ seedMask = seedMask' \ Y' = 0 \ maskedSeed = maskedSeed' \ EM = EM' \ m = m' \ c = c' \ seed = seed' \ M = M' \ lHash' = lHash \ C = PKCS1_RSAES_OAEP_Encrypt M L seed \ length C = k \ M' = PKCS1_RSAES_OAEP_Decrypt C \ PKCS1_OAEP_decode_validPS DB" proof - have c: "c = c'" by (metis I2OSP_OS2IP assms(16,17)) have m1: "m < n" using OAEP_Encrypt_IntVals assms(1,2,3,4,5,6,7,8,9,10,11,12,13,14,15) by blast have m: "m = m'" using RSA_inv(1) c m1 assms(15,18) by metis have EM1: "length EM = k" using OAEP_Encrypt_IntVals assms(1,2,3,4,5,6,7,8,9,10,11,12,13) by blast have EM2: "octets_valid EM" using OAEP_Encrypt_IntVals assms(1,2,3,4,5,6,7,8,9,10,11,12,13) by blast have EM: "EM = EM'" using m EM1 EM2 OS2IP_I2OSP assms(14,19) by force have Y: "Y' = 0" using EM OAEP_Encrypt_IntVals assms(1,2,3,4,5,6,7,8,9,10,11,12,13,20) by force have mS1: "length maskedSeed = hLen" using OAEP_maskedSeed_len PKCS1_OAEP_seedMask_def assms(3,11,12) mgf_len by presburger have mS: "maskedSeed = maskedSeed'" by (metis EM One_nat_def PKCS1_OAEP_EM_def PKCS1_OAEP_decode_maskedSeed_def Suc_eq_plus1 append_eq_conv_conj assms(13,21) list.size(3) list.size(4) mS1) have mDB: "maskedDB = maskedDB'" by (metis assms(13,22) EM PKCS1_OAEP_decode_maskedDB_def PKCS1_OAEP_EM_def Suc_eq_plus1 add.commute append.simps(1) append.simps(2) cancel_comm_monoid_add_class.diff_cancel drop0 drop_Suc_Cons drop_all drop_append mS1 order_refl) have sM: "seedMask = seedMask'" using assms(11,23) mDB PKCS1_OAEP_decode_seedMask_def PKCS1_OAEP_seedMask_def by presburger have S: "seed = seed'" by (metis assms(3,11,12,24) mS sM PKCS1_OAEP_decode_seed_def PKCS1_OAEP_maskedSeed_def PKCS1_OAEP_seedMask_def words_xor_inv xor_words_symm mgf_len) have dbM: "dbMask = dbMask'" using assms(9,15,25) S PKCS1_OAEP_dbMask_def PKCS1_OAEP_decode_dbMask_def by presburger have DB1: "DB = DB'" by (metis assms(1,2,3,4,5,6,7,8,9,10,26) dbM mDB mgf_len PKCS1_OAEP_maskedDB_def words_xor_inv2 PKCS1_OAEP_dbMask_def PKCS1_OAEP_decode_DB_def OAEP_Encrypt_IntVals) have DB2: "PKCS1_OAEP_decode_validPS DB" by (metis OAEP_Encrypt_IntVals assms(1,2,3,4,5,6,7,8)) have M: "M = M'" by (metis DB1 OAEP_Encrypt_IntVals assms(1,2,3,4,5,6,7,8,27)) have lH: "lHash' = lHash" using DB1 PKCS1_OAEP_DB_def PKCS1_OAEP_decode_lHash_def assms(6,8,28) hash_len by force have C1: "C = PKCS1_RSAES_OAEP_Encrypt M L seed" by (metis OAEPdefs(2) assms(5,6,7,8,9,10,11,12,13,14,15,16)) have C2: "length C = k" by (simp add: OAEP_decode_EM_len RSA_bnd(1) assms(15,16)) have "M' = PKCS1_RSAES_OAEP_Decrypt C" by (metis OAEPdefs(4) assms(17,18,19,21,22,23,24,25,26,27)) then show ?thesis using c m EM Y mS mDB sM S dbM DB1 DB2 M lH C1 C2 by blast qed lemma OAEP_Encrypt_validCipher_h1: assumes "C = PKCS1_RSAES_OAEP_Encrypt M L seed" shows "length C = k" by (metis OAEPdefs(2) assms OAEP_decode_EM_len RSA_bnd(1)) lemma OAEP_Encrypt_validCipher_h2: assumes "PKCS1_RSAES_OAEP_Encrypt_lengthValid mLen" shows "2*hLen + 2 \ k" by (meson Nat.le_diff_conv2 OAEPdefs(1) add_leD2 assms) lemma OAEP_Encrypt_validCipher_h3: assumes "C = PKCS1_RSAES_OAEP_Encrypt M L seed" "c = PKCS1_OS2IP C" "m = RSADECRYPT c" "EM = PKCS1_I2OSP m k" "Y = PKCS1_OAEP_decode_Y EM" "maskedSeed = PKCS1_OAEP_decode_maskedSeed EM" "maskedDB = PKCS1_OAEP_decode_maskedDB EM" "seedMask = PKCS1_OAEP_decode_seedMask maskedDB" "seed' = PKCS1_OAEP_decode_seed maskedSeed seedMask" "dbMask = PKCS1_OAEP_decode_dbMask seed'" "DB = PKCS1_OAEP_decode_DB maskedDB dbMask" "lHash' = PKCS1_OAEP_decode_lHash DB" "mLen = length M" "octets_valid M" "PKCS1_RSAES_OAEP_Encrypt_lengthValid mLen" "length seed = hLen" "octets_valid seed" "lHash = Hash L" shows "Y = 0 \ (lHash' = lHash) \ (PKCS1_OAEP_decode_validPS DB)" proof - let ?PS = "PKCS1_OAEP_PS mLen" let ?DB = "PKCS1_OAEP_DB lHash ?PS M" let ?dbMask = "PKCS1_OAEP_dbMask seed" let ?maskedDB = "PKCS1_OAEP_maskedDB ?DB ?dbMask" let ?seedMask = "PKCS1_OAEP_seedMask ?maskedDB" let ?maskedSeed = "PKCS1_OAEP_maskedSeed seed ?seedMask" let ?EM = "PKCS1_OAEP_EM ?maskedSeed ?maskedDB" let ?m = "PKCS1_OS2IP ?EM" let ?c = "RSAENCRYPT ?m" have "?DB = DB \ Y = 0 \ lHash' = lHash" using assms OAEPdefs(2) OAEP_EncryptDecrypt_IntValsMatch[of mLen M seed lHash L ?PS ?DB ?dbMask ?maskedDB ?seedMask ?maskedSeed ?EM ?m ?c C c m EM Y maskedSeed maskedDB seedMask seed' dbMask DB _ lHash'] by metis then show ?thesis using assms OAEP_Encrypt_IntVals by metis qed text\If you start with a valid message, when you encode and encrypt it with RSA-OAEP, the resulting ciphertext is a valid input for the RSA-OAEP decryption scheme.\ lemma OAEP_Encrypt_validCipher [OAEPthms]: assumes "PKCS1_RSAES_OAEP_Encrypt_lengthValid (length M)" "C = PKCS1_RSAES_OAEP_Encrypt M L seed" "octets_valid M" "octets_valid seed" "length seed = hLen" shows "PKCS1_RSAES_OAEP_Decrypt_validInput C (Hash L)" proof - let ?lHash = "Hash L" let ?c = "PKCS1_OS2IP C" let ?m = "RSADECRYPT ?c" let ?EM = "PKCS1_I2OSP ?m k" let ?Y = "PKCS1_OAEP_decode_Y ?EM" let ?maskedSeed = "PKCS1_OAEP_decode_maskedSeed ?EM" let ?maskedDB = "PKCS1_OAEP_decode_maskedDB ?EM" let ?seedMask = "PKCS1_OAEP_decode_seedMask ?maskedDB" let ?seed = "PKCS1_OAEP_decode_seed ?maskedSeed ?seedMask" let ?dbMask = "PKCS1_OAEP_decode_dbMask ?seed" let ?DB = "PKCS1_OAEP_decode_DB ?maskedDB ?dbMask" let ?lHash' = "PKCS1_OAEP_decode_lHash ?DB" have a1: "length C = k" using assms(2) OAEP_Encrypt_validCipher_h1 by blast have a2: "2*hLen + 2 \ k" using assms(1) OAEP_Encrypt_validCipher_h2 by blast have a3: "?c < n" by (metis assms(2) I2OSP_OS2IP OAEPdefs(2) RSA_bnd(1)) have a4: "(?Y = 0) \ (?lHash' = ?lHash) \ (PKCS1_OAEP_decode_validPS ?DB)" using assms OAEP_Encrypt_validCipher_h3 by force show ?thesis using a1 a2 a3 a4 OAEPdefs(3) by presburger qed text\Start with a valid message and seed. Encrypt the message with RSA-OAEP. Decrypt the resulting ciphertext with RSA-OAEP. The result will be the original message.\ lemma OAEP_Encrypt_Decrypt [OAEPthms]: assumes "PKCS1_RSAES_OAEP_Encrypt_lengthValid (length M)" "C = PKCS1_RSAES_OAEP_Encrypt M L seed" "octets_valid M" "octets_valid seed" "length seed = hLen" shows "PKCS1_RSAES_OAEP_Decrypt C = M" proof - let ?c' = "PKCS1_OS2IP C" let ?m' = "RSADECRYPT ?c'" let ?EM' = "PKCS1_I2OSP ?m' k" let ?maskedSeed' = "PKCS1_OAEP_decode_maskedSeed ?EM'" let ?maskedDB' = "PKCS1_OAEP_decode_maskedDB ?EM'" let ?seedMask' = "PKCS1_OAEP_decode_seedMask ?maskedDB'" let ?seed' = "PKCS1_OAEP_decode_seed ?maskedSeed' ?seedMask'" let ?dbMask' = "PKCS1_OAEP_decode_dbMask ?seed'" let ?DB' = "PKCS1_OAEP_decode_DB ?maskedDB' ?dbMask'" let ?M' = "PKCS1_OAEP_decode_M ?DB'" have 1: "M = ?M'" by (metis assms OAEPdefs(2) OAEP_EncryptDecrypt_IntValsMatch) have 2: "PKCS1_RSAES_OAEP_Decrypt C = ?M'" by (meson OAEPdefs(4)) show ?thesis using 1 2 by presburger qed text \In the following lemma, we put together all the facts that we know about the intermediate values when decrypting ciphertext C with RSA-OAEP.\ lemma OAEP_Decrypt_IntVals: assumes "PKCS1_RSAES_OAEP_Decrypt_validInput C lHash" "c = PKCS1_OS2IP C" "m = RSADECRYPT c" "EM = PKCS1_I2OSP m k" "Y = PKCS1_OAEP_decode_Y EM" "maskedSeed = PKCS1_OAEP_decode_maskedSeed EM" "maskedDB = PKCS1_OAEP_decode_maskedDB EM" "seedMask = PKCS1_OAEP_decode_seedMask maskedDB" "seed = PKCS1_OAEP_decode_seed maskedSeed seedMask" "dbMask = PKCS1_OAEP_decode_dbMask seed" "DB = PKCS1_OAEP_decode_DB maskedDB dbMask" "M = PKCS1_OAEP_decode_M DB" "lHash' = PKCS1_OAEP_decode_lHash DB" "PSlen = PKCS1_OAEP_decode_PSlen DB" "PS = replicate PSlen 0" shows "m < n \ Y = 0 \ lHash = lHash' \ octets_valid EM \ length EM = k \ octets_valid maskedSeed \ length maskedSeed = hLen \ octets_valid maskedDB \ length maskedDB = k - 1 - hLen \ octets_valid seedMask \ length seedMask = hLen \ octets_valid seed \ length seed = hLen \ octets_valid dbMask \ length dbMask = k - hLen - 1 \ octets_valid DB \ length DB = k - hLen - 1 \ octets_valid lHash' \ length lHash' = hLen \ octets_valid M \ length M = k - 2*hLen - PSlen - 2 \ M = PKCS1_RSAES_OAEP_Decrypt C \ PSlen \ k - 2*hLen - 2 \ DB = lHash @ PS @ [1] @ M \ EM = [0] @ maskedSeed @ maskedDB" proof - have l1: "2*hLen + 2 \ k" using assms(1,2,3) OAEPdefs(3) by meson have l2: "hLen + 1 \ k" using l1 by linarith have l3: "hLen + 1 \ k - hLen - 1" using l1 by auto have m: "m < n" using assms(1,2,3) RSA_bnd(2) by blast have EM1: "octets_valid EM" using assms(4) OAEP_decode_EM_octetsValid by fast have EM2: "length EM = k" using assms(4) m OAEP_decode_EM_len by auto have Y: "Y = 0" using OAEPdefs(3) assms(1,2,3,4,5) by meson have mS1: "octets_valid maskedSeed" using assms(6) EM1 OAEP_decode_maskedSeed_octetsValid by force have mS2: "length maskedSeed = hLen" using assms(6) l2 EM2 OAEP_decode_maskedSeed_len by presburger have mDB1: "octets_valid maskedDB" using assms(7) EM1 OAEP_decode_maskedDB_octetsValid by blast have mDB2: "length maskedDB = k - 1 - hLen" using assms(7) EM2 OAEP_decode_maskedDB_len by presburger have sM1: "octets_valid seedMask" using assms(8) mgf_valid OAEP_decode_seedMask_octetsValid by blast have sM2: "length seedMask = hLen" using assms(8) mgf_len OAEP_decode_seedMask_len by blast have s1: "octets_valid seed" using assms(9) mS1 sM1 OAEP_decode_seed_octetsValid by blast have s2: "length seed = hLen" using assms(9) mS2 sM2 OAEP_decode_seed_len by blast have dbM1: "octets_valid dbMask" using assms(10) OAEP_decode_dbMask_octetsValid by blast have dbM2: "length dbMask = k - hLen - 1" using assms(10) OAEP_decode_dbMask_len by blast have DB1: "octets_valid DB" using assms(11) mDB1 dbM1 OAEP_decode_DB_octetsValid by blast have DB2: "length DB = k - hLen - 1" using assms(11) mDB2 dbM2 OAEP_decode_DB_len by force have lH1: "octets_valid lHash'" using assms(13) DB1 OAEP_decode_lHash_octetsValid by fast have lH2: "length lHash' = hLen" using assms(13) l3 DB2 OAEP_decode_lHash_len by auto have lH3: "lHash = lHash'" by (metis OAEPdefs(3) assms(1,2,3,4,6,7,8,9,10,11,13)) have M1: "octets_valid M" using assms(12) DB1 OAEP_decode_M_octetsValid by fast have M2: "length M = k - 2*hLen - PSlen - 2" by (metis assms(12,14) DB2 OAEP_decode_M_len Suc_eq_plus1 add.left_neutral add_2_eq_Suc' diff_diff_left mult.commute mult.right_neutral mult_Suc_right) have M3: "M = PKCS1_RSAES_OAEP_Decrypt C" by (metis OAEPdefs(4) assms(2,3,4,6,7,8,9,10,11,12) ) have PSl1: "hLen + PSlen + 1 \ length DB" using assms OAEP_decode_validPS_DBlen OAEPdefs(3) by metis have PSl: "PSlen \ k - 2*hLen - 2" using PSl1 DB2 by linarith have DB3: "DB = lHash @ PS @ [1] @ M" using OAEP_decode_inputValid_DB assms by meson have EM3: "EM = [0] @ maskedSeed @ maskedDB" proof - have 1: "hd EM = 0" using Y PKCS1_OAEP_decode_Y_def assms(5) by presburger have 2: "take hLen (drop 1 EM) = maskedSeed" using PKCS1_OAEP_decode_maskedSeed_def assms(6) by presburger have 3: "drop (hLen + 1) EM = maskedDB" by (simp add: PKCS1_OAEP_decode_maskedDB_def assms(7)) show ?thesis by (metis 1 2 3 EM2 One_nat_def add_leD2 append_Cons append_Nil append_take_drop_id drop0 drop_Suc drop_drop hd_Cons_tl l2 list.size(3) not_one_le_zero) qed show ?thesis using m EM1 EM2 Y mS1 mS2 mDB1 mDB2 sM1 sM2 s1 s2 dbM1 dbM2 DB1 DB2 lH1 lH2 lH3 M1 M2 M3 PSl DB3 EM3 by blast qed text\If you start with a valid ciphertext, when you decode and decrypt it with RSA-OAEP, the resulting plaintexttext is a valid input for the RSA-OAEP encryption scheme.\ lemma OAEP_Decrypt_lengthValid [OAEPthms]: assumes "M = PKCS1_RSAES_OAEP_Decrypt C" "PKCS1_RSAES_OAEP_Decrypt_validInput C lHash" "mLen = length M" shows "PKCS1_RSAES_OAEP_Encrypt_lengthValid mLen" proof - let ?c = "PKCS1_OS2IP C" let ?m = "RSADECRYPT ?c" let ?EM = "PKCS1_I2OSP ?m k" let ?Y = "PKCS1_OAEP_decode_Y ?EM" let ?maskedSeed = "PKCS1_OAEP_decode_maskedSeed ?EM" let ?maskedDB = "PKCS1_OAEP_decode_maskedDB ?EM" let ?seedMask = "PKCS1_OAEP_decode_seedMask ?maskedDB" let ?seed = "PKCS1_OAEP_decode_seed ?maskedSeed ?seedMask" let ?dbMask = "PKCS1_OAEP_decode_dbMask ?seed" let ?DB = "PKCS1_OAEP_decode_DB ?maskedDB ?dbMask" let ?PSlen = "PKCS1_OAEP_decode_PSlen ?DB" let ?M = "PKCS1_OAEP_decode_M ?DB" have 1: "M = ?M" by (metis (no_types) OAEPdefs(4) assms(1)) have 2: "mLen = length ?M" using 1 assms(3) by fast have 3: "mLen = k - 2*hLen - ?PSlen - 2" using RSA_npos 2 assms(2,3) OAEP_Decrypt_IntVals by blast have 4: "2*hLen + 2 \ k" using assms(2) OAEPdefs(3) by meson show ?thesis using 3 4 OAEPdefs(1) by fastforce qed text \Assuming you have valid inputs, this lemma shows that the intermediate values computed while decrypting a ciphertext with RSA-OAEP match the corresponding intermediate values computed while encrypting the resulting plaintext.\ lemma OAEP_DecryptEncrypt_IntValsMatch: assumes "PKCS1_RSAES_OAEP_Decrypt_validInput C lHash" "octets_valid C" "c' = PKCS1_OS2IP C" "m' = RSADECRYPT c'" "EM' = PKCS1_I2OSP m' k" "Y' = PKCS1_OAEP_decode_Y EM'" "maskedSeed' = PKCS1_OAEP_decode_maskedSeed EM'" "maskedDB' = PKCS1_OAEP_decode_maskedDB EM'" "seedMask' = PKCS1_OAEP_decode_seedMask maskedDB'" "seed' = PKCS1_OAEP_decode_seed maskedSeed' seedMask'" "dbMask' = PKCS1_OAEP_decode_dbMask seed'" "DB' = PKCS1_OAEP_decode_DB maskedDB' dbMask'" "M' = PKCS1_OAEP_decode_M DB'" "lHash' = PKCS1_OAEP_decode_lHash DB'" "PSlen' = PKCS1_OAEP_decode_PSlen DB'" "mLen = length M'" "PS = PKCS1_OAEP_PS mLen" "DB = PKCS1_OAEP_DB lHash PS M'" "dbMask = PKCS1_OAEP_dbMask seed'" "maskedDB = PKCS1_OAEP_maskedDB DB dbMask" "seedMask = PKCS1_OAEP_seedMask maskedDB" "maskedSeed = PKCS1_OAEP_maskedSeed seed' seedMask" "EM = PKCS1_OAEP_EM maskedSeed maskedDB" "m = PKCS1_OS2IP EM" "c = RSAENCRYPT m" "C' = PKCS1_I2OSP c k" "Hash L = lHash" shows "DB = DB' \ dbMask = dbMask' \ maskedDB = maskedDB' \ seedMask = seedMask' \ Y' = 0 \ maskedSeed = maskedSeed' \ EM = EM' \ m = m' \ c = c' \ C = C' \ lHash' = lHash \ C' = PKCS1_RSAES_OAEP_Encrypt M' L seed' \ M' = PKCS1_RSAES_OAEP_Decrypt C \ (length PS = PSlen')" proof - have l1: "2*hLen + 2 \ k" by (meson OAEPdefs(3) assms(1)) have l2: "0 \ k - 2*hLen - 2" using l1 by blast let ?x = "k - 2*hLen - 2" have lH: "lHash' = lHash" by (metis assms(1,3,4,5,7,8,9,10,11,12,14) OAEPdefs(3)) have lH1: "hLen = length lHash'" using assms(27) hash_len lH by fastforce have Y: "Y' = 0" by (metis assms(1,3,4,5,6) OAEPdefs(3)) have M: "M' = PKCS1_RSAES_OAEP_Decrypt C" by (metis OAEPdefs(4) assms(3,4,5,7,8,9,10,11,12,13)) have M1: "mLen = k - 2*hLen - PSlen' - 2" using assms OAEP_Decrypt_IntVals by meson have M2: "mLen = ?x - PSlen'" using M1 by auto have M3: "mLen \ ?x" using M2 by linarith have PS1: "PSlen' \ k - 2*hLen - 2" using OAEP_Decrypt_IntVals assms(1,3,4,5,6,7,8,9,10,11,12,15) by blast have PS: "length PS = PSlen'" using M2 OAEP_PS_len PS1 assms(17) by auto have DB1: "DB = lHash @ PS @ [1] @ M'" using PKCS1_OAEP_DB_def assms(18) by presburger have DB2: "DB' = lHash' @ PS @ [1] @ M'" using assms(1,3,4,5,7,8,9,10,11,12,13,15,17) PS lH OAEP_decode_inputValid_DB PKCS1_OAEP_PS_def by auto have DB: "DB = DB'" using DB1 DB2 lH by fast have dbM: "dbMask = dbMask'" using DB PKCS1_OAEP_dbMask_def PKCS1_OAEP_decode_dbMask_def assms(11,19) by presburger have mDB: "maskedDB = maskedDB'" using DB OAEP.OAEP_Decrypt_IntVals OAEP.OAEP_dbMask_len OAEP_axioms PKCS1_OAEP_decode_DB_def PKCS1_OAEP_maskedDB_def assms(1,3,4,5,8,12,19,20) dbM words_xor_inv2 by auto have sM: "seedMask = seedMask'" by (simp add: assms(9,21) mDB PKCS1_OAEP_decode_seedMask_def PKCS1_OAEP_seedMask_def) have mS: "maskedSeed = maskedSeed'" by(metis sM assms(1,3,4,5,7,10,21,22) OAEP_Decrypt_IntVals OAEP_seedMask_len PKCS1_OAEP_decode_seed_def PKCS1_OAEP_maskedSeed_def words_xor_inv xor_words_symm) have EM: "EM = EM'" by (metis mDB mS assms(1,3,4,5,7,8,23) OAEP_Decrypt_IntVals PKCS1_OAEP_EM_def) have m: "m = m'" by (simp add: EM I2OSP_OS2IP assms(5,24)) have c: "c = c'" by (metis OAEPdefs(3) RSA_inv(2) assms(1,3,4,25) m) have C: "C = C'" by (metis OS2IP_I2OSP OAEPdefs(3) assms(1,2,3,26) c) have C2: "C' = PKCS1_RSAES_OAEP_Encrypt M' L seed'" by (metis OAEPdefs(2) assms(16,17,18,19,20,21,22,23,24,25,26,27)) show ?thesis using lH Y M PS DB dbM mDB sM mS EM m c C C2 by argo qed text\Start with a valid ciphertext. Remember that the ciphertext is only valid if the decoded lHash matches that hash of the known label L. Decrypt the ciphertext with RSA-OAEP. Remember the seed that you computed during that decryption. Then encrypt the plaintext that you got using that seed and label L. The result will be the original ciphertext.\ lemma OAEP_Decrypt_Encrypt [OAEPthms]: assumes "PKCS1_RSAES_OAEP_Decrypt_validInput C (Hash L)" "M = PKCS1_RSAES_OAEP_Decrypt C" "seed = PKCS1_RSAES_OAEP_Decrypt_seed C" "octets_valid C" shows "PKCS1_RSAES_OAEP_Encrypt M L seed = C" proof - let ?lHash = "Hash L" let ?mLen = "length M" let ?c' = "PKCS1_OS2IP C" let ?m' = "RSADECRYPT ?c'" let ?EM' = "PKCS1_I2OSP ?m' k" let ?Y' = "PKCS1_OAEP_decode_Y ?EM'" let ?maskedSeed' = "PKCS1_OAEP_decode_maskedSeed ?EM'" let ?maskedDB' = "PKCS1_OAEP_decode_maskedDB ?EM'" let ?seedMask' = "PKCS1_OAEP_decode_seedMask ?maskedDB'" let ?seed' = "PKCS1_OAEP_decode_seed ?maskedSeed' ?seedMask'" let ?dbMask' = "PKCS1_OAEP_decode_dbMask ?seed'" let ?DB' = "PKCS1_OAEP_decode_DB ?maskedDB' ?dbMask'" let ?M' = "PKCS1_OAEP_decode_M ?DB'" let ?lHash' = "PKCS1_OAEP_decode_lHash ?DB'" let ?PSlen' = "PKCS1_OAEP_decode_PSlen ?DB'" let ?mLen' = "length ?M'" let ?PS = "PKCS1_OAEP_PS ?mLen" let ?DB = "PKCS1_OAEP_DB ?lHash ?PS M" let ?dbMask = "PKCS1_OAEP_dbMask ?seed'" let ?maskedDB = "PKCS1_OAEP_maskedDB ?DB ?dbMask" let ?seedMask = "PKCS1_OAEP_seedMask ?maskedDB" let ?maskedSeed = "PKCS1_OAEP_maskedSeed ?seed' ?seedMask" let ?EM = "PKCS1_OAEP_EM ?maskedSeed ?maskedDB" let ?m = "PKCS1_OS2IP ?EM" let ?c = "RSAENCRYPT ?m" let ?C = "PKCS1_I2OSP ?c k" have 1: "seed = ?seed'" by (metis OAEPdefs(5) assms(3)) have 2: "C = ?C \ ?M' = PKCS1_RSAES_OAEP_Decrypt C \ ?C = PKCS1_RSAES_OAEP_Encrypt ?M' L seed" using 1 assms OAEPdefs(4) OAEP_DecryptEncrypt_IntValsMatch[of C ?lHash ?c' ?m' ?EM' ?Y' ?maskedSeed' ?maskedDB' ?seedMask' ?seed' ?dbMask' ?DB' ?M' ?lHash' ?PSlen' ?mLen' ?PS ?DB ?dbMask ?maskedDB ?seedMask ?maskedSeed ?EM ?m ?c ?C L] by metis show ?thesis using 2 assms(2) by argo qed end (*end of OAEP locale*) subsubsection \OAEP Interpretations\ text \To see an example of interpretation the OAEP locale, we define a very bad mask generating function, MFG0, and a very bad hash function, Hash0. These should never ever be used in a real application. We also need to have a valid RSA key.\ definition MGF0 :: mgf_type where "MGF0 os n = replicate n 0" lemma MGF0_len: "\x y. length (MGF0 x y) = y" using MGF0_def by simp lemma MGF0_valid: "\x y. octets_valid (MGF0 x y)" by (simp add: MGF0_def words_valid_def) definition Hash0 :: hash_type where "Hash0 os = [0]" lemma Hash0_len: "\x. length (Hash0 x) = 1" by (metis (no_types) Hash0_def One_nat_def Suc_eq_plus1 list.size(3) list.size(4)) lemma Hash0_valid: "\x. octets_valid (Hash0 x)" by (simp add: Hash0_def words_valid_def) locale OAEP_Basic_Example = fixes n d p q e :: nat assumes RSAkey: "PKCS1_validRSAprivateKey n d p q e" begin interpretation OAEP MGF0 Hash0 1 "PKCS1_RSAEP n e" "PKCS1_RSADP n d" n proof - have 1: "\x y. octets_valid (MGF0 x y)" using MGF0_valid by fast have 2: "\x y. length (MGF0 x y) = y" using MGF0_len by meson have 3: "\x. octets_valid (Hash0 x)" using Hash0_valid by blast have 4: "\x. length (Hash0 x) = 1" using Hash0_len by blast have 5: "0 < n" using n_positive2 RSAkey ValidKeyDefs(2) by meson have 6: "\m. PKCS1_RSAEP n e m < n" using 5 PKCS1_RSAEP_messageValid_def encryptValidCiphertext by presburger have 7: "\c. PKCS1_RSADP n d c < n" using 5 PKCS1_RSAEP_messageValid_def encryptValidCiphertext by presburger have 8: "\mcx y. octets_valid (MGF0 x y)" using MGF0_valid by fast have 2: "\x y. length (MGF0 x y) = y" using MGF0_len by meson have 3: "\x. octets_valid (Hash0 x)" using Hash0_valid by blast have 4: "\x. length (Hash0 x) = 1" using Hash0_len by blast have 5: "0 < n" by (metis RSAkey ValidKeyDefs(3) n_positive2) have 6: "\m. PKCS1_RSAEP n e m < n" using 5 PKCS1_RSAEP_messageValid_def encryptValidCiphertext by presburger have 7: "\c. PKCS1_RSADP_CRT p q dP dQ qInv c < n" using PKCS1_RSAEP_messageValid_def RSAkey decryptCRTmessageValid2 by auto have 8: "\mcSection 7.2: RSAES-PKCS1-v1.5\ text \"RSAES-PKCS1-v1_5 combines the RSAEP and RSADP primitives (Sections 5.1.1 and 5.1.2) with the EME-PKCS1-v1_5 encoding method (step 1 in Section 7.2.1 and step 3 in Section 7.2.2). It is mathematically equivalent to the encryption scheme in PKCS #1 v1.5. RSAES-PKCS1-v1_5 can operate on messages of length up to k – 11 octets (k is the octet length of the RSA modulus), although care should be taken to avoid certain attacks on low-exponent RSA due to Coppersmith, Franklin, Patarin, and Reiter when long messages are encrypted (see the third bullet in the notes below and [10]; [14] contains an improved attack). As a general rule, the use of this scheme for encrypting an arbitrary message, as opposed to a randomly generated key, is not recommended. The padding string PS in step 2 in Section 7.2.1 is at least eight octets long, which is a security condition for public-key operations that makes it difficult for an attacker to recover data by trying all possible encryption blocks." Here the padding string PS is randomly generated and assumed to be all non-zero. Just as with the seed in OAEP above, we include PS here with the list of inputs of this encryption scheme.\ locale RSAES_PKCS1_v1_5 = fixes RSAENCRYPT RSADECRYPT :: "nat \ nat" and n :: nat assumes RSA_mod: "0 < n" "11 < octet_length n" and RSA_bnd: "\m. (RSAENCRYPT m) < n" "\c. (RSADECRYPT c) < n" and RSA_inv: "\m. (m < n) \ RSADECRYPT (RSAENCRYPT m) = m" "\c. (c < n) \ RSAENCRYPT (RSADECRYPT c) = c" begin named_theorems v1_5defs named_theorems v1_5thms definition k where "k = octet_length n" lemma k_bnd: "11 < k" using k_def RSA_mod(2) by fastforce subsubsection \Section 7.2.1: Encryption Operation\ definition RSAES_PKCS1_v1_5_Encrypt_EM :: "octets \ octets \ octets" where [v1_5defs]: "RSAES_PKCS1_v1_5_Encrypt_EM PS M = [0,2] @ PS @ [0] @ M" definition RSAES_PKCS1_v1_5_Encrypt_inputValid :: "octets \ octets \ bool" where [v1_5defs]: "RSAES_PKCS1_v1_5_Encrypt_inputValid M PS = (let mLen = length M; psLen = length PS in (mLen + 11 \ k) \ (psLen + mLen + 3 = k) \ (8 \ psLen) \ (\p \ set PS. 0 < p) \ (octets_valid M) \ (octets_valid PS) )" definition RSAES_PKCS1_v1_5_Encrypt :: "octets \ octets \ octets" where [v1_5defs]: "RSAES_PKCS1_v1_5_Encrypt M PS = (let EM = RSAES_PKCS1_v1_5_Encrypt_EM PS M; m = PKCS1_OS2IP EM; c = RSAENCRYPT m in PKCS1_I2OSP c k )" lemma RSAES_v1_5_Encrypt_EM_valid: assumes "octets_valid M" "octets_valid PS" "EM = RSAES_PKCS1_v1_5_Encrypt_EM PS M" shows "octets_valid EM" using assms v1_5defs(1) words_valid_def by auto lemma RSAES_v1_5_Encrypt_EM_valid2: assumes "RSAES_PKCS1_v1_5_Encrypt_inputValid M PS" "EM = RSAES_PKCS1_v1_5_Encrypt_EM PS M" shows "octets_valid EM" using assms v1_5defs(2) RSAES_v1_5_Encrypt_EM_valid by meson lemma RSAES_v1_5_Encrypt_EM_len: "length (RSAES_PKCS1_v1_5_Encrypt_EM PS M) = (length PS) + (length M) + 3" by (metis One_nat_def v1_5defs(1) Suc_eq_plus1 add.assoc length_append list.size(3,4) numeral_3_eq_3 plus_1_eq_Suc) lemma RSAES_v1_5_Encrypt_EM_stripZero: assumes "EM = RSAES_PKCS1_v1_5_Encrypt_EM PS M" shows "words_strip_zero EM = [2] @ PS @ [0] @ M" by (simp add: v1_5defs(1) assms) lemma RSAES_v1_5_Encrypt_EM_stripLen: assumes "EM = RSAES_PKCS1_v1_5_Encrypt_EM PS M" shows "length (words_strip_zero EM) = (length PS) + (length M) + 2" by (metis assms RSAES_v1_5_Encrypt_EM_stripZero One_nat_def Suc_eq_plus1 add_Suc_right length_append list.size(3,4) numerals(2) plus_1_eq_Suc) lemma RSAES_v1_5_Encrypt_m_bnd: assumes "RSAES_PKCS1_v1_5_Encrypt_inputValid M PS" "EM = RSAES_PKCS1_v1_5_Encrypt_EM PS M" "m = PKCS1_OS2IP EM" shows "m < n" proof - let ?mLen = "length M" let ?psLen = "length PS" have 1: "(?psLen + ?mLen + 3 = k)" using assms(1) v1_5defs(2) by meson have 2: "length (words_strip_zero EM) = k - 1" using 1 RSAES_v1_5_Encrypt_EM_stripLen assms(2) by auto have 3: "octets_valid EM" using assms(1,2) RSAES_v1_5_Encrypt_EM_valid2 by simp show ?thesis by (metis 2 3 assms(3) k_def words_valid_def PKCS1_OS2IP_def RSA_mod(1) less_le_trans word_length_eq3 words_strip0_to_nat_len_bnd Twoto8 zero_less_numeral) qed lemma RSAES_v1_5_Encrypt_len: assumes "C = RSAES_PKCS1_v1_5_Encrypt M PS" shows "length C = k" by (metis I2OSP_valid_len PKCS1_RSAEP_def v1_5defs(3) RSA_bnd(1) assms encryptValidI2OSP k_def less_nat_zero_code mod_less nat_neq_iff power_one_right) subsubsection \Section 7.2.2: Decryption Operation\ definition RSAES_PKCS1_v1_5_Decode_M :: "octets \ octets" where [v1_5defs]: "RSAES_PKCS1_v1_5_Decode_M EM = drop 1 (words_strip_nonzero (drop 2 EM))" definition RSAES_PKCS1_v1_5_Decode_PS :: "octets \ octets" where [v1_5defs]: "RSAES_PKCS1_v1_5_Decode_PS EM = ( let emLen = length EM; M = RSAES_PKCS1_v1_5_Decode_M EM; mLen = length M; psLen = emLen - mLen - 3 in take psLen (drop 2 EM) )" definition RSAES_PKCS1_v1_5_Decode_validEM :: "octets \ bool" where [v1_5defs]: "RSAES_PKCS1_v1_5_Decode_validEM EM = ( let emLen = length EM; ZeroTwo = take 2 EM; Zero_M = words_strip_nonzero (drop 2 EM); M = RSAES_PKCS1_v1_5_Decode_M EM; mLen = length M; psLen = emLen - mLen - 3 in (ZeroTwo = [0,2]) \ (Zero_M \ []) \ (8 \ psLen) )" definition RSAES_PKCS1_v1_5_Decrypt_inputValid :: "octets \ bool" where [v1_5defs]: "RSAES_PKCS1_v1_5_Decrypt_inputValid C = ( let c = PKCS1_OS2IP C; m = RSADECRYPT c; EM = PKCS1_I2OSP m k in (length C = k) \ (octets_valid C) \ (c < n) \ (RSAES_PKCS1_v1_5_Decode_validEM EM) )" definition RSAES_PKCS1_v1_5_Decrypt :: "octets \ octets" where [v1_5defs]: "RSAES_PKCS1_v1_5_Decrypt C = ( let c = PKCS1_OS2IP C; m = RSADECRYPT c; EM = PKCS1_I2OSP m k in RSAES_PKCS1_v1_5_Decode_M EM )" text \It is convenient to have the following function to recover the padding string from the ciphertext. Or just the enocoded message.\ definition RSAES_PKCS1_v1_5_Decrypt_PS :: "octets \ octets" where "RSAES_PKCS1_v1_5_Decrypt_PS C = ( let c = PKCS1_OS2IP C; m = RSADECRYPT c; EM = PKCS1_I2OSP m k in RSAES_PKCS1_v1_5_Decode_PS EM )" definition RSAES_PKCS1_v1_5_Decrypt_EM :: "octets \ octets" where "RSAES_PKCS1_v1_5_Decrypt_EM C = ( let c = PKCS1_OS2IP C; m = RSADECRYPT c in PKCS1_I2OSP m k )" lemma RSAES_v1_5_Decode_EM: assumes "M = RSAES_PKCS1_v1_5_Decode_M EM" "PS = RSAES_PKCS1_v1_5_Decode_PS EM" "RSAES_PKCS1_v1_5_Decode_validEM EM" shows "EM = [0,2] @ PS @ [0] @ M" proof - let ?emLen = "length EM" let ?ZeroTwo = "take 2 EM" let ?Zero_M = "words_strip_nonzero (drop 2 EM)" let ?mLen = "length M" let ?psLen = "?emLen - ?mLen - 3" have 1: "?ZeroTwo = [0,2]" using assms(3) v1_5defs(6) by algebra have 2: "M = drop 1 ?Zero_M" using assms(1) v1_5defs(4) by blast have 3: "?Zero_M \ []" using assms(3) v1_5defs(6) by algebra have 4: "take 1 ?Zero_M = [0]" by (metis 3 strip_non0_head Cons_nth_drop_Suc cancel_comm_monoid_add_class.diff_cancel drop0 take0 length_greater_0_conv numerals(1) take_Cons_numeral) have 5: "PS = take ?psLen (drop 2 EM)" using assms(1,2) v1_5defs(5) by presburger have 6: "8 \ ?psLen" using assms(1,3) v1_5defs(6) by algebra show ?thesis by (metis (mono_tags, lifting) 1 2 4 5 One_nat_def Suc_eq_plus1 append_take_drop_id diff_diff_add length_Cons length_append length_drop list.size(3) numeral_2_eq_2 add.commute numeral_3_eq_3 strip_non0_drop) qed lemma RSAES_v1_5_Decode_EM_len: assumes "m = RSADECRYPT c" "EM = PKCS1_I2OSP m k" shows "length EM = k" by (metis I2OSP_valid_len PKCS1_I2OSP_inputValid_def RSA_bnd(2) assms k_def less_trans word_length_eq2 Twoto8 zero_less_numeral) lemma RSAES_v1_5_Decode_EM_valid: assumes "EM = PKCS1_I2OSP m k" shows "octets_valid EM" by (simp add: I2OSP_octets_valid assms) lemma RSAES_v1_5_Decode_M_valid: assumes "octets_valid EM" "M = RSAES_PKCS1_v1_5_Decode_M EM" shows "octets_valid M" by (metis v1_5defs(4) assms in_set_dropD words_valid_def strip_non0_drop) lemma RSAES_v1_5_Decode_PS_valid: assumes "octets_valid EM" "PS = RSAES_PKCS1_v1_5_Decode_PS EM" shows "octets_valid PS" by (metis v1_5defs(5) assms in_set_dropD in_set_takeD words_valid_def) lemma RSAES_v1_5_Decode_PS_nonzero: assumes "PS = RSAES_PKCS1_v1_5_Decode_PS EM" "RSAES_PKCS1_v1_5_Decode_validEM EM" shows "\p \ set PS. 0 < p" proof - let ?emLen = "length EM" let ?M = "drop 1 (words_strip_nonzero (drop 2 EM))" let ?mLen = "length ?M" let ?psLen = "?emLen - ?mLen - 3" have 1: "PS = take ?psLen (drop 2 EM)" using assms(1) v1_5defs(4,5) by presburger let ?X = "drop 2 EM" let ?Y = "words_strip_nonzero ?X" have 2: "length ?X = ?emLen - 2" using length_drop by blast have 3: "length ?Y = ?mLen + 1" by (metis One_nat_def v1_5defs(6) Suc_eq_plus1 Suc_pred assms(2) length_drop length_greater_0_conv) have 4: "\i < ((length ?X) - (length ?Y)). 0 < ?X ! i" using strip_non0_leading_pos by blast have 5: "(length ?X) - (length ?Y) = ?psLen" using 2 3 by fastforce have 6: "PS = take ?psLen ?X" using 1 by fast have 7: "\i < ?psLen. 0 < ?X ! i" using 4 5 by presburger have 8: "\i < ?psLen. ?X ! i = PS ! i" by (metis 6 nth_take) have 9: "\i < ?psLen. 0 < PS ! i" using 7 8 by algebra have 10: "length PS = ?psLen" by (metis 5 6 add_diff_cancel_right' append_take_drop_id length_append strip_non0_drop) show ?thesis by (metis 9 10 in_set_conv_nth) qed subsubsection \RSAES_PKCS1_v1_5 Lemmas\ lemma RSAES_PKCS1_v1_5_EncryptDecrypt_IntValsMatch [v1_5thms]: assumes "RSAES_PKCS1_v1_5_Encrypt_inputValid M PS" "EM = RSAES_PKCS1_v1_5_Encrypt_EM PS M" "m = PKCS1_OS2IP EM" "c = RSAENCRYPT m" "C = PKCS1_I2OSP c k" "c' = PKCS1_OS2IP C" "m' = RSADECRYPT c'" "EM' = PKCS1_I2OSP m' k" "M' = RSAES_PKCS1_v1_5_Decode_M EM'" shows "M = M' \ EM = EM' \ m = m' \ c = c' \ EM = [0,2] @ PS @ [0] @ M \ C = RSAES_PKCS1_v1_5_Encrypt M PS \ M' = RSAES_PKCS1_v1_5_Decrypt C \ M = RSAES_PKCS1_v1_5_Decode_M EM" proof - have c: "c = c'" by (simp add: I2OSP_OS2IP assms(5,6)) have m: "m = m'" using RSA_inv(1) assms(1,2,3,4,7) c RSAES_v1_5_Encrypt_m_bnd by force have EM1: "octets_valid EM" using assms(1,2) RSAES_v1_5_Encrypt_EM_valid2 by blast have EM2: "length EM = k" by (metis (no_types) v1_5defs(2) assms(1,2) RSAES_v1_5_Encrypt_EM_len) have EM: "EM = EM'" using m EM1 EM2 OS2IP_I2OSP assms(3,8) by fastforce have EM3: "EM = [0,2] @ PS @ [0] @ M" using assms(2) v1_5defs(1) by fast have M1: "drop 2 EM = PS @ [0] @ M" using EM3 by auto have M2: "\p \ set PS. 0 < p" using assms(1) v1_5defs(2) by meson have M3: "words_strip_nonzero (PS @ [0] @ M) = [0] @ M" using strip_non0_concat M2 by blast have M4: "M = RSAES_PKCS1_v1_5_Decode_M EM" using M1 M3 v1_5defs(4) by simp have M: "M = M'" using EM M1 M3 assms(9) v1_5defs(4) by simp have C1: "C = RSAES_PKCS1_v1_5_Encrypt M PS" using assms(2,3,4,5) v1_5defs(3) by presburger have C2: "M' = RSAES_PKCS1_v1_5_Decrypt C" using assms(6,7,8,9) v1_5defs(8) by presburger show ?thesis using c m EM EM3 M M4 C1 C2 by blast qed lemma RSAES_PKCS1_v1_5_Encrypt_validCipher [v1_5thms]: assumes "RSAES_PKCS1_v1_5_Encrypt_inputValid M PS" "C = RSAES_PKCS1_v1_5_Encrypt M PS" shows "RSAES_PKCS1_v1_5_Decrypt_inputValid C" proof - let ?EM = "RSAES_PKCS1_v1_5_Encrypt_EM PS M" let ?m = "PKCS1_OS2IP ?EM" let ?c = "RSAENCRYPT ?m" let ?C' = "PKCS1_I2OSP ?c k" let ?c' = "PKCS1_OS2IP ?C'" let ?m' = "RSADECRYPT ?c'" let ?EM' = "PKCS1_I2OSP ?m' k" let ?M' = "RSAES_PKCS1_v1_5_Decode_M ?EM'" have C1: "C = ?C'" using assms(2) v1_5defs(3) by presburger have C2: "length C = k" by (simp add: assms(2) RSAES_v1_5_Encrypt_len) have C3: "octets_valid C" by (simp add: C1 I2OSP_octets_valid) have c: "?c < n" using RSA_bnd(1) by blast have EM: "?EM = ?EM'" using v1_5thms(1) assms(1) by blast have EM1: "?EM = [0,2] @ PS @ [0] @ M" using v1_5defs(1) by blast let ?emLen = "length ?EM" let ?ZeroTwo = "take 2 ?EM" let ?Zero_M = "words_strip_nonzero (drop 2 ?EM)" let ?mLen = "length ?M'" let ?psLen = "?emLen - ?mLen - 3" have vEM1: "?ZeroTwo = [0,2]" using EM1 by simp have vEM2: "?Zero_M = [0] @ M" by (metis EM1 assms(1) One_nat_def v1_5defs(2) Suc_eq_plus1 append_eq_conv_conj list.size(3,4) numerals(2) strip_non0_concat) have vEM3: "?Zero_M \ []" using vEM2 by simp have vEM4: "8 \ ?psLen" by (metis assms(1) v1_5defs(2) RSAES_v1_5_Encrypt_EM_len v1_5thms(1) add_diff_cancel_right' diff_commute) show ?thesis using C1 C2 C3 c EM vEM1 vEM3 vEM4 v1_5defs(6,7) I2OSP_OS2IP by metis qed lemma RSAES_PKCS1_v1_5_Encrypt_Decrypt [v1_5thms]: assumes "RSAES_PKCS1_v1_5_Encrypt_inputValid M PS" "C = RSAES_PKCS1_v1_5_Encrypt M PS" shows "RSAES_PKCS1_v1_5_Decrypt C = M" by (metis (no_types) assms v1_5thms(1)) lemma RSAES_PKCS1_v1_5_DecryptEncrypt_IntValsMatch [v1_5thms]: assumes "RSAES_PKCS1_v1_5_Decrypt_inputValid C" "c' = PKCS1_OS2IP C" "m' = RSADECRYPT c'" "EM' = PKCS1_I2OSP m' k" "M' = RSAES_PKCS1_v1_5_Decode_M EM'" "PS' = RSAES_PKCS1_v1_5_Decode_PS EM'" "EM = RSAES_PKCS1_v1_5_Encrypt_EM PS' M'" "m = PKCS1_OS2IP EM" "c = RSAENCRYPT m" "C' = PKCS1_I2OSP c k" shows "C = C' \ EM = EM' \ m = m' \ c = c' \ EM = [0,2] @ PS' @ [0] @ M' \ C' = RSAES_PKCS1_v1_5_Encrypt M' PS' \ M' = RSAES_PKCS1_v1_5_Decrypt C" proof - have EM1: "RSAES_PKCS1_v1_5_Decode_validEM EM'" using assms(1,2,3,4) v1_5defs(7) by meson have EM2: "EM = [0,2] @ PS' @ [0] @ M'" using assms(7) v1_5defs(1) by fast have EM3: "EM = EM'" using EM1 EM2 RSAES_v1_5_Decode_EM assms(5,6) by presburger have m: "m = m'" using EM3 I2OSP_OS2IP assms(4,8) by presburger have c: "c = c'" by (metis m v1_5defs(7) RSA_inv(2) assms(1,2,3,9)) have C: "C = C'" by (metis (no_types) OS2IP_I2OSP v1_5defs(7) assms(1,2,10) c) have M1: "C' = RSAES_PKCS1_v1_5_Encrypt M' PS'" using assms(7,8,9,10) v1_5defs(3) by presburger have M2: "M' = RSAES_PKCS1_v1_5_Decrypt C" using assms(2,3,4,5) v1_5defs(8) by presburger show ?thesis using EM2 EM3 m c C M1 M2 by blast qed lemma RSAES_PKCS1_v1_5_Decrypt_validInput [v1_5thms]: assumes "M = RSAES_PKCS1_v1_5_Decrypt C" "RSAES_PKCS1_v1_5_Decrypt_inputValid C" "c = PKCS1_OS2IP C" "m = RSADECRYPT c" "EM = PKCS1_I2OSP m k" "PS = RSAES_PKCS1_v1_5_Decode_PS EM" shows "RSAES_PKCS1_v1_5_Encrypt_inputValid M PS" proof - have Mval: "octets_valid M" using I2OSP_octets_valid v1_5defs(8) assms(1) RSAES_v1_5_Decode_M_valid by presburger have PSval: "octets_valid PS" using I2OSP_octets_valid assms(5) assms(6) RSAES_v1_5_Decode_PS_valid by force have EMval: "RSAES_PKCS1_v1_5_Decode_validEM EM" using assms(2,3,4,5) v1_5defs(7) by meson have PSpos: "(\p \ set PS. 0 < p)" using RSAES_v1_5_Decode_PS_nonzero EMval assms(6) by blast let ?mLen = "length M" let ?psLen = "length PS" let ?emLen = "length EM" have l1: "?emLen = k" by (simp add: assms(4,5) RSAES_v1_5_Decode_EM_len) have l2: "?psLen = ?emLen - ?mLen - 3" by (metis assms v1_5thms(4) add_diff_cancel_right' diff_commute RSAES_v1_5_Encrypt_EM_len) have l3: "8 \ ?psLen" by (metis EMval v1_5defs(6,8) assms(1,3,4,5) l2) have l4: "?psLen + ?mLen + 3 = k" by (metis l1 v1_5thms(4) assms RSAES_v1_5_Encrypt_EM_len) have l5: "?mLen + 11 \ k" using l3 l4 by simp show ?thesis using Mval PSval PSpos l3 l4 l5 v1_5defs(2) by presburger qed lemma RSAES_PKCS1_v1_5_Decrypt_Encrypt [v1_5thms]: assumes "RSAES_PKCS1_v1_5_Decrypt_inputValid C" "RSAES_PKCS1_v1_5_Decrypt C = M" "c = PKCS1_OS2IP C" "m = RSADECRYPT c" "EM = PKCS1_I2OSP m k" "PS = RSAES_PKCS1_v1_5_Decode_PS EM" shows "C = RSAES_PKCS1_v1_5_Encrypt M PS" using v1_5thms(4) assms by force end (*v1_5 locale*) subsubsection \PSAES_PKCS1_v1_5 Interpretations\ text\We interpret the v1_5 locale for both the basic RSA decryption primitive and the decryption primitive that uses the Chinese Remainder Theorem.\ locale RSAES_PKCS1_v1_5_Basic = fixes n d p q e :: nat assumes RSAkey: "PKCS1_validRSAprivateKey n d p q e" "11 < octet_length n" begin interpretation RSAES_PKCS1_v1_5 "PKCS1_RSAEP n e" "PKCS1_RSADP n d" n proof - have 1: "0 < n" using n_positive2 RSAkey(1) ValidKeyDefs(2) by meson have 2: "11 < octet_length n" using RSAkey(2) by blast have 3: "\m. PKCS1_RSAEP n e m < n" using 1 PKCS1_RSAEP_messageValid_def encryptValidCiphertext by presburger have 4: "\c. PKCS1_RSADP n d c < n" using 1 PKCS1_RSAEP_messageValid_def encryptValidCiphertext by presburger have 5: "\mcm. PKCS1_RSAEP n e m < n" using 1 PKCS1_RSAEP_messageValid_def encryptValidCiphertext by presburger have 4: "\c. PKCS1_RSADP_CRT p q dP dQ qInv c < n" using PKCS1_RSAEP_messageValid_def RSAkey decryptCRTmessageValid2 by auto have 5: "\mcSection 9: Encoding Methods for Signatures\ text \Definitions in Section 8 depend on definitions in Section 9. So we include Section 9 first. "Encoding methods consist of operations that map between octet string messages and octet string encoded messages, which are converted to and from integer message representatives in the schemes. The integer message representatives are processed via the primitives. The encoding methods thus provide the connection between the schemes, which process messages, and the primitives." \ subsection \Section 9.1: EMSA-PSS\ locale EMSA_PSS = fixes MGF :: mgf_type and Hash :: hash_type and hLen :: nat assumes mgf_valid: "\x y. octets_valid (MGF x y)" and mgf_len: "\x y. length (MGF x y) = y" and hash_valid: "\x. octets_valid (Hash x)" and hash_len: "\x. length (Hash x) = hLen" begin subsubsection \Section 9.1.1: Encoding Operation\ definition PKCS1_EMSA_PSS_Encode_inputValid :: "nat \ nat \ bool" where "PKCS1_EMSA_PSS_Encode_inputValid sLen emLen = (hLen + sLen + 2 \ emLen)" definition PKCS1_EMSA_PSS_Encode_M' : "PKCS1_EMSA_PSS_Encode_M' mHash salt = (replicate 8 0) @ mHash @ salt" definition PKCS1_EMSA_PSS_Encode_H : "PKCS1_EMSA_PSS_Encode_H M' = Hash M'" definition PKCS1_EMSA_PSS_Encode_PS : "PKCS1_EMSA_PSS_Encode_PS emLen sLen = replicate (emLen - sLen - hLen - 2) 0" definition PKCS1_EMSA_PSS_Encode_DB : "PKCS1_EMSA_PSS_Encode_DB PS salt = PS @ [1] @ salt" definition PKCS1_EMSA_PSS_Encode_dbMask : "PKCS1_EMSA_PSS_Encode_dbMask H emLen = MGF H (emLen - hLen - 1)" definition PKCS1_EMSA_PSS_Encode_maskedDB : "PKCS1_EMSA_PSS_Encode_maskedDB DB dbMask = xor_octets DB dbMask" definition PKCS1_EMSA_PSS_Encode_EM :: "octets \ octets \ octets" where "PKCS1_EMSA_PSS_Encode_EM maskedDB_sethi H = maskedDB_sethi @ H @ [0xbc]" definition PKCS1_EMSA_PSS_Encode :: "octets \ octets \ nat \ octets" where "PKCS1_EMSA_PSS_Encode M salt emBits = ( let emLen = numBits_to_numOctets emBits; mHash = Hash M; sLen = length salt; M' = PKCS1_EMSA_PSS_Encode_M' mHash salt; H = PKCS1_EMSA_PSS_Encode_H M'; PS = PKCS1_EMSA_PSS_Encode_PS emLen sLen; DB = PKCS1_EMSA_PSS_Encode_DB PS salt; dbMask = PKCS1_EMSA_PSS_Encode_dbMask H emLen; maskedDB = PKCS1_EMSA_PSS_Encode_maskedDB DB dbMask; maskedDB' = SetLeftmostBits emLen emBits maskedDB in PKCS1_EMSA_PSS_Encode_EM maskedDB' H )" subsubsection \Section 9.1.2: Verification Operation\ definition PKCS1_EMSA_PSS_Verify_maskedDB : "PKCS1_EMSA_PSS_Verify_maskedDB emLen EM = take (emLen - hLen - 1) EM" definition PKCS1_EMSA_PSS_Verify_H : "PKCS1_EMSA_PSS_Verify_H emLen EM = take hLen (drop (emLen - hLen - 1) EM)" definition PKCS1_EMSA_PSS_Verify_dbMask : "PKCS1_EMSA_PSS_Verify_dbMask H emLen = MGF H (emLen - hLen - 1)" definition PKCS1_EMSA_PSS_Verify_DB : "PKCS1_EMSA_PSS_Verify_DB maskedDB dbMask = xor_octets maskedDB dbMask" definition PKCS1_EMSA_PSS_Verify_PS_0x01 : "PKCS1_EMSA_PSS_Verify_PS_0x01 emLen sLen DB_sethi = PKCS1_OS2IP (take (emLen - hLen - sLen -1) DB_sethi)" definition PKCS1_EMSA_PSS_Verify_salt : "PKCS1_EMSA_PSS_Verify_salt emLen sLen DB_sethi = drop (emLen - hLen - sLen - 1) DB_sethi" definition PKCS1_EMSA_PSS_Verify_M' : "PKCS1_EMSA_PSS_Verify_M' mHash salt = (replicate 8 0) @ mHash @ salt" definition PKCS1_EMSA_PSS_Verify_H' : "PKCS1_EMSA_PSS_Verify_H' M' = Hash M'" definition PKCS1_EMSA_PSS_Verify :: "nat \ octets \ octets \ nat \ bool" where "PKCS1_EMSA_PSS_Verify sLen M EM emBits = ( let emLen = numBits_to_numOctets emBits; mHash = Hash M; maskedDB = PKCS1_EMSA_PSS_Verify_maskedDB emLen EM; H = PKCS1_EMSA_PSS_Verify_H emLen EM; dbMask = PKCS1_EMSA_PSS_Verify_dbMask H emLen; DB = PKCS1_EMSA_PSS_Verify_DB maskedDB dbMask; DB' = SetLeftmostBits emLen emBits DB; PS_0x01 = PKCS1_EMSA_PSS_Verify_PS_0x01 emLen sLen DB'; salt = PKCS1_EMSA_PSS_Verify_salt emLen sLen DB'; M' = PKCS1_EMSA_PSS_Encode_M' mHash salt; H' = PKCS1_EMSA_PSS_Encode_H M' in (length EM = emLen) \ (hLen + sLen + 2 \ emLen) \ (last EM = 0xbc) \ PS_0x01 = 1 \ (H = H') \ (TestLeftmostBits emLen emBits maskedDB) )" subsubsection \EMSA-PSS Lemmas\ lemma EMSA_PSS_Encode_M'_valid: assumes "octets_valid mHash" "octets_valid salt" "M' = PKCS1_EMSA_PSS_Encode_M' mHash salt" shows "octets_valid M'" by (metis assms PKCS1_EMSA_PSS_Encode_M' words_valid_concat words_valid_zeros) lemma EMSA_PSS_Encode_M'_len: assumes "length mHash = hLen" "length salt = sLen" "M' = PKCS1_EMSA_PSS_Encode_M' mHash salt" shows "length M' = 8 + hLen + sLen" by (simp add: assms PKCS1_EMSA_PSS_Encode_M') lemma EMSA_PSS_Encode_H_valid: assumes "H = PKCS1_EMSA_PSS_Encode_H M'" shows "octets_valid H" using assms PKCS1_EMSA_PSS_Encode_H hash_valid by presburger lemma EMSA_PSS_Encode_H_len: assumes "H = PKCS1_EMSA_PSS_Encode_H M'" shows "length H = hLen" using assms PKCS1_EMSA_PSS_Encode_H hash_len by presburger lemma EMSA_PSS_Encode_PS_valid: assumes "PS = PKCS1_EMSA_PSS_Encode_PS emLen sLen" shows "octets_valid PS" using assms words_valid_zeros PKCS1_EMSA_PSS_Encode_PS by metis lemma EMSA_PSS_Encode_PS_len: assumes "PS = PKCS1_EMSA_PSS_Encode_PS emLen sLen" shows "length PS = emLen - sLen - hLen - 2" by (simp add: assms PKCS1_EMSA_PSS_Encode_PS) lemma EMSA_PSS_Encode_DB_valid: assumes "octets_valid PS" "octets_valid salt" "DB = PKCS1_EMSA_PSS_Encode_DB PS salt" shows "octets_valid DB" by (metis assms PKCS1_EMSA_PSS_Encode_DB words_valid_concat words_valid_cons words_valid_nil one_less_numeral_iff semiring_norm(76) Twoto8) lemma EMSA_PSS_Encode_DB_len: assumes "length PS = emLen - sLen - hLen - 2" "length salt = sLen" "DB = PKCS1_EMSA_PSS_Encode_DB PS salt" "PKCS1_EMSA_PSS_Encode_inputValid sLen emLen" shows "length DB = emLen - hLen - 1" proof - have 1: "hLen + sLen + 2 \ emLen" using assms(4) PKCS1_EMSA_PSS_Encode_inputValid_def by blast have 2: "length DB = (length PS) + 1 + (length salt)" by (simp add: assms(3) PKCS1_EMSA_PSS_Encode_DB) show ?thesis using 1 2 assms(1,2) by simp qed lemma EMSA_PSS_Encode_DB_hibits: assumes "emLen = numBits_to_numOctets emBits" "PS = PKCS1_EMSA_PSS_Encode_PS emLen sLen" "DB = PKCS1_EMSA_PSS_Encode_DB PS salt" shows "TestLeftmostBits emLen emBits DB" proof - have 1: "DB = PS @ [1] @ salt" using assms(3) PKCS1_EMSA_PSS_Encode_DB by blast let ?x = "8*emLen - emBits" have 2: "?x < 8 \ 0 \ ?x" using emLen_emBits assms(1) by blast let ?y = "8 - ?x" have 3: "1 \ ?y \ ?y \ 8" using 2 by force have 4: "2^1 \ (2::nat)^?y" using 3 one_le_numeral power_increasing by blast have 5: "(2::nat) \ 2^?y" using 4 by force have 6: "(1::nat) < 2^?y" using 5 by fastforce let ?h = "hd DB" show ?thesis proof (cases "PS = []") case True then have "?h = 1" using 1 by simp then show ?thesis using 6 TestLeftmostBits_def by algebra next case False then have "?h = 0" by (metis 1 assms(2) PKCS1_EMSA_PSS_Encode_PS EMSA_PSS_Encode_PS_len hd_append hd_replicate length_0_conv) then show ?thesis using 6 TestLeftmostBits_def by presburger qed qed lemma EMSA_PSS_Encode_dbMask_valid: assumes "dbMask = PKCS1_EMSA_PSS_Encode_dbMask H emLen" shows "octets_valid dbMask" using assms PKCS1_EMSA_PSS_Encode_dbMask mgf_valid by presburger lemma EMSA_PSS_Encode_dbMask_len: assumes "dbMask = PKCS1_EMSA_PSS_Encode_dbMask H emLen" shows "length dbMask = emLen - hLen - 1" using assms PKCS1_EMSA_PSS_Encode_dbMask mgf_len by presburger lemma EMSA_PSS_Encode_maskedDB_valid: assumes "maskedDB = PKCS1_EMSA_PSS_Encode_maskedDB DB dbMask" "octets_valid DB" "octets_valid dbMask" shows "octets_valid maskedDB" by (simp add: PKCS1_EMSA_PSS_Encode_maskedDB assms xor_valid_words) lemma EMSA_PSS_Encode_maskedDB_len: assumes "maskedDB = PKCS1_EMSA_PSS_Encode_maskedDB DB dbMask" "length DB = x" "length dbMask = x" shows "length maskedDB = x" by (simp add: PKCS1_EMSA_PSS_Encode_maskedDB assms xor_words_length) lemma EMSA_PSS_Encode_maskedDB_sethi_valid: assumes "octets_valid maskedDB" shows "octets_valid (SetLeftmostBits emLen emBits maskedDB)" by (simp add: SetLeftmostBits_valid assms) lemma EMSA_PSS_Encode_maskedDB_sethi_len: "length (SetLeftmostBits emLen emBits maskedDB) = length maskedDB" by (simp add: SetLeftmostBits_len) lemma EMSA_PSS_Encode_maskedDB_sethi_hd: assumes "maskedDB_sethi = SetLeftmostBits emLen emBits maskedDB" "maskedDB \ []" shows "hd maskedDB_sethi < 2^(8- (8*emLen - emBits))" using SetLeftmostBits_hd assms by presburger lemma EMSA_PSS_Encode_EM_valid: assumes "octets_valid maskedDB" "octets_valid H" "EM = PKCS1_EMSA_PSS_Encode_EM maskedDB H" shows "octets_valid EM" proof - have "(0xbc::nat) < 256" by auto then show ?thesis using words_valid_concat PKCS1_EMSA_PSS_Encode_EM_def assms words_valid_cons words_valid_nil Twoto8 by presburger qed lemma EMSA_PSS_Encode_EM_len: assumes "length maskedDB = emLen - hLen - 1" "length H = hLen" "EM = PKCS1_EMSA_PSS_Encode_EM maskedDB H" "PKCS1_EMSA_PSS_Encode_inputValid sLen emLen" shows "length EM = emLen" proof - have 1: "hLen + sLen + 2 \ emLen" using assms(4) PKCS1_EMSA_PSS_Encode_inputValid_def by blast have 2: "hLen + 1 < emLen" using 1 by auto have 3: "length EM = (length maskedDB) + (length H) + 1" using assms(3) PKCS1_EMSA_PSS_Encode_EM_def by simp show ?thesis using 2 3 assms(1,2) by fastforce qed text \In the following, we collect up all the facts we know about the intermediate values computed in the EMAS_PSS_Encode operation.\ lemma EMSA_PSS_Encode_IntVals: assumes "PKCS1_EMSA_PSS_Encode_inputValid sLen emLen" "octets_valid salt" "emLen = numBits_to_numOctets emBits" "mHash = Hash M" "sLen = length salt" "M' = PKCS1_EMSA_PSS_Encode_M' mHash salt" "H = PKCS1_EMSA_PSS_Encode_H M'" "PS = PKCS1_EMSA_PSS_Encode_PS emLen sLen" "DB = PKCS1_EMSA_PSS_Encode_DB PS salt" "dbMask = PKCS1_EMSA_PSS_Encode_dbMask H emLen" "maskedDB = PKCS1_EMSA_PSS_Encode_maskedDB DB dbMask" "maskedDB' = SetLeftmostBits emLen emBits maskedDB" "DB' = SetLeftmostBits emLen emBits DB" "dbMask' = SetLeftmostBits emLen emBits dbMask" "EM = PKCS1_EMSA_PSS_Encode_EM maskedDB' H" shows "octets_valid mHash \ length mHash = hLen \ octets_valid M' \ length M' = 8 + hLen + sLen \ octets_valid H \ length H = hLen \ octets_valid PS \ length PS = emLen - sLen - hLen - 2 \ octets_valid DB \ length DB = emLen - hLen - 1 \ octets_valid dbMask \ length dbMask = emLen - hLen - 1 \ octets_valid maskedDB \ length maskedDB = emLen - hLen - 1 \ octets_valid maskedDB' \ length maskedDB' = emLen - hLen - 1 \ octets_valid EM \ length EM = emLen \ EM = PKCS1_EMSA_PSS_Encode M salt emBits \ last EM = 0xbc \ bit_length (PKCS1_OS2IP EM) \ emBits" proof - have mH1: "octets_valid mHash" by (simp add: assms(4) hash_valid) have mH2: "length mHash = hLen" by (simp add: assms(4) hash_len) have M1: "octets_valid M'" using EMSA_PSS_Encode_M'_valid assms(2,6) mH1 by presburger have M2: "length M' = 8 + hLen + sLen" using EMSA_PSS_Encode_M'_len assms(5,6) mH2 by blast have H1: "octets_valid H" by (simp add: EMSA_PSS_Encode_H_valid assms(7)) have H2: "length H = hLen" by (simp add: EMSA_PSS_Encode_H_len assms(7)) have PS1: "octets_valid PS" by (simp add: EMSA_PSS_Encode_PS_valid assms(8)) have PS2: "length PS = emLen - sLen - hLen - 2" by (simp add: EMSA_PSS_Encode_PS_len assms(8)) have PS3: "octets_to_nat PS = 0" by (metis assms(8) PKCS1_EMSA_PSS_Encode_PS words_to_nat_prepend_zeros append_Nil2 words_to_nat_empty) have DB1: "octets_valid DB" using EMSA_PSS_Encode_DB_valid PS1 assms(2,9) by blast have DB2: "length DB = emLen - hLen - 1" using EMSA_PSS_Encode_DB_len PS2 assms(1,5,9) by blast have DB3: "DB \ []" by (metis Nil_is_append_conv PKCS1_EMSA_PSS_Encode_DB assms(9) list.distinct(1)) have DB4: "hd DB' < 2^(8 - (8*emLen - emBits))" using DB3 SetLeftmostBits_hd assms(13) by presburger have DB5: "DB = PS @ [1] @ salt" by (simp add: PKCS1_EMSA_PSS_Encode_DB assms(9)) have DB6: "take (emLen - hLen - sLen -1) DB = PS @ [1]" using DB5 PS2 DB2 assms(5) by auto have DB7: "PKCS1_OS2IP (PS @ [1]) = 1" using words_to_nat_prepend_zeros assms(8) PKCS1_OS2IP_def one_word_to_nat by (simp add: PKCS1_EMSA_PSS_Encode_PS) have DB8: "PKCS1_EMSA_PSS_Verify_PS_0x01 emLen sLen DB = 1" using DB6 DB7 PKCS1_EMSA_PSS_Verify_PS_0x01 by presburger have DB9: "TestLeftmostBits emLen emBits DB" using EMSA_PSS_Encode_DB_hibits assms(3,8,9) by blast have DB10: "DB = DB'" using DB9 SetTestLeftmostBits assms(13) by simp have dbM1: "octets_valid dbMask" by (simp add: EMSA_PSS_Encode_dbMask_valid assms(10)) have dbM2: "length dbMask = emLen - hLen - 1" using EMSA_PSS_Encode_dbMask_len assms(10) by presburger have dbM3: "dbMask \ []" using DB2 DB3 dbM2 by force have dbM4: "hd dbMask' < 2^(8 - (8*emLen - emBits))" using dbM3 SetLeftmostBits_hd assms(14) by presburger have mDB1: "octets_valid maskedDB" using assms(11) EMSA_PSS_Encode_maskedDB_valid DB1 dbM1 by blast have mDB2: "length maskedDB = emLen - hLen - 1" using assms(11) EMSA_PSS_Encode_maskedDB_len DB2 dbM2 by blast have mDB3: "maskedDB \ []" using DB2 DB3 mDB2 by fastforce have mDBsh1: "octets_valid maskedDB'" by (simp add: EMSA_PSS_Encode_maskedDB_sethi_valid assms(12) mDB1) have mDBsh2: "length maskedDB' = emLen - hLen - 1" using EMSA_PSS_Encode_maskedDB_sethi_len assms(12) mDB2 by presburger have mDBsh3: "hd maskedDB' < 2^(8 - (8*emLen - emBits))" using assms(12) mDB3 EMSA_PSS_Encode_maskedDB_sethi_hd by blast have mDBsh4: "maskedDB = xor_octets DB dbMask" by (simp add: PKCS1_EMSA_PSS_Encode_maskedDB assms(11)) have mDBsh5: "length DB = length dbMask" using DB2 dbM2 by argo have mDBsh6: "maskedDB' = xor_octets DB' dbMask'" using assms(12,13,14) mDBsh4 mDBsh5 SetLeftmostBits_xor[of DB dbMask emLen emBits] by fast have mDBsh7: "maskedDB' = xor_octets DB dbMask'" using mDBsh6 DB10 by blast have mDBsh8: "TestLeftmostBits emLen emBits maskedDB'" using TestLeftmostBits_def mDBsh3 by presburger have EM1: "octets_valid EM" using EMSA_PSS_Encode_EM_valid H1 assms(15) mDBsh1 by blast have EM2: "length EM = emLen" using EMSA_PSS_Encode_EM_len H2 assms(1,15) mDBsh2 by blast have EM3: "EM = PKCS1_EMSA_PSS_Encode M salt emBits" by (metis assms(3,4,5,6,7,8,9,10,11,12,15) PKCS1_EMSA_PSS_Encode_def) have EM4: "last EM = 0xbc" by (simp add: PKCS1_EMSA_PSS_Encode_EM_def assms(15)) have EM5: "EM \ []" using DB2 DB3 EM2 by force have EM6: "hd EM = hd maskedDB'" by (metis DB2 DB3 PKCS1_EMSA_PSS_Encode_EM_def assms(15) hd_append2 length_0_conv mDBsh2) have EM7: "TestLeftmostBits emLen emBits EM" using EM6 TestLeftmostBits_def mDBsh3 by presburger let ?n = "octets_to_nat EM" have n: "?n = PKCS1_OS2IP EM" using PKCS1_OS2IP_def by presburger have EM8: "bit_length (?n) \ emBits" using setleftmost_bit_len3[of emLen emBits EM ?n] assms(3) EM1 EM2 EM5 EM7 by presburger have EM9: "bit_length (PKCS1_OS2IP EM) \ emBits" using n EM8 by presburger show ?thesis using mH1 mH2 M1 M2 H1 H2 PS1 PS2 DB1 DB2 DB4 DB8 DB9 DB10 dbM1 dbM2 dbM4 mDB1 mDB2 mDB3 mDBsh1 mDBsh2 mDBsh3 mDBsh6 mDBsh7 mDBsh8 EM1 EM2 EM3 EM4 EM7 EM9 by blast qed lemma EMSA_PSS_Encode_Verify_IntValsMatch: assumes "PKCS1_EMSA_PSS_Encode_inputValid sLen emLen" "octets_valid salt" "emLen = numBits_to_numOctets emBits" "mHash = Hash M" "sLen = length salt" "M' = PKCS1_EMSA_PSS_Encode_M' mHash salt" "H = PKCS1_EMSA_PSS_Encode_H M'" "PS = PKCS1_EMSA_PSS_Encode_PS emLen sLen" "DB = PKCS1_EMSA_PSS_Encode_DB PS salt" "DB' = SetLeftmostBits emLen emBits DB" "dbMask = PKCS1_EMSA_PSS_Encode_dbMask H emLen" "dbMask' = SetLeftmostBits emLen emBits dbMask" "maskedDB = PKCS1_EMSA_PSS_Encode_maskedDB DB dbMask" "maskedDB' = SetLeftmostBits emLen emBits maskedDB" "EM = PKCS1_EMSA_PSS_Encode_EM maskedDB' H" "vmaskedDB = PKCS1_EMSA_PSS_Verify_maskedDB emLen EM" "vmaskedDB' = SetLeftmostBits emLen emBits vmaskedDB" "vH = PKCS1_EMSA_PSS_Verify_H emLen EM" "vdbMask = PKCS1_EMSA_PSS_Verify_dbMask vH emLen" "vdbMask' = SetLeftmostBits emLen emBits vdbMask" "vDB = PKCS1_EMSA_PSS_Verify_DB vmaskedDB vdbMask" "vDB' = SetLeftmostBits emLen emBits vDB" "PS_0x01 = PKCS1_EMSA_PSS_Verify_PS_0x01 emLen sLen vDB'" "vsalt = PKCS1_EMSA_PSS_Verify_salt emLen sLen vDB'" "vM' = PKCS1_EMSA_PSS_Encode_M' mHash vsalt" "vH' = PKCS1_EMSA_PSS_Encode_H vM'" shows "maskedDB' = vmaskedDB \ H = vH \ dbMask = vdbMask \ DB = vDB' \ PS_0x01 = 1 \ salt = vsalt \ M' = vM' \ vH' = H" proof - have EM1: "EM = maskedDB' @ H @ [0xbc]" using PKCS1_EMSA_PSS_Encode_EM_def assms(15) by presburger have mDB1: "length maskedDB = emLen - hLen - 1" using EMSA_PSS_Encode_IntVals assms(1,2,3,4,5,6,7,8,9,11,13) by blast have mDB2: "maskedDB' = vmaskedDB" by (simp add: EM1 PKCS1_EMSA_PSS_Verify_maskedDB SetLeftmostBits_len assms(14,16) mDB1) have mDB3: "vmaskedDB = vmaskedDB'" using mDB2 SetLeftmostBits_idem assms(14,17) by auto have H1: "H = vH" by (simp add: EM1 PKCS1_EMSA_PSS_Encode_H PKCS1_EMSA_PSS_Verify_H SetLeftmostBits_len assms(7,14,18) hash_len mDB1) have dbM1: "dbMask = vdbMask" using H1 PKCS1_EMSA_PSS_Encode_dbMask PKCS1_EMSA_PSS_Verify_dbMask assms(11,19) by presburger have dbM2: "dbMask' = vdbMask'" using dbM1 assms(12,20) by blast have DB0: "DB = DB'" by (simp add: EMSA_PSS_Encode_DB_hibits SetTestLeftmostBits assms(3,8,9,10)) have DB1: "maskedDB = xor_octets DB dbMask" by (simp add: PKCS1_EMSA_PSS_Encode_maskedDB assms(13)) have DB2: "length DB = length dbMask" using EMSA_PSS_Encode_IntVals EMSA_PSS_Encode_dbMask_len assms(1,2,3,5,8,9,11) by presburger have DB3: "DB = xor_octets maskedDB dbMask" using DB1 DB2 words_xor_inv xor_words_symm by presburger have DB4: "DB' = xor_octets maskedDB' dbMask'" by (simp add: EMSA_PSS_Encode_dbMask_len SetLeftmostBits_xor assms(10,11,12,14) mDB1 DB3) have DB5: "DB' = xor_octets vmaskedDB' vdbMask'" using DB4 mDB2 mDB3 dbM1 dbM2 by argo have DB6: "vDB = xor_octets vmaskedDB vdbMask" by (simp add: PKCS1_EMSA_PSS_Verify_DB assms(21)) have DB7: "vDB' = xor_octets vmaskedDB' vdbMask'" using DB6 SetLeftmostBits_xor DB2 EMSA_PSS_Encode_maskedDB_len SetLeftmostBits_len assms(13,14,17,20,22) dbM1 mDB2 by auto have DB8: "DB' = vDB'" using DB5 DB7 by fast have DB9: "DB = vDB'" using DB0 DB8 by fast have DB10: "DB = PS @ [1] @ salt" by (simp add: PKCS1_EMSA_PSS_Encode_DB assms(9)) have s: "salt = vsalt" by (metis (no_types, lifting) EMSA_PSS_Encode_maskedDB_len PKCS1_EMSA_PSS_Verify_salt add.commute add_diff_cancel_right' append.assoc append_eq_conv_conj diff_diff_left length_append assms(5,13,24) DB2 mDB1 DB9 DB10) have M: "M' = vM'" using s assms(6,25) by metis have H1: "vH' = H" using M assms(7,26) by argo have H2: "H = vH" by (simp add: EM1 PKCS1_EMSA_PSS_Encode_H PKCS1_EMSA_PSS_Verify_H SetLeftmostBits_len assms(7,14,18) hash_len mDB1) have PS1: "PS_0x01 = PKCS1_OS2IP(PS @ [1])" by (smt (verit, del_insts) DB10 DB2 DB9 mDB1 assms(5,13,23) EMSA_PSS_Encode_maskedDB_len PKCS1_EMSA_PSS_Verify_PS_0x01 add_diff_cancel_right' append.assoc append_eq_conv_conj cancel_ab_semigroup_add_class.diff_right_commute length_append ) have PS2: "PS_0x01 = 1" using PS1 by (simp add: PKCS1_EMSA_PSS_Encode_PS PKCS1_OS2IP_def assms(8) words_to_nat_prepend_zeros) show ?thesis using H1 H2 M s mDB2 dbM1 DB9 PS2 by blast qed lemma PKCS1_EMSA_PSS_Encode_Verify: assumes "EM = PKCS1_EMSA_PSS_Encode M salt emBits" "sLen = length salt" "octets_valid salt" "PKCS1_EMSA_PSS_Encode_inputValid sLen emLen" "emLen = numBits_to_numOctets emBits" shows "PKCS1_EMSA_PSS_Verify sLen M EM emBits" proof - have EM1: "length EM = emLen" using assms EMSA_PSS_Encode_IntVals by force have l: "hLen + sLen + 2 \ emLen" using assms(4) PKCS1_EMSA_PSS_Encode_inputValid_def by blast have EM2: "last EM = 0xbc" using EMSA_PSS_Encode_IntVals assms by force let ?mHash = "Hash M" let ?maskedDB = "PKCS1_EMSA_PSS_Verify_maskedDB emLen EM" let ?H = "PKCS1_EMSA_PSS_Verify_H emLen EM" let ?dbMask = "PKCS1_EMSA_PSS_Verify_dbMask ?H emLen" let ?DB = "PKCS1_EMSA_PSS_Verify_DB ?maskedDB ?dbMask" let ?DB' = "SetLeftmostBits emLen emBits ?DB" let ?PS_0x01 = "PKCS1_EMSA_PSS_Verify_PS_0x01 emLen sLen ?DB'" let ?salt = "PKCS1_EMSA_PSS_Verify_salt emLen sLen ?DB'" let ?M' = "PKCS1_EMSA_PSS_Encode_M' ?mHash salt" let ?H' = "PKCS1_EMSA_PSS_Encode_H ?M'" have PS: "?PS_0x01 = 1" by (metis EMSA_PSS_Encode_Verify_IntValsMatch PKCS1_EMSA_PSS_Encode_def assms) have H: "?H = ?H'" by (metis (no_types) EMSA_PSS_Encode_Verify_IntValsMatch PKCS1_EMSA_PSS_Encode_def assms) have mDB: "TestLeftmostBits emLen emBits ?maskedDB" by (metis assms TestLeftmostBits_def EMSA_PSS_Encode_Verify_IntValsMatch PKCS1_EMSA_PSS_Encode_def SetLeftmostBits_hd SetTestLeftmostBits) show ?thesis using EM1 EM2 l PS H mDB PKCS1_EMSA_PSS_Verify_def[of sLen M EM emBits] by (metis EMSA_PSS_Encode_Verify_IntValsMatch PKCS1_EMSA_PSS_Encode_def assms) qed end (* of locale EMSA *) subsection \Section 9.2: EMSA-PKCS1-v1_5\ text \"This encoding method is deterministic and only has an encoding operation." The algorithmID for the nine hash functions mentioned in Appendix B.1 are given on page 40 of the standard. We include them in section B.1 of this theory. \ locale EMSA_v1_5 = fixes Hash :: hash_type and hLen :: nat and algorithmID :: octets assumes hash_valid: "\x. octets_valid (Hash x)" and hash_len: "\x. length (Hash x) = hLen" and IDvalid: "octets_valid algorithmID" begin definition PKCS1_EMSA_v1_5_Encode_inputValid :: "octets \ nat \ bool" where "PKCS1_EMSA_v1_5_Encode_inputValid M emLen = ( let H = Hash M; T = algorithmID @ H; tLen = length T; PS = replicate (emLen - tLen - 3) (0xff::nat) in (tLen + 11 \ emLen) \ (8 \ length PS) )" definition PKCS1_EMSA_v1_5_Encode :: "octets \ nat \ octets" where "PKCS1_EMSA_v1_5_Encode M emLen = ( let H = Hash M; T = algorithmID @ H; tLen = length T; PS = replicate (emLen - tLen - 3) (0xff::nat) in [0, 1] @ PS @ [0] @ T )" text \Only having an encoding method means that there's not a whole lot that we might prove about it. The one thing we can show is that given valid input, the length of the result of decoding is known. We also know that the encoded message has valid octets.\ lemma PKCS1_EMSA_v1_5_Encode_valid_len: assumes "PKCS1_EMSA_v1_5_Encode_inputValid M emLen" "EM = PKCS1_EMSA_v1_5_Encode M emLen" shows "length EM = emLen" proof - let ?H = "Hash M" let ?T = "algorithmID @ ?H" let ?tLen = "length ?T" let ?PS = "replicate (emLen - ?tLen - 3) (0xff::nat)" let ?psLen = "length ?PS" have 1: "(?tLen + 11 \ emLen) \ (8 \ ?psLen)" using assms(1) PKCS1_EMSA_v1_5_Encode_inputValid_def by meson have 2: "EM = [0, 1] @ ?PS @ [0] @ ?T" using assms(2) PKCS1_EMSA_v1_5_Encode_def by meson have 3: "length EM = 3 + ?psLen + ?tLen" using 2 by auto have 4: "?psLen = emLen - ?tLen - 3" by simp show ?thesis using 1 3 4 by presburger qed lemma PKCS1_EMSA_v1_5_Encode_octetsValid: assumes "EM = PKCS1_EMSA_v1_5_Encode M emLen" shows "octets_valid EM" proof - let ?H = "Hash M" let ?T = "algorithmID @ ?H" let ?tLen = "length ?T" let ?PS = "replicate (emLen - ?tLen - 3) (0xff::nat)" have 1: "EM = [0, 1] @ ?PS @ [0] @ ?T" by (metis (no_types) PKCS1_EMSA_v1_5_Encode_def assms) have 2: "octets_valid ?PS" by (simp add: words_valid_def) have 3: "octets_valid [0,1] \ octets_valid [0]" by (simp add: words_valid_def) have 4: "octets_valid ?T" by (simp add: IDvalid hash_valid words_valid_concat) show ?thesis using 1 2 3 4 words_valid_concat by presburger qed text \Well, I suppose we could unwrap the encoding. It's obvious so maybe that's why the standard did not bother to write it down. We provide two different forms.\ lemma PKCS1_EMSA_v1_5_Decode: assumes "EM = PKCS1_EMSA_v1_5_Encode M emLen" shows "Hash M = drop (1+(length algorithmID)) (words_strip_nonzero (drop 2 EM))" proof - let ?H = "Hash M" let ?T = "algorithmID @ ?H" let ?tLen = "length ?T" let ?PS = "replicate (emLen - ?tLen - 3) (0xff::nat)" have 0: "EM = [0, 1] @ ?PS @ [0] @ ?T" using assms PKCS1_EMSA_v1_5_Encode_def by meson have 1: "\p \ set ?PS. 0 < p" by fastforce have 2: "drop 2 EM = ?PS @ [0] @ ?T" using 0 by auto have 3: "words_strip_nonzero (drop 2 EM) = [0] @ ?T" using 2 strip_non0_concat by force show ?thesis using 3 by simp qed lemma PKCS1_EMSA_v1_5_Decode2: assumes "EM = PKCS1_EMSA_v1_5_Encode M emLen" "PKCS1_EMSA_v1_5_Encode_inputValid M emLen" "H = Hash M" "hLen = length H" shows "H = drop (emLen - hLen) EM" proof - let ?aLen = "length algorithmID" let ?T = "algorithmID @ H" let ?tLen = "length ?T" have T: "?tLen = ?aLen + hLen" using assms(4) by simp let ?PS = "replicate (emLen - ?tLen - 3) (0xff::nat)" have EM1: "EM = [0, 1] @ ?PS @ [0] @ algorithmID @ H" using assms(1,3) PKCS1_EMSA_v1_5_Encode_def by meson have EM2: "?tLen + 11 \ emLen" using assms(2,3) PKCS1_EMSA_v1_5_Encode_inputValid_def by meson have EM3: "?aLen + hLen + 11 \ emLen" using T EM2 by presburger have EM4: "?aLen + hLen + 3 < emLen" using EM3 by auto have PS: "length ?PS = emLen - ?aLen - hLen - 3" using T length_replicate by force have 1: "length ([0, 1] @ ?PS @ [0] @ algorithmID) = 2 + (emLen - ?aLen - hLen - 3) + 1 + ?aLen" by (metis PS assms(4) One_nat_def Suc_eq_plus1 ab_semigroup_add_class.add_ac(1) length_append list.size(3,4) numeral_2_eq_2) have 2: "length ([0, 1] @ ?PS @ [0] @ algorithmID) = emLen - hLen" using 1 EM4 by simp show ?thesis using 2 EM1 by (smt (verit, ccfv_threshold) append.assoc cancel_comm_monoid_add_class.diff_cancel drop0 drop_append eq_Nil_appendI length_0_conv length_drop) qed end (* of locale EMSA_v1_5 *) section \Section 8: Signature Schemes\ text\"For the purposes of this document, a signature scheme with appendix consists of a signature generation operation and a signature verification operation, where the signature generation operation produces a signature from a message with a signer's RSA private key, and the signature verification operation verifies the signature on the message with the signer's corresponding RSA public key. To verify a signature constructed with this type of scheme it is necessary to have the message itself. In this way, signature schemes with appendix are distinguished from signature schemes with message recovery, which are not supported in this document. Two signature schemes with appendix are specified in this document: RSASSA-PSS and RSASSA-PKCS1-v1_5. Although no attacks are known against RSASSA-PKCS1-v1_5, in the interest of increased robustness, RSASSA-PSS is required in new applications. RSASSA-PKCS1-v1_5 is included only for compatibility with existing applications."\ subsection \Section 8.1: RSASSA-PSS\ text\"RSASSA-PSS combines the RSASP1 and RSAVP1 primitives with the EMSA-PSS encoding method. It is compatible with the IFSSA scheme as amended in the IEEE 1363a-2004, where the signature and verification primitives are IFSP-RSA1 and IFVP-RSA1 as defined in IEEE 1363-2000 and the message encoding method is EMSA4. EMSA4 is slightly more general than EMSA-PSS as it acts on bit strings rather than on octet strings. EMSA-PSS is equivalent to EMSA4 restricted to the case that the operands as well as the hash and salt values are octet strings."\ locale RSASSA_PSS = EMSA_PSS + fixes RSAENCRYPT RSADECRYPT :: "nat \ nat" and n :: nat assumes RSA_mod: "0 < n" and RSA_bnd: "\m. (RSAENCRYPT m) < n" "\c. (RSADECRYPT c) < n" and RSA_inv: "\m. (m < n) \ RSADECRYPT (RSAENCRYPT m) = m" "\c. (c < n) \ RSAENCRYPT (RSADECRYPT c) = c" begin definition k : "k = octet_length n" definition modBits : "modBits = bit_length n" subsubsection \Section 8.1.1: Signature Generation Operation\ definition PKCS1_RSASSA_PSS_Sign_inputValid: "PKCS1_RSASSA_PSS_Sign_inputValid salt = ( let sLen = length salt; emLen = numBits_to_numOctets (modBits-1) in (PKCS1_EMSA_PSS_Encode_inputValid sLen emLen) \ (octets_valid salt))" definition PKCS1_RSASSA_PSS_Sign: "PKCS1_RSASSA_PSS_Sign M salt = ( let EM = PKCS1_EMSA_PSS_Encode M salt (modBits-1); m = PKCS1_OS2IP EM; s = RSADECRYPT m in PKCS1_I2OSP s k)" subsubsection \Section 8.1.2: Signature Verification Operation\ definition PKCS1_RSASSA_PSS_Verify : "PKCS1_RSASSA_PSS_Verify M S sLen = ( let s = PKCS1_OS2IP S; m = RSAENCRYPT s; emLen = numBits_to_numOctets (modBits - 1); EM = PKCS1_I2OSP m emLen in (length S = k) \ (s < n) \ (PKCS1_I2OSP_inputValid m emLen) \ (PKCS1_EMSA_PSS_Verify sLen M EM (modBits - 1)) )" subsubsection \RSASSA_PSS Lemmas\ lemma RSASSA_PSS_IntVals: assumes "PKCS1_RSASSA_PSS_Sign_inputValid salt" "sLen = length salt" "emLen = numBits_to_numOctets (modBits-1)" "EM = PKCS1_EMSA_PSS_Encode M salt (modBits-1)" "m = PKCS1_OS2IP EM" "s = RSADECRYPT m" "S = PKCS1_I2OSP s k" "s' = PKCS1_OS2IP S" "m' = RSAENCRYPT s'" "EM' = PKCS1_I2OSP m' emLen" shows "s = s' \ s < n \ m = m' \ m < n \ EM = EM' \ length EM = emLen \ octets_valid EM \ length S = k \ S = PKCS1_RSASSA_PSS_Sign M salt" proof - have s1: "s = s'" by (simp add: I2OSP_OS2IP assms(7,8)) have s2: "s < n" by (simp add: RSA_bnd(2) assms(6)) have m1: "bit_length m \ (modBits -1)" by (metis assms(1,4,5) EMSA_PSS_Encode_IntVals PKCS1_RSASSA_PSS_Sign_inputValid) have m2: "bit_length m < bit_length n" by (metis m1 RSA_bnd(1) bit_len_zero_eq diff_less leD less_imp_diff_less less_nat_zero_code modBits nat_neq_iff zero_less_one) have m3: "m < n" using m2 word_len_comp by blast have m4: "m = m'" using RSA_inv(2) assms(6,9) m3 s1 by auto have EM1: "length EM = emLen" by (metis EMSA_PSS_Encode_IntVals PKCS1_RSASSA_PSS_Sign_inputValid assms(1,3,4)) have EM2: "EM = EM'" by (metis EM1 EMSA_PSS_Encode_DB_valid EMSA_PSS_Encode_EM_valid EMSA_PSS_Encode_H_valid EMSA_PSS_Encode_PS_valid EMSA_PSS_Encode_dbMask_valid EMSA_PSS_Encode_maskedDB_sethi_valid EMSA_PSS_Encode_maskedDB_valid OS2IP_I2OSP PKCS1_EMSA_PSS_Encode_def PKCS1_RSASSA_PSS_Sign_inputValid assms(1,4,5,10) m4) have EM3: "octets_valid EM" by (simp add: EM2 I2OSP_octets_valid assms(10)) have S1: "length S = k" by (metis I2OSP_valid_len PKCS1_I2OSP_inputValid_def RSA_bnd(2) assms(6,7) k less_trans word_length_eq2 Twoto8 zero_less_numeral) have S2: "S = PKCS1_RSASSA_PSS_Sign M salt" by (simp add: PKCS1_RSASSA_PSS_Sign assms(4,5,6,7)) show ?thesis using s1 s2 m3 m4 EM1 EM2 EM3 S1 S2 by blast qed lemma RSASSA_PSS_SigVerifies: assumes "PKCS1_RSASSA_PSS_Sign_inputValid salt" "S = PKCS1_RSASSA_PSS_Sign M salt" "sLen = length salt" shows "PKCS1_RSASSA_PSS_Verify M S sLen" by (metis assms EMSA_PSS.PKCS1_EMSA_PSS_Encode_Verify EMSA_PSS_axioms PKCS1_I2OSP_inputValid_def PKCS1_OS2IP_def PKCS1_RSASSA_PSS_Verify RSASSA_PSS.PKCS1_RSASSA_PSS_Sign_inputValid Twoto8 RSASSA_PSS_IntVals RSASSA_PSS_axioms words_to_nat_len_bnd words_valid_def zero_less_numeral) end (* of locale RSASSA_PSS *) subsection \Section 8.2: RSASSA-PKCS1-v1_5\ locale RSASSA_v1_5 = EMSA_v1_5 + fixes RSAENCRYPT RSADECRYPT :: "nat \ nat" and n :: nat assumes RSA_mod: "0 < n" and RSA_bnd: "\m. (RSAENCRYPT m) < n" "\c. (RSADECRYPT c) < n" and RSA_inv: "\m. (m < n) \ RSADECRYPT (RSAENCRYPT m) = m" "\c. (c < n) \ RSAENCRYPT (RSADECRYPT c) = c" begin definition k : "k = octet_length n" subsubsection \Section 8.2.1: Signature Generation Operation\ definition PKCS1_RSASSA_v1_5_Sign : "PKCS1_RSASSA_v1_5_Sign M = (let EM = PKCS1_EMSA_v1_5_Encode M k; m = PKCS1_OS2IP EM; s = RSADECRYPT m in PKCS1_I2OSP s k )" subsubsection \Section 8.2.2: Signature Verification Operation\ definition PKCS1_RSASSA_v1_5_Verify : "PKCS1_RSASSA_v1_5_Verify M S = (let s = PKCS1_OS2IP S; m = RSAENCRYPT s; EM = PKCS1_I2OSP m k; EM' = PKCS1_EMSA_v1_5_Encode M k in (length S = k) \ (s < n) \ (PKCS1_I2OSP_inputValid m k) \ (EM = EM') \ (PKCS1_EMSA_v1_5_Encode_inputValid M k) )" subsubsection \RSASSA_v1_5 Lemmas\ lemma RSASSA_v1_5_IntVals: assumes "PKCS1_EMSA_v1_5_Encode_inputValid M k" "EM = PKCS1_EMSA_v1_5_Encode M k" "m = PKCS1_OS2IP EM" "s = RSADECRYPT m" "S = PKCS1_I2OSP s k" "s' = PKCS1_OS2IP S" "m' = RSAENCRYPT s'" "EM' = PKCS1_I2OSP m' k" shows "s = s' \ m = m' \ EM = EM' \ length S = k \ m < n \ s < n \ length EM = k" proof - have s1: "s = s'" by (simp add: I2OSP_OS2IP assms(5,6)) have s2: "s < n" by (simp add: RSA_bnd(2) assms(4)) have s3: "s' < n" using s1 s2 by fast have S: "length S = k" by (metis I2OSP_valid_len PKCS1_I2OSP_inputValid_def assms(5) k less_trans word_length_eq2 s2 Twoto8 zero_less_numeral) have EM1: "length EM = k" using assms(1,2) PKCS1_EMSA_v1_5_Encode_valid_len by blast have EM2: "octets_valid EM" by (simp add: PKCS1_EMSA_v1_5_Encode_octetsValid assms(2)) let ?H = "Hash M" let ?T = "algorithmID @ ?H" let ?tLen = "length ?T" let ?PS = "replicate (k - ?tLen - 3) (0xff::nat)" let ?psLen = "length ?PS" have EM3: "EM = [0, 1] @ ?PS @ [0] @ ?T" by (metis (no_types) PKCS1_EMSA_v1_5_Encode_def assms(2)) have EM4: "words_strip_zero EM = [1] @ ?PS @ [0] @ ?T" using EM3 by simp have EM5: "length (words_strip_zero EM) = k-1" by (metis EM1 EM3 EM4 append_Cons length_tl list.sel(3)) have m1: "m < 256^(k-1)" by (metis EM2 EM5 PKCS1_OS2IP_def assms(3) words_strip0_to_nat_len_bnd words_valid_def Twoto8) have m2: "m < n" by (metis (no_types, lifting) RSA_mod k less_le_trans m1 word_length_eq3 Twoto8 zero_less_numeral) have m3: "m = m'" using RSA_inv(2) m2 assms(4,7) s1 by auto have EM4: "EM = EM'" using EM1 EM2 OS2IP_I2OSP assms(3,8) m3 by fastforce show ?thesis using s1 s2 S m2 m3 EM1 EM4 by blast qed lemma RSASSA_v1_5_SigVerifies: assumes "PKCS1_EMSA_v1_5_Encode_inputValid M k" "S = PKCS1_RSASSA_v1_5_Sign M" shows "PKCS1_RSASSA_v1_5_Verify M S" by (metis PKCS1_I2OSP_inputValid_def PKCS1_RSASSA_v1_5_Sign PKCS1_RSASSA_v1_5_Verify RSASSA_v1_5_IntVals assms k less_trans word_length_eq2 Twoto8 zero_less_numeral) end (* of locale RSASSA_v1_5 *) section \Appendix B: Supporting Techniques\ subsection \Appendix B.1: Hash Functions\ text \"Nine hash functions are given as examples for the encoding methods in this document: MD2, MD5, SHA-1, SHA-224, SHA-256, SHA-384, SHA-512, SHA-512/224, and SHA-512/256. For the RSAES-OAEP encryption scheme and EMSA-PSS encoding method, only SHA-1, SHA-224, SHA-256, SHA-384, SHA-512, SHA-512/224, and SHA-512/256 are recommended. For the EMSA-PKCS1-v1_5 encoding method, SHA-224, SHA-256, SHA-384, SHA-512, SHA-512/224, and SHA-512/256 are recommended for new applications. MD2, MD5 and SHA-1 are recommended only for compatibility with existing applications based on PKCS #1 v1.5."\ datatype HashAlg = tMD2 | tMD5 | tSHA1 | tSHA224 | tSHA256 | tSHA384 | tSHA512 | tSHA512_224 | tSHA512_256 text \This is the DER encoding of the algorithm IDs for the 9 has algorithms given in the appendix of PKCS #1.\ fun PKCS1_AlgorithmID :: "HashAlg \ octets" where "PKCS1_AlgorithmID tMD2 = [0x30, 0x20, 0x30, 0x0c, 0x06, 0x08, 0x2a, 0x86, 0x48, 0x86, 0xf7, 0x0d, 0x02, 0x02, 0x05, 0x00, 0x04, 0x10]" | "PKCS1_AlgorithmID tMD5 = [0x30, 0x20, 0x30, 0x0c, 0x06, 0x08, 0x2a, 0x86, 0x48, 0x86, 0xf7, 0x0d, 0x02, 0x05, 0x05, 0x00, 0x04, 0x10]" | "PKCS1_AlgorithmID tSHA1 = [0x30, 0x21, 0x30, 0x09, 0x06, 0x05, 0x2b, 0x0e, 0x03, 0x02, 0x1a, 0x05, 0x00, 0x04, 0x14]" | "PKCS1_AlgorithmID tSHA224 = [0x30, 0x2d, 0x30, 0x0d, 0x06, 0x09, 0x60, 0x86, 0x48, 0x01, 0x65, 0x03, 0x04, 0x02, 0x04, 0x05, 0x00, 0x04, 0x1c]" | "PKCS1_AlgorithmID tSHA256 = [0x30, 0x31, 0x30, 0x0d, 0x06, 0x09, 0x60, 0x86, 0x48, 0x01, 0x65, 0x03, 0x04, 0x02, 0x01, 0x05, 0x00, 0x04, 0x20]" | "PKCS1_AlgorithmID tSHA384 = [0x30, 0x41, 0x30, 0x0d, 0x06, 0x09, 0x60, 0x86, 0x48, 0x01, 0x65, 0x03, 0x04, 0x02, 0x02, 0x05, 0x00, 0x04, 0x30]" | "PKCS1_AlgorithmID tSHA512 = [0x30, 0x51, 0x30, 0x0d, 0x06, 0x09, 0x60, 0x86, 0x48, 0x01, 0x65, 0x03, 0x04, 0x02, 0x03, 0x05, 0x00, 0x04, 0x40]" | "PKCS1_AlgorithmID tSHA512_224 = [0x30, 0x2d, 0x30, 0x0d, 0x06, 0x09, 0x60, 0x86, 0x48, 0x01, 0x65, 0x03, 0x04, 0x02, 0x05, 0x05, 0x00, 0x04, 0x1c]" | "PKCS1_AlgorithmID tSHA512_256 = [0x30, 0x31, 0x30, 0x0d, 0x06, 0x09, 0x60, 0x86, 0x48, 0x01, 0x65, 0x03, 0x04, 0x02, 0x06, 0x05, 0x00, 0x04, 0x20]" lemma PKCS1_AlgorithmID_Valid: "octets_valid (PKCS1_AlgorithmID X)" apply (cases X) by (simp_all add: words_valid_def) text\This gives the default values for the hash of the label L when no L is given. Because MD2 and MD5 are not recommended for RSA-OAEP or EMSAPSS, there is no lHash given for those hash algorithms.\ fun PKCS1_emptyL_lHash :: "HashAlg \ octets" where "PKCS1_emptyL_lHash tMD2 = []" | "PKCS1_emptyL_lHash tMD5 = []" | "PKCS1_emptyL_lHash tSHA1 = [0xda, 0x39, 0xa3, 0xee, 0x5e, 0x6b, 0x4b, 0x0d, 0x32, 0x55, 0xbf, 0xef, 0x95, 0x60, 0x18, 0x90, 0xaf, 0xd8, 0x07, 0x09]" | "PKCS1_emptyL_lHash tSHA224 = [0xd1, 0x4a, 0x02, 0x8c, 0x2a, 0x3a, 0x2b, 0xc9, 0x47, 0x61, 0x02, 0xbb, 0x28, 0x82, 0x34, 0xc4, 0x15, 0xa2, 0xb0, 0x1f, 0x82, 0x8e, 0xa6, 0x2a, 0xc5, 0xb3, 0xe4, 0x2f]" | "PKCS1_emptyL_lHash tSHA256 = [0xe3, 0xb0, 0xc4, 0x42, 0x98, 0xfc, 0x1c, 0x14, 0x9a, 0xfb, 0xf4, 0xc8, 0x99, 0x6f, 0xb9, 0x24, 0x27, 0xae, 0x41, 0xe4, 0x64, 0x9b, 0x93, 0x4c, 0xa4, 0x95, 0x99, 0x1b, 0x78, 0x52, 0xb8, 0x55]" | "PKCS1_emptyL_lHash tSHA384 = [0x38, 0xb0, 0x60, 0xa7, 0x51, 0xac, 0x96, 0x38, 0x4c, 0xd9, 0x32, 0x7e, 0xb1, 0xb1, 0xe3, 0x6a, 0x21, 0xfd, 0xb7, 0x11, 0x14, 0xbe, 0x07, 0x43, 0x4c, 0x0c, 0xc7, 0xbf, 0x63, 0xf6, 0xe1, 0xda, 0x27, 0x4e, 0xde, 0xbf, 0xe7, 0x6f, 0x65, 0xfb, 0xd5, 0x1a, 0xd2, 0xf1, 0x48, 0x98, 0xb9, 0x5b]" | "PKCS1_emptyL_lHash tSHA512 = [0xcf, 0x83, 0xe1, 0x35, 0x7e, 0xef, 0xb8, 0xbd, 0xf1, 0x54, 0x28, 0x50, 0xd6, 0x6d, 0x80, 0x07, 0xd6, 0x20, 0xe4, 0x05, 0x0b, 0x57, 0x15, 0xdc, 0x83, 0xf4, 0xa9, 0x21, 0xd3, 0x6c, 0xe9, 0xce, 0x47, 0xd0, 0xd1, 0x3c, 0x5d, 0x85, 0xf2, 0xb0, 0xff, 0x83, 0x18, 0xd2, 0x87, 0x7e, 0xec, 0x2f, 0x63, 0xb9, 0x31, 0xbd, 0x47, 0x41, 0x7a, 0x81, 0xa5, 0x38, 0x32, 0x7a, 0xf9, 0x27, 0xda, 0x3e]" | "PKCS1_emptyL_lHash tSHA512_224 = [0x6e, 0xd0, 0xdd, 0x02, 0x80, 0x6f, 0xa8, 0x9e, 0x25, 0xde, 0x06, 0x0c, 0x19, 0xd3, 0xac, 0x86, 0xca, 0xbb, 0x87, 0xd6, 0xa0, 0xdd, 0xd0, 0x5c, 0x33, 0x3b, 0x84, 0xf4]" | "PKCS1_emptyL_lHash tSHA512_256 = [0xc6, 0x72, 0xb8, 0xd1, 0xef, 0x56, 0xed, 0x28, 0xab, 0x87, 0xc3, 0x62, 0x2c, 0x51, 0x14, 0x06, 0x9b, 0xdd, 0x3a, 0xd7, 0xb8, 0xf9, 0x73, 0x74, 0x98, 0xd0, 0xc0, 0x1e, 0xce, 0xf0, 0x96, 0x7a]" lemma PKCS1_emptyL_lHash_Valid: "octets_valid (PKCS1_emptyL_lHash X)" apply (cases X) by (simp_all add: words_valid_def) subsection \Appendix B.2: Mask Generation Functions\ subsubsection \Appendix B.2.1: MGF1\ text\ "Options: Hash hash function (hLen denotes the length in octets of the hash function output) Input: mgfSeed seed from which mask is generated, an octet string* maskLen intended length in octets of the mask, at most (2^32)*hLen"\ named_theorems MGF1thms definition PKCS1_MGF1_maskLenValid :: "nat \ nat \ bool" where "PKCS1_MGF1_maskLenValid maskLen hLen = (maskLen \ (2^32)*hLen)" fun PKCS1_MGF1_rec :: "hash_type \ octets \ nat \ octets" where "PKCS1_MGF1_rec Hash mgfSeed 0 = Hash ( mgfSeed @ (PKCS1_I2OSP 0 4) )" | "PKCS1_MGF1_rec Hash mgfSeed (Suc n) = (PKCS1_MGF1_rec Hash mgfSeed n) @ (Hash (mgfSeed @ (PKCS1_I2OSP (Suc n) 4)))" definition PKCS1_MGF1_MaxCounter :: "nat \ nat \ nat" where "PKCS1_MGF1_MaxCounter maskLen hLen = ( if (maskLen mod hLen = 0) then (maskLen div hLen) - 1 else (maskLen div hLen) )" definition PKCS1_MGF1 :: "hash_type \ nat \ octets \ nat \ octets" where "PKCS1_MGF1 Hash hLen mgfSeed maskLen = take maskLen ( PKCS1_MGF1_rec Hash mgfSeed (PKCS1_MGF1_MaxCounter maskLen hLen) )" text \What we want to know about any mask generating function is that it produces valid octets (meaning the value of every octet is < 256) and that the length of the mask produced is given by the input maskLen. These facts are true for MGF1 assuming that the underlying Hash produces valid octets, and the hash of any input has length hLen where 0 < hLen.\ lemma MGF1_rec_valid: assumes "\x. octets_valid (Hash x)" shows "octets_valid (PKCS1_MGF1_rec Hash mgfSeed n)" using assms proof (induction n) case 0 then show ?case by simp next case (Suc n) then show ?case by (simp add: words_valid_concat) qed text \Here is the first of the two things we want to know about MGF1. It produces valid octets, assuming that the underlying hash produces valid octets.\ lemma MGF1_valid [MGF1thms]: assumes "\x. octets_valid (Hash x)" shows "octets_valid (PKCS1_MGF1 Hash hLen mgfSeed maskLen)" using MGF1_rec_valid PKCS1_MGF1_def assms words_valid_take by presburger lemma MGF1_rec_len: assumes "\x. length (Hash x) = hLen" shows "length (PKCS1_MGF1_rec Hash mgfSeed n) = (n+1)*hLen" using assms proof (induction n) case 0 then show ?case by simp next case (Suc n) then show ?case by (metis PKCS1_MGF1_rec.simps(2) Suc_eq_plus1 comm_monoid_mult_class.mult_1 comm_semiring_class.distrib length_append) qed lemma MaxCounter_zero: assumes "m < h" shows "PKCS1_MGF1_MaxCounter m h = 0" by (simp add: PKCS1_MGF1_MaxCounter_def assms) lemma MaxCounter_zero2: "PKCS1_MGF1_MaxCounter h h = 0" by (metis PKCS1_MGF1_MaxCounter_def Suc_diff_1 div_0 lessI less_one mod_mult_self2_is_0 mult.commute nat_mult_1 nonzero_mult_div_cancel_left zero_diff) lemma MGF1_len_h: assumes "\x. length (Hash x) = hLen" "0 < hLen" shows "maskLen \ length (PKCS1_MGF1_rec Hash mgfSeed (PKCS1_MGF1_MaxCounter maskLen hLen))" proof (cases) assume 10: "maskLen \ hLen" have 11: "PKCS1_MGF1_MaxCounter maskLen hLen = 0" by (metis 10 MaxCounter_zero MaxCounter_zero2 le_eq_less_or_eq) have 12: "length (PKCS1_MGF1_rec Hash mgfSeed (PKCS1_MGF1_MaxCounter maskLen hLen)) = hLen" using 11 MGF1_rec_len assms(1) by presburger show ?thesis using 10 12 by argo next assume "\ maskLen \ hLen" then have 20: "hLen < maskLen" by auto have 21: "1 \ maskLen div hLen" by (metis One_nat_def Suc_leI assms(2) 20 div_greater_zero_iff less_imp_le_nat) let ?n = "PKCS1_MGF1_MaxCounter maskLen hLen" show ?thesis proof (cases "maskLen mod hLen = 0") case True then have "?n + 1 = maskLen div hLen" by (metis 21 One_nat_def PKCS1_MGF1_MaxCounter_def Suc_diff_1 add.right_neutral add_Suc_right less_le_trans zero_less_one) then have "length (PKCS1_MGF1_rec Hash mgfSeed ?n) = maskLen" by (simp add: MGF1_rec_len True assms(1) mod_0_imp_dvd) then show ?thesis by simp next case 30: False let ?x = "maskLen div hLen" let ?y = "maskLen mod hLen" have 31: "maskLen = ?x * hLen + ?y" by presburger have 32: "?n + 1 = ?x + 1" using 30 PKCS1_MGF1_MaxCounter_def by presburger have 33: "(?n + 1) * hLen = ?x*hLen + hLen" by (metis 32 Suc_eq_plus1 add.commute mult.commute mult_Suc_right) have 34: "?y < hLen" using assms(2) mod_less_divisor by blast have 35: "?x * hLen + ?y < ?x*hLen + hLen" using 34 by linarith have 36: "maskLen < (?n + 1) * hLen" using 33 35 by presburger then show ?thesis by (metis MGF1_rec_len assms(1) less_imp_le) qed qed text \Here is the second of the two things we want to know about MGF1. It produces an octet string of length maskLen (the final input value), as long as the two conditions are true of the underlying Hash.\ lemma MGF1_len [MGF1thms]: assumes "\x. length (Hash x) = hLen" "0 < hLen" shows "length (PKCS1_MGF1 Hash hLen mgfSeed maskLen) = maskLen" proof - have "maskLen \ length (PKCS1_MGF1_rec Hash mgfSeed (PKCS1_MGF1_MaxCounter maskLen hLen))" using assms MGF1_len_h by blast then show ?thesis using PKCS1_MGF1_def by simp qed end diff --git a/thys/Crypto_Standards/SEC1v2_0.thy b/thys/Crypto_Standards/SEC1v2_0.thy --- a/thys/Crypto_Standards/SEC1v2_0.thy +++ b/thys/Crypto_Standards/SEC1v2_0.thy @@ -1,3407 +1,3407 @@ theory SEC1v2_0 imports "Words" "EC_Common" begin text \ https://www.secg.org/sec1-v2.pdf SEC 1 v2.0 is the current standard from the Standards for Efficient Cryptography Group (SECG) for elliptic curve cryptography (ECC). The text of the standard is available at the above link. In this theory, we translate the functions defined in SEC 1 v2.0 to HOL. We adhere as closely to the written standard as possible, including function and variable names. It should be clear to anyone, regardless of their experience with HOL, that this translation exactly matches the standard. Variances from the written standard are (over-)explained in comments. Note: We only consider curves of the form \y\<^sup>2 \ x\<^sup>3 + ax + b (mod p)\ over a finite prime field \F\<^sub>p\, 2 < p, where \4a\<^sup>3 + 27b\<^sup>2 \ 0 (mod p)\. The standard also considers curves over finite fields of characteristic 2. The residues_prime_gt2 locale fixes a prime p, where 2 < p, and the corresponding residue ring labeled R. That is, \R = \\<^sub>p\, the integers modulo p. Note that SEC 1 sometimes uses the letter R to represent points on an elliptic curve. Because R is used for the residue ring here, we will need to choose another letter when SEC 1 uses R. Otherwise we strive to keep the same variable names as the standard.\ context residues_prime_gt2 begin section \2. Mathematical Foundations\ text \The majority of this chapter of the standard is, as the title suggests, background information about finite fields and elliptic curves. HOL/Isabelle has well-built theories of finite fields and elliptic curves. We don't need to translate any of that now. We use what HOL already has. There are a few things in Section 2.3 Data Types and Conversions which need translation. See below.\ subsection \2.1-2.2 Finite Fields and Elliptic Curves\ text \The language used in the HOL formalization of finite fields might be opaque to non-mathematicians. We use this section only to point out things to the reader who might appreciate some reminders on the language used for finite fields and elliptic curves. First we state a few simple things about the prime p that will be useful in some proof below. For example, that p is odd and that the number of octets (8-bit bytes) needed to represent p is greater than 0. The third lemma here is useful when thinking about the functions that convert between points in the elliptic curve and a string of octets.\ lemma p_mod2: "p mod 2 = 1" using gt2 p_prime prime_odd_nat by presburger lemma octet_len_p: "0 < octet_length p" by (metis gt2 nat_to_words_nil_iff_zero2 neq0_conv not_less_iff_gr_or_eq zero_less_numeral) lemma octet_len_p': assumes "l = octet_length p" shows "1 \ l + 1 \ 1 \ 2*l + 1 \ l + 1 \ 2*l + 1" using assms octet_len_p by linarith text \The term "carrier" might be unknown to the non-math reader. For the finite field with p elements, where p is prime, the carrier is integers mod p, or simply the interval [0, p-1].\ lemma inCarrier: "(a \ carrier R) = (0 \ a \ a < p)" by (simp add: res_carrier_eq) text \The carrier R is a ring of integers. But also, it's only the integers in [0,p-1]. Isabelle/HOL is a typed language, so we might need to convert an integer (int) in the carrier to a natural number (nat). This is not a problem, since all elements of the carrier are non-negative. So we can convert to a nat and back to an int and we are back to where we started.\ lemma inCarrierNatInt: assumes "a \ carrier R" shows "int (nat a) = a" using assms by (simp add: inCarrier) text \The term nonsingular is defined in Elliptic_Locale.thy and is done for a generic type of curve. Then Elliptic_Test.thy builds on Elliptic_Locale.thy builds on that theory but specifically for curves over the integers. So just to make it more clear for the case of integers, the definition of nonsingular here is exactly the familiar notion for elliptic curves as found in SEC 1, page 16.\ lemma nonsingularEq: assumes "nonsingular a b" shows "(4*a^3 + 27*b^2) mod p \ 0" by (metis nonsingular_def add_cong assms mult_cong res_of_integer_eq res_pow_eq zero_cong) lemma nonsingularEq_nat: fixes a b :: nat assumes "nonsingular (int a) (int b)" shows "(4*a^3 + 27*b^2) mod p \ 0" proof - have "(4*(int a)^3 + 27*(int b)^2) mod p \ 0" using assms nonsingularEq by meson then show ?thesis by (metis (mono_tags, opaque_lifting) cring_class.of_nat_add cring_class.of_nat_mult of_nat_mod of_nat_numeral of_nat_power semiring_1_class.of_nat_0) qed text \As we did for "nonsingular", we see that the definition "on_curve" is exactly the usual definition for curves over the integers, for example on page 6 of SEC 1.\ lemma onCurveEq: assumes "on_curve a b P" "P = Point x y" shows "y^2 mod p = (x^3 + a*x + b) mod p" proof - have "y [^] (2::nat) = x [^] (3::nat) \ a \ x \ b" using assms(1,2) on_curve_def by simp then show ?thesis by (metis add.commute mod_add_right_eq res_add_eq res_mult_eq res_pow_eq) qed lemma onCurveEq2: assumes "on_curve a b (Point x y)" shows "(x \ carrier R) \ (y \ carrier R) \ (y^2 mod p = (x^3 + a*x + b) mod p)" proof - have 1: "y^2 mod p = (x^3 + a*x + b) mod p" using assms onCurveEq by blast have 2: "x \ carrier R \ y \ carrier R" using assms on_curve_def by auto show ?thesis using 1 2 on_curve_def by fast qed lemma onCurveEq3: assumes "(x \ carrier R) \ (y \ carrier R) \ (y^2 mod p = (x^3 + a*x + b) mod p)" shows "on_curve a b (Point x y)" by (smt (z3) assms on_curve_def mod_add_right_eq point.case(2) res_add_eq res_mult_eq res_pow_eq) lemma onCurveEq4: "on_curve a b (Point x y) = ((x \ carrier R) \ (y \ carrier R) \ (y^2 mod p = (x^3 + a*x + b) mod p))" using onCurveEq2 onCurveEq3 by blast text \The number of points on an elliptic curve over a finite field is finite.\ lemma CurveFinite: "finite {P. on_curve a b P}" proof - let ?S = "{P. on_curve a b P}" have 1: "Infinity \ ?S" using on_curve_def by force let ?S' = "{P. P \ Infinity \ on_curve a b P}" have 2: "?S = {Infinity} \ ?S'" using 1 by fast have 3: "?S' = {P. \x y. (P = Point x y) \ (x \ carrier R) \ (y \ carrier R) \ (y^2 mod p = (x^3 + a*x + b) mod p)}" by (metis (no_types, lifting) onCurveEq2 onCurveEq3 point.distinct(2) point.exhaust) have 4: "?S' \ {P. \x y. (P = Point x y) \ (x \ carrier R) \ (y \ carrier R)}" using 3 by auto have 5: "finite {P. \x y. (P = Point x y) \ (x \ carrier R) \ (y \ carrier R)}" by (simp add: finite_image_set2) have 6: "finite ?S'" using 4 5 finite_subset by blast show ?thesis using 2 6 by force qed text \The points on a non-singular elliptic curve form a group. So if the number of points is a prime n, every point on the curve has order n.\ lemma CurveOrderPrime: assumes "a \ carrier R" "b \ carrier R" "nonsingular a b" "C = {P. on_curve a b P}" "prime n" "card C = n" "Q \ C" shows "point_mult a n Q = Infinity" proof - have 1: "ell_prime_field p (nat a) (nat b)" using assms(1,2,3) ell_prime_field_def R_def ell_prime_field_axioms_def inCarrierNatInt nonsingularEq_nat residues_prime_gt2_axioms by algebra interpret EPF: ell_prime_field p R "(nat a)" "(nat b)" using 1 apply blast using R_def by presburger have 2: "C = carrier EPF.curve" using assms(1,2,4) EPF.curve_def inCarrierNatInt by simp have 3: "Q [^]\<^bsub>EPF.curve\<^esub>n = \\<^bsub>EPF.curve\<^esub>" using 2 assms(5,6,7) EPF.curve.power_order_eq_one EPF.finite_curve by presburger show ?thesis by (metis 2 3 EPF.in_carrier_curve EPF.multEqPowInCurve EPF.one_curve assms(1,7) inCarrierNatInt) qed text \"opp" is defined in Elliptic_Locale. This is the "opposite" of a point on the curve, meaning you just negate the y coordinate (modulo p).\ lemma oppEq: "opp (Point x y) = Point x ((-y) mod p)" using opp_Point res_neg_eq by presburger lemma oppEq': fixes x y :: int assumes "0 < y" "y < p" shows "opp (Point x y) = Point x (p-y)" proof - let ?y = "(-y) mod p" have "?y = p - y" using assms(1,2) zmod_zminus1_eq_if by auto then show ?thesis using oppEq by presburger qed subsection \2.3 Data Types and Conversions\ text \We have translated routines for converting natural numbers from/to words of a given number of bits for many crypto standards. This can be found in Words.thy. We have abbreviations for converting natural numbers to/from octet string, and to/from bit strings. For ease of reference, we will discuss in comments the existing conversion methods for each mentioned in section 2.3 of SEC 1. New in this standard is converting between elliptic curve points and octet strings. We define those primitives below.\ subsubsection \2.3.1 Bit-String-to-Octet-String Conversion\ text\"Bit strings should be converted to octet strings as described in this section. Informally the idea is to pad the bit string with 0's on the left to make its length a multiple of 8, then chop the result up into octets." This conversion in Words.thy is named bits_to_octets_len. Note that there is another routine called bits_to_octets. This second routine will convert to octets but will "wash away" any "extra" 0 bits on the left so that the leftmost byte of the output is not zero. Whereas, the output of bits_to_octets_len will always be \l/8\, where l is the length of the input bit string, given that the input bit string is valid. A bit string is valid when each value in the list is either 0 or 1. Note bit strings are represented as lists, as are octet strings, so bits_to_octets_len [0,0,0,0,0,0,0,0,0,0,0,1,0,0,1] = [0,9] bits_to_octets [0,0,0,0,0,0,0,0,0,0,0,1,0,0,1] = [9] Bottom line: when this standard calls for conversion from a bit string to an octet string, we use bits_to_octets_len. \ subsubsection \2.3.2 Octet-String-to-Bit-String Conversion\ text \Similar to above, to convert from a valid octet string to a bit string, we use octets_to_bits_len in Words.thy. Again, there is a second routine called octets_to_bits that will "wash away" any high 0s. So octets_to_bits_len [0,9] = [0,0,0,0,0,0,0,0,0,0,0,0,1,0,0,1] octets_to_bits [0,9] = [1,0,0,1] The output of octets_to_bits_len will always be 8x as long as the input, when the input is a valid string of octets. Octets are valid when each octet is in the list is a value in [0, 255]. The output of octets_to_bits will always start with 1. Bottom line: when the standard calls for conversion from an octet string to a bit string, we use octets_to_bits_len. \ subsubsection \2.3.3 Elliptic-Curve-Point-to-Octet-String Conversion\ text \We have not dealt with storing elliptic curve points as octet strings in other standards. Thus this is a new data conversion primitive. Note that the standard writes ceiling(log2 p / 8). This is the same as the octet_length when p is an odd prime. (See Words/octetLenLog2Prime.) So we write octet_length instead. Also the standard says: Assign the value 0x02 to the single octet Y if (y mod 2) = 0, or the value 0x03 if (y mod 2) = 1. Here we write that as Y = [2 + (y mod 2)]. Remember that octets are modeled in Isabelle as a list of natural numbers. So Y will be [2] if y is even and will be [3] when y is odd. We hope that this is clear to the reader that it exactly matches the standard as written.\ definition point_to_octets_nocomp :: "nat \ nat \ octets" where "point_to_octets_nocomp x y = ( let l = octet_length p; X = nat_to_octets_len x l; Y = nat_to_octets_len y l in 4 # X @ Y )" definition point_to_octets_comp :: "nat \ nat \ octets" where "point_to_octets_comp x y = (let l = octet_length p; X = nat_to_octets_len x l; Y = [2 + (y mod 2)] in Y @ X )" definition point_to_octets :: "int point \ bool \ octets" where "point_to_octets P Compress = ( case P of Infinity \ [0] | Point x y \ ( if Compress then (point_to_octets_comp (nat x) (nat y)) else (point_to_octets_nocomp (nat x) (nat y)) ) )" text \Now a few lemmas about converting a point to octets. The first set of lemmas shows that we know the length of the output when the input is a point on the curve. For points not on the curve, we have a bound on the length of the resulting octets.\ lemma point_to_octets_len_Inf: "length (point_to_octets Infinity C) = 1" using point_to_octets_def by simp lemma point_to_octets_nocomp_len_bnd: assumes "l = octet_length p" shows "2*l + 1 \ length (point_to_octets_nocomp x y)" proof - let ?X = "nat_to_octets_len x l" have 1: "l \ length ?X" using assms(1) nat_to_words_len_len by auto let ?Y = "nat_to_octets_len y l" have 2: "l \ length ?Y" using assms(1) nat_to_words_len_len by auto have 3: "point_to_octets_nocomp x y = 4 # ?X @ ?Y" using point_to_octets_nocomp_def assms(1) by meson then show ?thesis using 1 2 by simp qed lemma point_to_octets_nocomp_len: assumes "x < p" "y < p" "l = octet_length p" shows "length (point_to_octets_nocomp x y) = 2*l + 1" proof - let ?X = "nat_to_octets_len x l" have 1: "length ?X = l" using assms(1,3) nat_to_word_len_mono' zero_less_numeral by blast let ?Y = "nat_to_octets_len y l" have 2: "length ?Y = l" using assms(2,3) nat_to_word_len_mono' zero_less_numeral by blast have 3: "point_to_octets_nocomp x y = 4 # ?X @ ?Y" using point_to_octets_nocomp_def assms(3) by meson show ?thesis using 1 2 3 by auto qed lemma point_to_octets_len_F_bnd: assumes "P = Point x y" "l = octet_length p" shows "2*l + 1 \ length (point_to_octets P False)" using assms point_to_octets_nocomp_len_bnd point_to_octets_def by auto lemma point_to_octets_len_F': assumes "P = Point (x::int) (y::int)" "x < p" "y < p" "l = octet_length p" shows "length (point_to_octets P False) = 2*l + 1" proof - have x2: "(nat x) < p" using assms(2) m_gt_one by linarith have y2: "(nat y) < p" using assms(3) m_gt_one by linarith have "length (point_to_octets P False) = length (point_to_octets_nocomp (nat x) (nat y))" using assms(1) point_to_octets_def by simp then show ?thesis using assms(4) point_to_octets_nocomp_len x2 y2 by presburger qed lemma point_to_octets_len_F: assumes "on_curve a b P" "l = octet_length p" "P \ Infinity" shows "length (point_to_octets P False) = 2*l + 1" proof - obtain x and y where 1: "P = Point x y" using assms(3) by (meson point.exhaust) have x1: "x < p" using 1 assms(1) on_curve_def inCarrier by auto have y1: "y < p" using 1 assms(1) on_curve_def inCarrier by auto show ?thesis using 1 x1 y1 assms(2) point_to_octets_len_F' by blast qed lemma point_to_octets_comp_len_bnd: assumes "l = octet_length p" shows "l + 1 \ length (point_to_octets_comp x y)" proof - let ?X = "nat_to_octets_len x l" have 1: "l \ length ?X" using assms(1) nat_to_words_len_len by auto let ?Y = "[2 + (y mod 2)]" have 2: "length ?Y = 1" by simp have 3: "point_to_octets_comp x y = ?Y @ ?X" using point_to_octets_comp_def assms(1) by meson then show ?thesis using 1 2 by simp qed lemma point_to_octets_comp_len: assumes "x < p" "l = octet_length p" shows "length (point_to_octets_comp x y) = l + 1" proof - let ?X = "nat_to_octets_len x l" have 1: "length ?X = l" using assms(1,2) nat_to_word_len_mono' zero_less_numeral by blast then show ?thesis by (metis 1 assms(2) point_to_octets_comp_def One_nat_def length_append list.size(3,4) Suc_eq_plus1 plus_1_eq_Suc) qed lemma point_to_octets_len_T_bnd: assumes "P = Point x y" "l = octet_length p" shows "l + 1 \ length (point_to_octets P True)" using assms point_to_octets_comp_len_bnd point_to_octets_def by auto lemma point_to_octets_len_T': assumes "P = Point (x::int) (y::int)" "x < p" "l = octet_length p" shows "length (point_to_octets P True) = l + 1" proof - have 1: "(nat x) < p" using assms(2) m_gt_one by linarith have "length (point_to_octets P True) = length (point_to_octets_comp (nat x) (nat y))" using assms(1) point_to_octets_def by simp then show ?thesis using assms(3) point_to_octets_comp_len 1 by presburger qed lemma point_to_octets_len_T: assumes "on_curve a b P" "l = octet_length p" "P \ Infinity" shows "length (point_to_octets P True) = l + 1" proof - obtain x and y where 1: "P = Point x y" using assms(3) by (meson point.exhaust) have x1: "x < p" using 1 assms(1) on_curve_def inCarrier by auto show ?thesis using 1 x1 assms(2) point_to_octets_len_T' by blast qed lemma point_to_octets_len1: assumes "length (point_to_octets P C) = 1" shows "P = Infinity" apply (cases P) apply simp by (metis (full_types) add_diff_cancel_right' assms diff_is_0_eq' mult_is_0 neq0_conv octet_len_p point_to_octets_len_F_bnd point_to_octets_len_T_bnd zero_neq_numeral) lemma point_to_octets_len_oncurve: assumes "on_curve a b P" "l = octet_length p" "L = length (point_to_octets P C)" shows "L = 1 \ L = l+1 \ L = 2*l+1" apply (cases P) using point_to_octets_len_Inf assms(3) apply presburger apply (cases C) using point_to_octets_len_T assms(1,2,3) apply simp using point_to_octets_len_F assms(1,2,3) by simp text \Next set of lemmas: We know the output of point_to_octets is always a valid string of octets, no matter what is input. Recall that an octet is valid if its value is < 256 = 2^8.\ lemma point_to_octets_comp_valid: "octets_valid (point_to_octets_comp x y)" proof - have "2 + (y mod 2) < 256" by linarith then show ?thesis using point_to_octets_comp_def nat_to_words_len_valid words_valid_cons words_valid_concat by force qed lemma point_to_octets_nocomp_valid: "octets_valid (point_to_octets_nocomp x y)" proof - have "(4::nat) < 256" by fastforce then show ?thesis by (metis point_to_octets_nocomp_def nat_to_words_len_valid words_valid_cons words_valid_concat Twoto8) qed lemma point_to_octets_valid: "octets_valid (point_to_octets P C)" apply (cases P) apply (simp add: point_to_octets_def words_valid_cons words_valid_nil) apply (cases C) apply (simp add: point_to_octets_comp_valid point_to_octets_def) by (simp add: point_to_octets_nocomp_valid point_to_octets_def) text \Next up: converting a point to octets without compression is injective. A similar lemma when compression is used is a bit more complicated and is shown below.\ lemma point_to_octets_nocomp_inj: assumes "point_to_octets_nocomp x y = point_to_octets_nocomp x' y'" "x < p" "x' < p" shows "x = x' \ y = y'" proof - let ?l = "octet_length p" let ?M = "point_to_octets_nocomp x y" let ?M' = "point_to_octets_nocomp x' y'" let ?X = "nat_to_octets_len x ?l" let ?X' = "nat_to_octets_len x' ?l" let ?Y = "nat_to_octets_len y ?l" let ?Y' = "nat_to_octets_len y' ?l" have l: "length ?X = ?l \ length ?X' = ?l" by (meson assms(2,3) le_trans less_or_eq_imp_le nat_lowbnd_word_len2 nat_to_words_len_upbnd not_le zero_less_numeral) have m1: "?M = 4 # ?X @ ?Y" using point_to_octets_nocomp_def by meson have m2: "?M' = 4 # ?X' @ ?Y'" using point_to_octets_nocomp_def by meson have x1: "?X = take ?l (drop 1 ?M)" using m1 l by simp have x2: "?X' = take ?l (drop 1 ?M')" using m2 l by simp have x3: "?X = ?X'" using x1 x2 assms(1) by argo have x4: "x = x'" using nat_to_words_len_inj x3 by auto have y1: "?Y = drop (1+?l) ?M" using m1 l by simp have y2: "?Y' = drop (1+?l) ?M'" using m2 l by simp have y3: "?Y = ?Y'" using y1 y2 assms(1) by argo have y4: "y = y'" using nat_to_words_len_inj y3 by auto show ?thesis using x4 y4 by fast qed lemma point_to_octets_nocomp_inj': assumes "point_to_octets_nocomp x y = point_to_octets_nocomp x' y'" "x < p" "y < p" shows "x = x' \ y = y'" proof - let ?l = "octet_length p" let ?M = "point_to_octets_nocomp x y" let ?M' = "point_to_octets_nocomp x' y'" let ?X = "nat_to_octets_len x ?l" let ?X' = "nat_to_octets_len x' ?l" let ?Y = "nat_to_octets_len y ?l" let ?Y' = "nat_to_octets_len y' ?l" have l1: "length ?X = ?l \ length ?Y = ?l" by (meson assms(2,3) le_trans less_or_eq_imp_le nat_lowbnd_word_len2 nat_to_words_len_upbnd not_le zero_less_numeral) have l2: "?l \ length ?X' \ ?l \ length ?Y'" using nat_to_words_len_len by auto have m1: "?M = 4 # ?X @ ?Y" using point_to_octets_nocomp_def by meson have m2: "?M' = 4 # ?X' @ ?Y'" using point_to_octets_nocomp_def by meson have l3: "length ?M = 1 + length ?X + length ?Y" using m1 by force have l4: "length ?M' = 1 + length ?X' + length ?Y'" using m2 by force have l5: "length ?X + length ?Y = length ?X' + length ?Y'" using assms(1) l3 l4 by simp have l6: "length ?X' = ?l \ length ?Y' = ?l" using l5 l2 l1 by simp have x1: "?X = take ?l (drop 1 ?M)" using m1 l1 by simp have x2: "?X' = take ?l (drop 1 ?M')" using m2 l6 by simp have x3: "?X = ?X'" using x1 x2 assms(1) by argo have x4: "x = x'" using nat_to_words_len_inj x3 by auto have y1: "?Y = drop (1+?l) ?M" using m1 l1 by simp have y2: "?Y' = drop (1+?l) ?M'" using m2 l6 by simp have y3: "?Y = ?Y'" using y1 y2 assms(1) by argo have y4: "y = y'" using nat_to_words_len_inj y3 by auto show ?thesis using x4 y4 by fast qed text \Now a few lemmas about converting the "opposite" of a point to octets using compression.\ definition FlipY :: "nat \ nat" where "FlipY Y = (if Y = 2 then 3 else 2)" definition FlipYhd :: "octets \ octets" where "FlipYhd M = (FlipY (hd M)) # (tl M)" lemma point_to_octets_comp_opp: assumes "M = point_to_octets_comp x y" "M' = point_to_octets_comp x (p-y)" "y < p" "0 < y" shows "M = FlipYhd M' \ M' = FlipYhd M" proof - let ?l = "octet_length p" let ?X = "nat_to_octets_len x ?l" let ?Y = "2 + (y mod 2)" let ?y = "(p-y)" let ?Y' = "2 + (?y mod 2)" have 1: "M = ?Y # ?X \ M' = ?Y' # ?X" using assms(1,2) point_to_octets_comp_def by force have 2: "(y mod 2 = 0) = (?y mod 2 = 1)" by (metis p_mod2 assms(3) dvd_imp_mod_0 le_add_diff_inverse less_imp_le_nat odd_add odd_iff_mod_2_eq_one) have 3: "(?Y = 2) = (?Y' = 3)" using 2 by auto have 4: "(?Y = FlipY ?Y') \ (?Y' = FlipY ?Y)" using FlipY_def 3 by fastforce show ?thesis using FlipYhd_def 4 1 by force qed lemma point_to_octets_comp_opp': assumes "P = Point x y" "0 < y" "y < p" "M = point_to_octets P True" "M' = point_to_octets (opp P) True" shows "M = FlipYhd M' \ M' = FlipYhd M" by (simp add: assms point_to_octets_comp_opp oppEq' point_to_octets_def nat_minus_as_int) subsubsection \2.3.4 Octet-String-to-Elliptic-Curve-Point Conversion\ text \This conversion function needs to know the elliptic curve coefficients a and b to check if the recovered point is really on the curve. We also check that the high octet is 4. And we insist that the octet string is valid (meaning every octet is a value < 256) and the number of octets must be 2*l+ 1, where l is the length of the modulus p in octets.\ definition octets_to_point_nocomp_validInput :: "int \ int \ octets \ bool" where "octets_to_point_nocomp_validInput a b M = ( let W = hd M; M' = drop 1 M; l = octet_length p; X = take l M'; x = int (octets_to_nat X); Y = drop l M'; y = int (octets_to_nat Y); P = Point x y in (octets_valid M) \ (W = 4) \ (on_curve a b P) \ (length M = 2*l + 1) )" definition octets_to_point_nocomp :: "int \ int \ octets \ int point option" where "octets_to_point_nocomp a b M = ( let W = hd M; M' = drop 1 M; l = octet_length p; X = take l M'; x = int (octets_to_nat X); Y = drop l M'; y = int (octets_to_nat Y); P = Point x y in if (octets_to_point_nocomp_validInput a b M) then Some P else None )" text \Recovering a point when compression was used is not as straight forward. We need to find a square root mod p, which, depending on p, is not always simple to do. If p mod 4 = 3, then it is easy to write down the square root of a quadratic residue, but the SEC 1 standard does not require that p meet that restriction. In this definition, we check if \ is a quadratic residue and let HOL return one of its square roots, \. There is an issue that is not discussed in the standard. In the case when \ = 0, there is only one square root < p, namely 0. If \ = 0 and y = 1, the standard says to set y' = p. This will solve the elliptic curve equation but p is not a field element modulo p. To handle this, we add a check that is not described in the standard. If \ = 0, we insist that y must be 0 (correspondingly Y must be 2). Note that we define a variable T below that does not appear in the standard explicitly. T is for "test" and it just tests that the high octet is either 2 or 3. (See step 2.3 in section 2.3.4 in the standard.)\ definition octets_to_point_comp_validInput :: "int \ int \ octets \ bool" where "octets_to_point_comp_validInput a b M = ( let Y = hd M; T = (Y = 2) \ (Y = 3); l = octet_length p; X = drop 1 M; x = int (octets_to_nat X); \ = (x^3 + a*x + b) mod p in (octets_valid M) \ (x < p) \ T \ (\ = 0 \ Y = 2) \ (QuadRes' p \) \ (length M = l + 1) )" definition octets_to_point_comp :: "int \ int \ octets \ int point option" where "octets_to_point_comp a b M = ( let Y = hd M; y = Y - 2; X = drop 1 M; x = int (octets_to_nat X); \ = (x^3 + a*x + b) mod p; \ = some_root_nat p \; y' = int (if \ mod 2 = y then \ else (p-\)); P = Point x y' in if (octets_to_point_comp_validInput a b M) then Some P else None )" definition octets_to_point_validInput :: "int \ int \ octets \ bool" where "octets_to_point_validInput a b M \ (M = [0]) \ (octets_to_point_comp_validInput a b M) \ (octets_to_point_nocomp_validInput a b M)" definition octets_to_point :: "int \ int \ octets \ int point option" where "octets_to_point a b M = ( let lp = octet_length p; lM = length M in if M = [0] then Some Infinity else ( if lM = lp + 1 then octets_to_point_comp a b M else ( if lM = 2*lp + 1 then octets_to_point_nocomp a b M else None ) ) )" text \octets_to_point produces the point at infinity in only one way.\ lemma octets2PointNoCompNotInf: assumes "octets_to_point_nocomp a b M = Some P" shows "P \ Infinity" by (smt (z3) assms octets_to_point_nocomp_def option.distinct(1) option.inject point.distinct(1)) lemma octets2PointCompNotInf: assumes "octets_to_point_comp a b M = Some P" shows "P \ Infinity" by (smt (z3) assms octets_to_point_comp_def option.distinct(1) option.inject point.distinct(1)) lemma octets2PointInf1: "octets_to_point a b [0] = Some Infinity" using octets_to_point_def by presburger lemma octets2PointInf2: assumes "octets_to_point a b M = Some Infinity" shows "M = [0]" by (smt (z3) assms octets_to_point_def octets2PointNoCompNotInf octets2PointCompNotInf option.distinct(1)) lemma octets2PointInf: "(octets_to_point a b M = Some Infinity) = (M = [0])" using octets2PointInf1 octets2PointInf2 by blast text \octets_to_point produces a point on the curve or None.\ lemma octets2PointNoCompOnCurve: assumes "octets_to_point_nocomp a b M = Some P" shows "on_curve a b P" by (smt (z3) assms octets_to_point_nocomp_def octets_to_point_nocomp_validInput_def option.distinct(1) option.inject) lemma octets2PointCompOnCurve: assumes "octets_to_point_comp a b M = Some P" shows "on_curve a b P" proof - let ?Y = "hd M" let ?T = "(?Y = 2) \ (?Y = 3)" let ?y = "?Y - 2" let ?X = "drop 1 M" let ?x = "int (octets_to_nat ?X)" let ?\ = "(?x^3 + a*?x + b) mod p" let ?\ = "some_root_nat p ?\" let ?y' = "int (if ?\ mod 2 = ?y then ?\ else (p-?\))" have 1: "P = Point ?x ?y'" by (smt (z3) assms(1) octets_to_point_comp_def option.distinct(1) option.inject) have 2: "octets_to_point_comp_validInput a b M" by (smt (verit, best) assms(1) octets_to_point_comp_def option.simps(3)) have 3: "QuadRes' p ?\" using 2 octets_to_point_comp_validInput_def by meson have a1: "?\ = 0 \ ?y = 0" by (metis 2 octets_to_point_comp_validInput_def diff_self_eq_0) have a2: "?\ \ 0 \ 0 < ?\" - by (metis (mono_tags, opaque_lifting) 3 QuadRes'HasNatRoot bits_mod_0 mod_mod_trivial - bot_nat_0.not_eq_extremum cong_def int_ops(1) power_zero_numeral) + using QuadRes'HasNatRoot [OF 3] + by (auto simp flip: res_eq_to_cong) (smt (verit) gr0I int_ops(1) mod_less res_eq_to_cong zero_power2 zmod_int) have x1: "0 \ ?x" using of_nat_0_le_iff by blast have x2: "?x < p" using 2 octets_to_point_comp_validInput_def by meson have x3: "?x \ carrier R" using x1 x2 inCarrier by blast have y1: "0 \ ?y'" using of_nat_0_le_iff by blast have y2: "?y' < p" by (smt (verit, ccfv_threshold) 3 a1 a2 QuadRes'HasNatRoot dvd_imp_mod_0 even_diff_nat gr0I int_ops(6) of_nat_0_less_iff zero_less_diff) have y3: "?y' \ carrier R" using y1 y2 inCarrier by blast have y4: "?y'^2 mod p = (?x^3 + a*?x + b) mod p" by (smt (z3) 3 QuadRes'HasNatRoot QuadRes'HasTwoNatRoots cong_def mod_mod_trivial of_nat_power) show ?thesis using 1 y4 x3 y3 onCurveEq3 by presburger qed lemma octets2PointOnCurve: assumes "octets_to_point a b M = Some P" shows "on_curve a b P" by (smt (verit) assms octets2PointCompOnCurve octets2PointNoCompOnCurve octets_to_point_def on_curve_infinity option.inject option.simps(3)) text \Now we have some lemmas about converting a point to octets then back to a point. First when compression is not used.\ lemma assumes "on_curve a b P" "P = Point x y" "M = point_to_octets_nocomp (nat x) (nat y)" shows point2OctetsNoComp_validInput: "octets_to_point_nocomp_validInput a b M" and point2Octets2PointNoComp: "octets_to_point_nocomp a b M = Some P" proof - let ?l = "octet_length p" let ?X = "nat_to_octets_len (nat x) ?l" let ?Y = "nat_to_octets_len (nat y) ?l" have x1: "nat x < p" using assms(1,2) on_curve_def inCarrier by fastforce have x2: "length ?X = ?l" using nat_to_word_len_mono' x1 zero_less_numeral by blast have y1: "nat y < p" using assms(1,2) on_curve_def inCarrier by fastforce have y2: "length ?Y = ?l" using nat_to_word_len_mono' y1 zero_less_numeral by blast have m1: "M = 4 # ?X @ ?Y" using assms(3) point_to_octets_nocomp_def by meson have m2: "octets_valid M" using assms(3) point_to_octets_nocomp_valid by blast let ?M' = "drop 1 M" let ?X' = "take ?l ?M'" let ?Y' = "drop ?l ?M'" have x3: "?X = ?X'" using m1 x2 by fastforce have y3: "?Y = ?Y'" using m1 x2 by fastforce let ?x = "int (octets_to_nat ?X')" let ?y = "int (octets_to_nat ?Y')" have x4: "?x = int (nat x)" by (metis x3 nat_to_words_len_to_nat zero_less_numeral) have x5: "?x = x" using assms(1,2) on_curve_def inCarrierNatInt x4 by auto have y4: "?y = int (nat y)" by (metis y3 nat_to_words_len_to_nat zero_less_numeral) have y5: "?y = y" using assms(1,2) on_curve_def inCarrierNatInt y4 by auto show "octets_to_point_nocomp_validInput a b M" using m2 m1 octets_to_point_nocomp_validInput_def assms(1,2) x2 x5 y2 y5 by auto then show "octets_to_point_nocomp a b M = Some P" using assms(2) octets_to_point_nocomp_def x5 y5 by presburger qed text \Continuing lemmas about converting a point to octets then back to a point: Now for when compression is used.\ lemma assumes "on_curve a b P" "P = Point x y" "M = point_to_octets_comp (nat x) (nat y)" shows point2OctetsValidInputComp: "octets_to_point_comp_validInput a b M" and point2Octets2PointComp: "octets_to_point_comp a b M = Some P" proof - let ?l = "octet_length p" let ?X = "nat_to_octets_len (nat x) ?l" let ?Y = "[2 + ((nat y) mod 2)]" have m1: "M = ?Y @ ?X" using assms(3) point_to_octets_comp_def by presburger let ?Y' = "hd M" have y1: "?Y' = 2 + (nat y) mod 2" using m1 by fastforce let ?T = "?Y' = 2 \ ?Y' = 3" have T1: "?T" using y1 by force let ?X' = "drop 1 M" have x1: "?X' = ?X" using m1 by auto let ?x = "int (octets_to_nat ?X')" have x2: "?x = int (nat x)" by (metis x1 nat_to_words_len_to_nat zero_less_numeral) have x3: "?x = x" using assms(1,2) on_curve_def inCarrierNatInt x2 by auto have x4: "?x < p" using x3 assms(1,2) inCarrier on_curve_def by auto let ?\ = "((?x^3 + a*?x + b) mod p)" have y2: "y^2 mod p = ?\" using assms(1,2) onCurveEq x3 by presburger have y3: "0 \ y \ y < p" using assms(1,2) inCarrier onCurveEq2 by blast have a1: "QuadRes' p ?\" by (metis y2 QuadRes'_def QuadResImplQuadRes'2 QuadRes_def cong_def mod_mod_trivial p_prime) have y4: "?\ = 0 \ y^2 mod p = 0" using y2 by argo have y5: "?\ = 0 \ y = 0" by (meson y4 y3 ZeroSqrtMod p_prime prime_nat_int_transfer) have y6: "?\ = 0 \ ?Y' = 2" using y1 y5 by fastforce have l1: "length ?X = ?l" using nat_to_word_len_mono' x3 x4 by force have l2: "length M = ?l + 1" using l1 m1 by auto show A1: "octets_to_point_comp_validInput a b M" by (metis assms(3) x4 T1 y6 a1 l2 octets_to_point_comp_validInput_def point_to_octets_comp_valid) let ?y = "?Y' - 2" have y7: "?y = (nat y) mod 2" using y1 by force let ?\ = "some_root_nat p ?\" have b1: "?\ < p" using QuadRes'HasNatRoot a1 by blast let ?y' = "int (if ?\ mod 2 = ?y then ?\ else (p-?\))" let ?P = "Point ?x ?y'" have A2: "octets_to_point_comp a b M = Some ?P" by (smt (verit) A1 octets_to_point_comp_def) show "octets_to_point_comp a b M = Some P" proof (cases "?\ = 0") case T0: True have T1: "y = 0" by (metis ZeroSqrtMod y4 y3 p_prime T0 prime_nat_int_transfer) show ?thesis using A2 T0 T1 ZeroSqrtMod' assms(2) p_prime x3 y7 by force next case False then have "0 < ?\" using m_gt_one by (auto intro: le_neq_trans) then have "(nat y) = ?\ \ (nat y) = p - ?\" by (metis Euclidean_Rings.pos_mod_bound QuadResHasExactlyTwoRoots2 a1 gt2 int_ops(1) nat_int not_less of_nat_le_0_iff p_prime y2 y3) then show ?thesis by (smt (verit) A2 assms(2) b1 diff_is_0_eq' even_diff_nat int_nat_eq less_imp_le_nat less_numeral_extra(3) odd_add odd_iff_mod_2_eq_one p_mod2 x3 y7 y3 zero_less_diff) qed qed text \Now it's easy to show that for points on the curve, point_to_octets_comp is injective.\ lemma point_to_octets_comp_inj: assumes "point_to_octets_comp x y = point_to_octets_comp x' y'" "on_curve a b (Point (int x) (int y))" "on_curve a b (Point (int x') (int y'))" shows "x = x' \ y = y'" by (metis assms(1,2,3) nat_int option.inject point.inject point2Octets2PointComp) lemma point_to_octets_comp_inj': assumes "point_to_octets_comp (nat x) (nat y) = point_to_octets_comp (nat x') (nat y')" "on_curve a b (Point x y)" "on_curve a b (Point x' y')" shows "x = x' \ y = y'" by (metis assms(1,2,3) option.inject point.inject point2Octets2PointComp) text \In the Public Key Recovery Operation, you have "forgotten" y. So you just try to recover the whole point by guessing that y is even. So we have a few lemmas here that will be useful for that recovery operation.\ lemma point2Octets2PointComp_forgot_Y2: assumes "on_curve a b P" "P = Point x y" "y mod 2 = 0" "l = octet_length p" "X = nat_to_octets_len (nat x) l" "M = 2 # X" shows "octets_to_point_comp a b M = Some P" proof - have "M = point_to_octets_comp (nat x) (nat y)" by (metis assms(3,4,5,6) point_to_octets_comp_def add.right_neutral append_Cons append_Nil dvd_imp_mod_0 even_mod_2_iff even_of_nat int_nat_eq odd_add) then show ?thesis using assms(1,2) point2Octets2PointComp by blast qed lemma point2Octets2PointComp_forgot_Y2': assumes "on_curve a b P" "P = Point x y" "y mod 2 = 1" "l = octet_length p" "X = nat_to_octets_len (nat x) l" "M = 2 # X" shows "octets_to_point_comp a b M = Some (opp P)" proof - have 0: "y \ 0" using assms(3) by force have 1: "0 < y" using 0 assms(1,2) on_curve_def inCarrier by force have 2: "y < p" using assms(1,2) on_curve_def inCarrier by force let ?y = "p - y" let ?P = "Point x ?y" have 3: "?P = opp P" using oppEq'[OF 1 2] assms(2) by presburger have 4: "(nat ?y) mod 2 = 0" by (metis assms(3) cancel_comm_monoid_add_class.diff_cancel even_of_nat int_nat_eq mod_0 mod_diff_right_eq not_mod_2_eq_1_eq_0 odd_iff_mod_2_eq_one p_mod2) have 5: "M = point_to_octets_comp (nat x) (nat ?y)" using 4 assms(4,5,6) point_to_octets_comp_def by fastforce show ?thesis using 3 5 assms(1,2) point2Octets2PointComp opp_closed by auto qed lemma point2Octets2PointComp_PoppP: assumes "on_curve a b P" "P = Point x y" "l = octet_length p" "X = nat_to_octets_len (nat x) l" "M = 2 # X" shows "octets_to_point_comp a b M = Some P \ octets_to_point_comp a b M = Some (opp P)" apply (cases "y mod 2 = 0") using assms(1,2,3,4,5) point2Octets2PointComp_forgot_Y2 apply presburger using assms(1,2,3,4,5) point2Octets2PointComp_forgot_Y2' by presburger lemma point2Octets2PointComp_PoppP': assumes "on_curve a b P" "P = Point x y" "l = octet_length p" "X = nat_to_octets_len (nat x) l" "M = 2 # X" shows "octets_to_point a b M = Some P \ octets_to_point a b M = Some (opp P)" proof - have "x < p" using assms(1,2) inCarrier onCurveEq2 by blast then have "nat x < p" using m_gt_one by linarith then have "length X = l" using assms(3,4) nat_to_word_len_mono' zero_less_numeral by blast then have "octets_to_point a b M = octets_to_point_comp a b M" by (simp add: octets_to_point_def assms(3,5)) then show ?thesis using assms point2Octets2PointComp_PoppP by algebra qed text \Putting together what we have proved above, if P is on the curve, then converting P to octets and then back to a point, whether or not compression is used, will get back the original P.\ lemma point2Octets2Point: assumes "on_curve a b P" "M = point_to_octets P C" shows "octets_to_point a b M = Some P" proof (cases P) case Infinity then show ?thesis by (simp add: assms(2) octets_to_point_def point_to_octets_def) next case P: (Point x y) then show ?thesis proof (cases C) case T0: True let ?lp = "octet_length p" have T1: "length M = ?lp + 1" using P assms(1,2) T0 point_to_octets_len_T by fastforce have T2: "octets_to_point a b M = octets_to_point_comp a b M" by (smt (verit, best) octet_len_p' T1 octets_to_point_def One_nat_def Suc_eq_plus1 list.size(3,4)) show ?thesis using T2 P T0 assms(1,2) point2Octets2PointComp point_to_octets_def by auto next case F0: False let ?lp = "octet_length p" have l: "0 < ?lp" by (metis gr_implies_not_zero gt2 nat_to_words_nil_iff_zero2 neq0_conv zero_less_numeral) have F1: "length M = 2*?lp + 1" using P assms(1,2) F0 point_to_octets_len_F by fastforce have F2: "octets_to_point a b M = octets_to_point_nocomp a b M" by (smt (verit, best) F1 octet_len_p' octets_to_point_def One_nat_def Suc_eq_plus1 list.size(3,4)) then show ?thesis using F2 P F0 assms(1,2) point2Octets2PointNoComp point_to_octets_def by simp qed qed text \Now we think about the other direction. We start with octets. If the octets are a valid input for when compression is or is not used, then if we convert the octets to a point and then back again, we will get back the original octets.\ lemma octets2PointNoCompValidInIFF1: "(octets_to_point_nocomp a b M \ None) = (octets_to_point_nocomp_validInput a b M)" by (smt (z3) not_None_eq octets_to_point_nocomp_def) lemma octets2PointNoCompalidInIFF2: "(\P. octets_to_point_nocomp a b M = Some P) = (octets_to_point_nocomp_validInput a b M)" using octets2PointNoCompValidInIFF1 by force lemma octets2Point2OctetsNoComp: assumes "octets_to_point_nocomp a b M = Some P" shows "point_to_octets P False = M" proof - let ?W = "hd M" let ?M' = "drop 1 M" let ?l = "octet_length p" let ?X = "take ?l ?M'" let ?x = "int (octets_to_nat ?X)" let ?Y = "drop ?l ?M'" let ?y = "int (octets_to_nat ?Y)" let ?P = "Point ?x ?y" have V: "octets_to_point_nocomp_validInput a b M" using assms octets2PointNoCompValidInIFF1 by blast have P1: "?P = P" by (smt (z3) assms(1) V octets_to_point_nocomp_def option.inject) have M1: "length M = 2*?l + 1" using V octets_to_point_nocomp_validInput_def by meson have M2: "octets_valid M" using V octets_to_point_nocomp_validInput_def by meson have W1: "?W = 4" using V octets_to_point_nocomp_validInput_def by meson have X1: "length ?X = ?l" using M1 by force have X2: "octets_valid ?X" using M2 words_valid_take words_valid_drop by blast have Y1: "length ?Y = ?l" using M1 by force have Y2: "octets_valid ?Y" using M2 words_valid_drop by blast have M3: "M = ?W # ?X @ ?Y" by (metis M1 One_nat_def append_take_drop_id drop_0 drop_Suc_Cons hd_Cons_tl leD le_add2 list.size(3) zero_less_one) have O1: "point_to_octets P False = point_to_octets_nocomp (nat ?x) (nat ?y)" using P1 point_to_octets_def by force have x1: "nat ?x = octets_to_nat ?X" by (meson nat_int) have y1: "nat ?y = octets_to_nat ?Y" by (meson nat_int) let ?X' = "nat_to_octets_len (nat ?x) ?l" have X3: "?X' = ?X" by (metis X1 X2 x1 words_to_nat_to_words_len2) let ?Y' = "nat_to_octets_len (nat ?y) ?l" have Y3: "?Y' = ?Y" by (metis Y1 Y2 y1 words_to_nat_to_words_len2) show ?thesis using M3 O1 X3 Y3 W1 point_to_octets_nocomp_def by algebra qed lemma octets2PointCompValidInIFF1: "(octets_to_point_comp a b M \ None) = (octets_to_point_comp_validInput a b M)" by (smt (z3) not_None_eq octets_to_point_comp_def) lemma octets2PointCompValidInIFF2: "(\P. octets_to_point_comp a b M = Some P) = (octets_to_point_comp_validInput a b M)" using octets2PointCompValidInIFF1 by force lemma octets2Point2OctetsComp: assumes "octets_to_point_comp a b M = Some P" shows "point_to_octets P True = M" proof - let ?Y = "hd M" let ?T = "(?Y = 2) \ (?Y = 3)" let ?y = "?Y - 2" let ?l = "octet_length p" let ?X = "drop 1 M" let ?x = "int (octets_to_nat ?X)" let ?\ = "(?x^3 + a*?x + b) mod p" let ?\ = "some_root_nat p ?\" let ?y' = "int (if ?\ mod 2 = ?y then ?\ else (p-?\))" let ?P = "Point ?x ?y'" have V: "octets_to_point_comp_validInput a b M" using assms octets2PointCompValidInIFF1 by blast have P1: "?P = P" by (smt (z3) assms(1) V octets_to_point_comp_def option.inject) let ?X' = "nat_to_octets_len (nat ?x) ?l" let ?Y' = "[2 + ((nat ?y') mod 2)]" let ?M = "?Y' @ ?X'" have O1: "point_to_octets P True = ?M" using P1 point_to_octets_def point_to_octets_comp_def by force have x1: "?X = ?X'" by (metis V add_diff_cancel_right' length_drop nat_int octets_to_point_comp_validInput_def words_to_nat_to_words_len2 words_valid_drop) have a1: "QuadRes' p ?\" by (meson V octets_to_point_comp_validInput_def) have M1: "length M = ?l + 1" by (meson V octets_to_point_comp_validInput_def) have y1: "?y = 0 \ ?y = 1" by (metis V add_diff_cancel_right' diff_self_eq_0 numeral_2_eq_2 numeral_3_eq_3 octets_to_point_comp_validInput_def plus_1_eq_Suc) have y3: "(nat ?y') mod 2 = ?y" by (smt (verit, ccfv_threshold) QuadRes'HasNatRoot a1 even_add even_diff_nat less_imp_of_nat_less mod2_eq_if nat_int p_mod2 y1) have y4: "[?Y] = ?Y'" by (metis V add.commute add.left_neutral add_diff_cancel_right' numeral_2_eq_2 numeral_3_eq_3 octets_to_point_comp_validInput_def plus_1_eq_Suc y3) have M2: "?M = M" by (metis (mono_tags, opaque_lifting) x1 y4 M1 One_nat_def Suc_eq_plus1 append_Cons append_Nil drop0 drop_Suc list.exhaust_sel list.size(3) nat.simps(3)) show ?thesis using O1 M2 by argo qed lemma octets2PointValidInIFF1: "(octets_to_point a b M \ None) = (octets_to_point_validInput a b M)" by (smt (verit, best) R_def octets2Point2OctetsNoComp octets2PointCompValidInIFF1 octets2PointNoCompOnCurve octets_to_point_comp_validInput_def octets_to_point_def octets_to_point_nocomp_def octets_to_point_validInput_def option.simps(3) point2Octets2Point residues_prime_gt2_axioms) lemma octets2PointValidInIFF2: "(\P. octets_to_point a b M = Some P) = (octets_to_point_validInput a b M)" using octets2PointValidInIFF1 by auto lemma octets2Point2Octets: assumes "octets_to_point a b M = Some P" "lM = length M" "lp = octet_length p" "C = (if lM = lp + 1 then True else False)" shows "point_to_octets P C = M" proof - have 1: "(M = [0]) \ (lM = lp + 1) \ (lM = 2*lp + 1)" by (smt (z3) assms(1,2,3) octets_to_point_def option.simps(3)) have 2: "M = [0] \ point_to_octets P C = M" using assms(1) octets_to_point_def point_to_octets_def by force have 3: "lM = lp + 1 \ point_to_octets P C = M" using 2 assms(1,2,3,4) octets_to_point_def octets2Point2OctetsComp by force have 4: "lM = 2*lp + 1 \ point_to_octets P C = M" using 2 3 assms(1,2,3,4) octets_to_point_def octets2Point2OctetsNoComp by force show ?thesis using 1 2 3 4 by fast qed subsubsection \2.3.5 Field-Element-to-Octet-String Conversion\ text \In this translation of SEC 1, we only consider prime fields. So a "field element" is simply a natural number in the range [0, p-1]. So converting field elements to or from octets is the same as converting between nats and octets. The only detail to keep in mind is that the resulting octet string should have the same number of octets as the prime p. So if x is a field element (natural number in [0,p-1]), then we convert x to an octet string M with: M = nat_to_octets_len x (octet_length p) \ subsubsection \2.3.6 Octet-String-to-Field-Element Conversion\ text \Again, we only consider prime fields in this translation. So a field element is a natural number in the range [0,p-1]. We can convert any octet string to a nat using octets_to_nat, defined in Words.thy. If the resulting nat is p or greater, it should be rejected. Note that you might need to check if the input octets were valid using octets_valid, which checks that each octet is a value no greater than 255.\ subsubsection \2.3.7 Integer-to-Octet-String Conversion\ text \Again, there are two conversion primitives for converting a non-negative integer, a.k.a. a natural number, to an octet string. nat_to_octets will convert a nat to an octet string so that the high octet is non-zero. nat_to_octets_len will produce an octet string that has at least a minimum number of octets. In either case, if you translate back to a natural number, you will get back the original value. For example: nat_to_octets 256 = [1,0] nat_to_octets_len 256 5 = [0,0,0,1,0] nat_to_octets_len 256 1 = [1,0] and octets_to_nat [1,0] = 256 octets_to_nat [0,0,0,1,0] = 256 \ subsubsection \2.3.8 Octet-String-to-Integer Conversion\ text \As described above, the routine to convert an octet string to a natural number is octets_to_nat. Of note, octets_valid ensures that a string of octets is valid, meaning each octet has the value 255 or less. So octets_valid [0,1,255,3] = True and octets_valid [0,1,256,3] = False.\ subsubsection \2.3.9 Field-Element-to-Integer Conversion\ text \Again, we only consider prime fields in this translation of SEC 1. So a field element is simply an integer in [0,p-1]. So there is no conversion required.\ section \3. Cryptographic Components\ subsection \3.1 Elliptic Curve Domain Parameters\ text \3.1.1.2.1 Elliptic Curve Domain Parameters over \F\<^sub>p\ Validation Primitive SEC 1 only allows a security level t of 80, 112, 128, 192, or 256. Then "twice the security level" will actually be 2*t except for t = 80 and 256, where it's merely close to 2*t.\ definition twiceSecurityLevel :: "nat \ nat" where "twiceSecurityLevel t = ( if (t = 80) then 192 else ( if (t = 256) then 521 else 2*t ) )" text \The standard writes \log 2 p\. Note that because p is odd (thus not a power of 2), \log 2 p\ = bit_length p. See Words.bitLenLog2Prime. Also, the standard explains that the number of points on the elliptic curve is n*h. (h is called the "cofactor".) It doesn't make a lot of sense to consider an elliptic curve that has zero points on it. So we must have 0 < h. (We know that 0 < n because the standard insists that n is prime.) So while not explicitly written in the standard, we include the requirement that 0 < h.\ definition ECdomainParametersValid :: "nat \ nat \ int point \ nat \ nat \ nat \ bool" where "ECdomainParametersValid a b G n h t \ ( (t \ {80, 112, 128, 192, 256}) \ (\log 2 p\ = twiceSecurityLevel t) \ (a \ carrier R) \ (b \ carrier R) \ (nonsingular a b) \ (n*h \ p) \ (prime n) \ card {P. on_curve a b P} = n*h \ (on_curve a b G) \ (G \ Infinity) \ (point_mult a n G = Infinity) \ (h \ 2^(t div 8)) \ (h = \((sqrt p + 1)^2) div n\) \ 0 < h \ (\B\{1..99}. (p^B mod n \ 1)) )" text \The security level t is not always explicitly listed in elliptic curve parameters.\ definition ECdomainParametersValid' :: "nat \ nat \ int point \ nat \ nat \ bool" where "ECdomainParametersValid' a b G n h \ (\t. ECdomainParametersValid a b G n h t)" lemma ECparamsValidImplies': "ECdomainParametersValid a b G n h t \ ECdomainParametersValid' a b G n h" using ECdomainParametersValid'_def by blast definition twiceSecurityLevelInv :: "nat \ nat" where "twiceSecurityLevelInv twoT = ( if (twoT = 192) then 80 else ( if (twoT = 521) then 256 else (twoT div 2) ) )" lemma twiceSecurityLevelInv: assumes "t \ {80, 112, 128, 192, 256}" shows "twiceSecurityLevelInv (twiceSecurityLevel t) = t" using assms twiceSecurityLevelInv_def twiceSecurityLevel_def by fastforce text \Even when not explicitly stated, the security level can be recovered from the bit length of the prime modulus p.\ lemma ECparamsValid'Implies: assumes "ECdomainParametersValid' a b G n h" "t = twiceSecurityLevelInv (nat \log 2 p\)" shows "ECdomainParametersValid a b G n h t" proof - obtain t' where t1: "ECdomainParametersValid a b G n h t'" using ECdomainParametersValid'_def assms(1) by presburger have t2: "t' \ {80, 112, 128, 192, 256}" using ECdomainParametersValid_def t1 by presburger have t3: "\log 2 p\ = twiceSecurityLevel t'" using ECdomainParametersValid_def t1 by presburger have t4: "t' = twiceSecurityLevelInv (nat \log 2 p\)" using nat_int t2 t3 twiceSecurityLevelInv by presburger show ?thesis using assms(2) t1 t4 by argo qed lemma ECparamsValid'Implies2: assumes "ECdomainParametersValid' a b G n h" "t = twiceSecurityLevelInv (bit_length p)" shows "ECdomainParametersValid a b G n h t" by (metis ECparamsValid'Implies assms(1,2) bitLenLog2Prime gt2 nat_int p_prime) lemma ECparamsValid_min_p: assumes "ECdomainParametersValid a b G n h t" shows "2^191 \ p" proof - have 1: "t \ {80, 112, 128, 192, 256}" using assms ECdomainParametersValid_def by algebra have 2: "bit_length p = twiceSecurityLevel t" by (metis assms bitLenLog2Prime gt2 nat_int p_prime ECdomainParametersValid_def) have 3: "192 \ twiceSecurityLevel t" using 1 twiceSecurityLevel_def by fastforce have 4: "192 \ bit_length p" using 2 3 by fastforce show ?thesis by (meson 4 less_bit_len less_le_trans numeral_less_iff semiring_norm(68,79,81)) qed text \The following lemma is convenient below when showing that valid EC domain parameters meet the definition of "Elliptic Prime Field" defined in EC_Common.\ lemma paramsValidEllPrimeField: assumes "ECdomainParametersValid a b G n h t" shows "a \ carrier R \ b \ carrier R \ (4*a^3 + 27*b^2) mod p \ 0" using ECdomainParametersValid_def assms nonsingularEq_nat by presburger lemma paramsValidEllPrimeField': assumes "ECdomainParametersValid a b G n h t" shows "ell_prime_field p a b" using assms ell_prime_field.intro ell_prime_field_axioms_def residues_prime_gt2.paramsValidEllPrimeField residues_prime_gt2_axioms by blast text \The following lemmas are useful for the Public Key Recovery Operation below.\ lemma p_div_n: assumes "ECdomainParametersValid a b G n h t" shows "p div n \ h" proof - let ?h = "((sqrt p + 1)^2) / n" have h1: "h = nat (floor ?h)" using assms(1) ECdomainParametersValid_def by fastforce have h2: "?h < h + 1" using h1 by linarith have f1: "(sqrt p + 1)^2 = p + 1 + 2*(sqrt p)" by (simp add: comm_semiring_1_class.power2_sum) have h3: "?h = (p + 2*(sqrt p) + 1) / n" using f1 by force have h4: "?h = (real p / real n) + ((2*(sqrt p) + 1)/n)" by (simp add: add_divide_distrib h3) have h5: "(real p / real n) \ ?h" using h4 by auto have h6: "(real p / real n) < h + 1" using h5 h2 by linarith have a1: "real (p div n) \ (real p / real n)" using real_of_nat_div4 by blast have a2: "p div n < h + 1" using h6 a1 by auto show ?thesis using a2 by simp qed lemma less_p_div_n: assumes "ECdomainParametersValid a b G n h t" "x < p" shows "x div n \ h" by (meson assms p_div_n div_le_mono le_trans less_imp_le_nat) lemma less_p_div_n': assumes "ECdomainParametersValid a b G n h t" "(int x) \ carrier R" shows "x div n \ h" by (meson assms(1,2) inCarrier less_p_div_n of_nat_less_iff) text \Next up we show that h < n. The number of points on the elliptic curve is n*h. For recommended curves, the cofactor h is typically small, like h < 10. The prime n is very large, around the size of p. So in some sense, it's not surprising that h < n. The only issue here is to show that given the definition of valid parameters.\ lemma h_less_n: assumes "ECdomainParametersValid a b G n h t" shows "h < n" proof - let ?t2 = "twiceSecurityLevel t" have t: "t \ {80, 112, 128, 192, 256}" using assms ECdomainParametersValid_def by algebra have t2: "2*t \ ?t2" using t twiceSecurityLevel_def by force have h: "h \ 2^(t div 8)" using assms ECdomainParametersValid_def by algebra have h1: "real h \ real 2^(t div 8)" using h by simp have h2: "0 < real (h + 1)" by auto have h3: "real (h+1) \ real (2^(t div 8) + 1)" using h by fastforce have p1: "\log 2 p\ = ?t2" using assms ECdomainParametersValid_def by algebra have p2: "bit_length p = ?t2" by (metis p1 bitLenLog2Prime gt2 of_nat_eq_iff p_prime) have p3: "2^(?t2-1) \ p" using bit_len_exact3 gt2 p2 by force have p4: "2^(2*t-1) \ p" by (meson p3 t2 diff_le_mono le_trans one_le_numeral power_increasing) have n1: "p div n \ h" using assms p_div_n by fast have n2: "\(real p / real n)\ \ h" by (simp add: n1 floor_divide_of_nat_eq) have n3: "(real p / real n) < h + 1" using n2 by linarith have n4: "prime n" using assms ECdomainParametersValid_def by algebra have n5: "0 < n" by (simp add: n4 prime_gt_0_nat) have n6: "0 < real n" using n5 by fastforce have n7: "real p < real n * (h+1)" using n3 n6 by (simp add: mult.commute pos_divide_less_eq) have n: "real p / real (h+1) < real n" using n7 h2 pos_divide_less_eq by blast have a1: "real 2^(2*t-1) / real (h+1) < real n" by (metis n p4 h2 less_le_trans not_less of_nat_power_less_of_nat_cancel_iff pos_divide_less_eq) have a2: "real 2^(2*t-1) / real (2^(t div 8)+1) \ real 2^(2*t-1) / real (h+1)" using h3 by (simp add: frac_le) have a3: "real 2^(2*t-1) / real (2^(t div 8)+1) < real n" using a1 a2 by argo have a4: "real 2^(t div 8) < real 2^(2*t-1) / real (2^(t div 8)+1)" using t by fastforce have "real h < real n" using a3 a4 h1 by argo then show ?thesis by simp qed text \It is sometimes helpful to know that n is an odd prime. In the following proof, we see if n = 2, then h = 1 and p = 2. We already know that p cannot be that small, so we are done. But also, if n = 2, the n*h = p, which also contradicts the definition of valid elliptic curve parameters.\ lemma n_not_2: assumes "ECdomainParametersValid a b G n h t" shows "2 < n" proof - have n0: "prime n" using assms ECdomainParametersValid_def by algebra have h0: "0 < h" using assms ECdomainParametersValid_def by algebra have p0: "2^191 \ p" using assms ECparamsValid_min_p by fast have n2: "n = 2 \ p \ 2" proof assume A: "n = 2" have A1: "h = 1" using A h0 h_less_n assms by fastforce have A2: "1 = \((sqrt p + 1)^2) div 2\" using A A1 assms ECdomainParametersValid_def by simp have A3: "((sqrt p + 1)^2) div 2 \ 2" using A2 by linarith have A4: "((sqrt p + 1)^2) \ 4" using A3 by simp show "p \ 2" using A4 real_le_rsqrt by force qed have "n \ 2" using p0 n2 by fastforce then show ?thesis using n0 le_neq_implies_less prime_ge_2_nat by presburger qed lemma n_odd: assumes "ECdomainParametersValid a b G n h t" shows "odd n" by (metis assms n_not_2 ECdomainParametersValid_def prime_odd_nat) text \In fact, n must be a great deal larger than 2.\ lemma ECparamsValid_min_n: assumes "ECdomainParametersValid a b G n h t" shows "2^95 < n" proof - have n1: "p div n \ h" using assms p_div_n by fast have n2: "\(real p / real n)\ \ h" by (simp add: n1 floor_divide_of_nat_eq) have n3: "(real p / real n) < h + 1" using n2 by linarith have n4: "prime n" using assms ECdomainParametersValid_def by algebra have n5: "0 < n" by (simp add: n4 prime_gt_0_nat) have n6: "0 < real n" using n5 by fastforce have n7: "real p < real n * (h+1)" using n3 n6 by (simp add: mult.commute pos_divide_less_eq) have n8: "p < n *(h+1)" by (metis n7 of_nat_less_imp_less semiring_1_class.of_nat_mult) have h1: "h+1 \ n" using assms h_less_n by fastforce have a1: "p < n^2" by (metis n8 h1 less_le_trans nat_1_add_1 nat_mult_le_cancel_disj power_add power_one_right) have a2: "2^191 < n^2" using a1 assms ECparamsValid_min_p by fastforce have a3: "((2::nat)^95)^2 = 2^190" by auto have a4: "((2::nat)^95)^2 < n^2" using a2 a3 by force show ?thesis using a4 power_less_imp_less_base by blast qed end (*residues_prime_gt2 context*) text \Now we know what valid elliptic curve parameters are. So we fix a set of parameters and define the cryptographic primitives below using those valid parameters.\ locale SEC1 = residues_prime_gt2 + fixes a b n h t :: nat and G :: "int point" assumes ECparamsValid: "ECdomainParametersValid a b G n h t" begin text \Elliptic_Locale defines elliptic curve operations. In this locale, we have fixed a set of elliptic curve parameters. We use the coefficients a and b to make convenient abbreviations for definitions found in Elliptic_Locale as well as a data conversion primitive defined above.\ abbreviation on_curve' :: "int point \ bool" where "on_curve' Q \ on_curve a b Q" abbreviation point_mult' :: "nat \ int point \ int point" where "point_mult' x Q \ point_mult a x Q" abbreviation add' :: "int point \ int point \ int point" where "add' P Q \ add a P Q" abbreviation octets_to_point' :: "octets \ int point option" where "octets_to_point' M \ octets_to_point a b M" text \We show that the valid elliptic curve parameters meet the assumptions of the ell_prime_field locale in EC_Common. We then write some properties in terms of "curve" which is defined in ell_prime_field. Of note: the definitions in Elliptic_Locale use additive notation for the elliptic curve group. For example, \add' P Q\, or "P + Q". In contrast, ell_prime_field uses multiplicative notation, so that the same operation is denoted \P\\<^bsub>EPF.curve\<^esub>Q\. \ sublocale EPF: ell_prime_field p R a b using ECparamsValid paramsValidEllPrimeField' apply blast by (simp add: R_def) lemma order_EPF_curve: "order EPF.curve = n*h" using order_def ECparamsValid ECdomainParametersValid_def EPF.curve_def by (metis partial_object.select_convs(1)) lemma order_EPF_curve_h1: assumes "h = 1" shows "order EPF.curve = n" using assms order_EPF_curve by simp lemma order_EPF_curve_h1': assumes "h = 1" shows "prime (order EPF.curve)" using assms order_EPF_curve_h1 ECparamsValid ECdomainParametersValid_def by blast text \When h=1, then all points on the curve can be written as some power of the generator G. (And any point on the curve can be picked as the generator, for that matter. It's just a cyclic group.)\ lemma EC_h1_cyclic: assumes "h = 1" shows "carrier EPF.curve = {Q. (\d G \ Infinity \ point_mult' n G = Infinity\ prime (n::nat)" using ECparamsValid ECdomainParametersValid_def by blast let ?S1 = "carrier EPF.curve" let ?S2 = "{Q. (\dd. on_curve' (point_mult' d G)" by (simp add: H point_mult_closed) have 4: "?S2 \ ?S1" using 3 EPF.curve_def by force show ?thesis using 1 2 4 EPF.finite_curve card_subset_eq by blast qed lemma EC_h1_cyclic': assumes "h = 1" "on_curve' P" shows "\dNext up, we write down some facts about the generator G, and its multiples P = d*G (or powers when written with multiplicative notation, P = G^d).\ lemma GonEPFcurve: "G \ carrier EPF.curve" using ECdomainParametersValid_def ECparamsValid EPF.in_carrier_curve by presburger lemma dGonEPFcurve: assumes "P = point_mult' d G" shows "P \ carrier EPF.curve" using GonEPFcurve EPF.multEqPowInCurve EPF.curve.nat_pow_closed EPF.in_carrier_curve assms by force lemma Gnot1onEPFcurve: "G \ \\<^bsub>EPF.curve\<^esub>" using ECdomainParametersValid_def ECparamsValid EPF.one_curve by presburger lemma nGeq1onEPFcurve: "G [^]\<^bsub>EPF.curve\<^esub> n = \\<^bsub>EPF.curve\<^esub>" proof - have "point_mult' n G = Infinity" using ECdomainParametersValid_def ECparamsValid by presburger then show ?thesis using GonEPFcurve EPF.in_carrier_curve EPF.one_curve EPF.multEqPowInCurve by auto qed lemma invG: "inv\<^bsub>EPF.curve\<^esub> G = G [^]\<^bsub>EPF.curve\<^esub> (n-1)" proof - have "prime n" using ECparamsValid ECdomainParametersValid_def by blast then have "1 < n" using prime_gt_1_nat by simp then have "1 + (n-1) = n" by auto then have "G \\<^bsub>EPF.curve\<^esub> (G [^]\<^bsub>EPF.curve\<^esub> (n-1)) = \\<^bsub>EPF.curve\<^esub>" by (metis nGeq1onEPFcurve EPF.curve.nat_pow_Suc2 GonEPFcurve plus_1_eq_Suc) then show ?thesis using EPF.curve.comm_inv_char EPF.curve.nat_pow_closed GonEPFcurve by presburger qed lemma ordGisn: assumes "x < n" "0 < x" shows "point_mult' x G \ Infinity" using ECdomainParametersValid_def ECparamsValid assms(1,2) EPF.order_n by presburger lemma ordGisn': assumes "x < n" "0 < x" shows "G [^]\<^bsub>EPF.curve\<^esub> x \ \\<^bsub>EPF.curve\<^esub>" using assms ordGisn EPF.one_curve EPF.multEqPowInCurve GonEPFcurve EPF.in_carrier_curve by algebra lemma curve_ord_G: "EPF.curve.ord G = n" using ECdomainParametersValid_def ECparamsValid EPF.curve_ord_n5 by presburger lemma multGmodn: "point_mult' x G = point_mult' (x mod n) G" by (metis ECdomainParametersValid_def ECparamsValid EPF.multQmodn) lemma multGmodn': "G[^]\<^bsub>EPF.curve\<^esub>x = G[^]\<^bsub>EPF.curve\<^esub>(x mod n)" using multGmodn EPF.in_carrier_curve GonEPFcurve EPF.multEqPowInCurve by fastforce lemma multdGmodn: assumes "P = point_mult' d G" shows "point_mult' x P = point_mult' (x mod n) P" by (metis ECdomainParametersValid_def ECparamsValid EPF.in_carrier_curve EPF.multQmodn EPF.order_n_cycle assms dGonEPFcurve) lemma multdGmodn': assumes "P = point_mult' d G" shows "P[^]\<^bsub>EPF.curve\<^esub>x = P[^]\<^bsub>EPF.curve\<^esub>(x mod n)" using assms multdGmodn EPF.in_carrier_curve dGonEPFcurve EPF.multEqPowInCurve by fastforce lemma multGmodn'int: "G[^]\<^bsub>EPF.curve\<^esub>(x::int) = G[^]\<^bsub>EPF.curve\<^esub>(x mod n)" by (metis EPF.in_carrier_curve EPF.one_curve GonEPFcurve EPF.multEqPowInCurve EPF.multQmodn'int nGeq1onEPFcurve prime_gt_1_nat ECparamsValid ECdomainParametersValid_def) lemma multdGmodn'int: assumes "P = point_mult' d G" shows "P[^]\<^bsub>EPF.curve\<^esub>(x::int) = P[^]\<^bsub>EPF.curve\<^esub>(x mod n)" by (metis assms ECparamsValid EPF.in_carrier_curve EPF.multQmodn'int EPF.order_n_cycle dGonEPFcurve prime_gt_1_nat ECdomainParametersValid_def) text \Computations done on the components of elliptic curve points are done modulo p. The order of the curve, however, is n, a prime different from p. So, for example, the scalar multiplication (n+a)G = aG, because the generator G has order n. So when computations are done on scalars intended for a scalar multiplication of a point on the curve, those should be reduced modulo n. In particular, we will need to compute inverses of scalars modulo n.\ definition "Rn = residue_ring (int n)" sublocale N: residues_prime n Rn using ECdomainParametersValid_def ECparamsValid residues_prime_def apply presburger using Rn_def by presburger abbreviation inv_mod_n :: "nat \ nat" where "inv_mod_n x \ inv_mod n x" lemma inv_mod_n_inv: assumes "x \ carrier Rn" "x \ 0" shows "inv\<^bsub>Rn\<^esub> x = inv_mod_n x" using N.inv_mod_p_inv_in_ring assms(1,2) by presburger lemma inv_mod_n_inv': assumes "0 < x" "x < n" shows "inv\<^bsub>Rn\<^esub> x = inv_mod_n x" using assms inv_mod_n_inv by (simp add: N.res_carrier_eq) subsection \3.2 Elliptic Curve Key Pairs\ text \3.2.1 Elliptic Curve Key Pair Generation Primitive This section doesn't exactly define a validation primitive. But from the Key Pair Generation Primitive, we can write down what it means to be a valid EC key pair. d is the private key and Q is the public key. We can also write down separately what it means for d to be a valid private key. Also note that the key generation method in 3.2.1 says to randomly select d. As we said above, we do not model random selection here, so we include d as an input.\ definition ECkeyGen :: "nat \ int point" where "ECkeyGen d = point_mult' d G" definition ECkeyPairValid :: "nat \ int point \ bool" where "ECkeyPairValid d Q \ (0 < d) \ (d < n) \ (point_mult' d G = Q)" definition ECprivateKeyValid :: "nat \ bool" where "ECprivateKeyValid d \ (0 < d) \ (d < n)" lemma ECkeyPairImpliesKeyGen: assumes "ECkeyPairValid d Q" shows "Q = ECkeyGen d" using ECkeyGen_def ECkeyPairValid_def assms by presburger lemma ECkeyPairEqKeyGen: "(ECkeyPairValid d Q) = ((Q = ECkeyGen d) \ ECprivateKeyValid d)" using ECkeyGen_def ECkeyPairValid_def ECprivateKeyValid_def by auto lemma ECkeyPairImpliesPrivateKeyValid: assumes "ECkeyPairValid d Q" shows "ECprivateKeyValid d" using assms ECkeyPairValid_def ECprivateKeyValid_def by auto lemma ECkeyPairNotInf: assumes "ECkeyPairValid d Q" shows "Q \ Infinity" by (metis assms ECkeyPairValid_def ordGisn) lemma ECkeyPairOnCurve: assumes "ECkeyPairValid d Q" shows "on_curve' Q" by (metis ECdomainParametersValid_def ECkeyPairValid_def ECparamsValid assms point_mult_closed) lemma ECkeyPairOrd_n: assumes "ECkeyPairValid d Q" shows "point_mult' n Q = Infinity" by (metis ECkeyPairValid_def assms mod_self multdGmodn point_mult.simps(1)) lemma ECkeyPair_dInRn': assumes "ECprivateKeyValid d" shows "d \ carrier Rn" using ECprivateKeyValid_def N.res_carrier_eq assms by force lemma ECkeyPair_dInRn: assumes "ECkeyPairValid d Q" shows "d \ carrier Rn" using ECkeyPairValid_def N.res_carrier_eq assms by force lemma ECkeyPair_invRn': assumes "ECprivateKeyValid d" shows "inv\<^bsub>Rn\<^esub>d \ carrier Rn \ d\\<^bsub>Rn\<^esub>inv\<^bsub>Rn\<^esub>d = \\<^bsub>Rn\<^esub>" using ECprivateKeyValid_def N.inv_closed N.res_carrier_eq N.zero_cong assms by force lemma ECkeyPair_invRn: assumes "ECkeyPairValid d Q" shows "inv\<^bsub>Rn\<^esub>d \ carrier Rn \ d\\<^bsub>Rn\<^esub>inv\<^bsub>Rn\<^esub>d = \\<^bsub>Rn\<^esub>" using ECkeyPairValid_def N.inv_closed N.res_carrier_eq N.zero_cong assms by force lemma ECkeyPair_inv_d': assumes "ECprivateKeyValid d" shows "inv_mod_n d = inv\<^bsub>Rn\<^esub>d" using ECprivateKeyValid_def N.res_carrier_eq assms inv_mod_n_inv by force lemma ECkeyPair_inv_d: assumes "ECkeyPairValid d Q" shows "inv_mod_n d = inv\<^bsub>Rn\<^esub>d" using ECkeyPairValid_def N.res_carrier_eq assms inv_mod_n_inv by force lemma ECkeyPair_curveMult: assumes "ECkeyPairValid d Q" shows "G[^]\<^bsub>EPF.curve\<^esub>d = Q" using ECkeyPairValid_def EPF.in_carrier_curve GonEPFcurve assms EPF.multEqPowInCurve by auto lemma ECkeyGen_mod_n: "ECkeyGen d = ECkeyGen (d mod n)" using ECkeyGen_def multGmodn by presburger lemma ECkeyGen_valid_mod_n: "(d mod n \ 0) = ECprivateKeyValid (d mod n)" by (simp add: ECprivateKeyValid_def N.p_prime prime_gt_0_nat) lemma ECkeyGen_int: fixes d :: int shows "G[^]\<^bsub>EPF.curve\<^esub>d = ECkeyGen (nat (d mod int n))" using EPF.ell_prime_field_axioms GonEPFcurve N.p_prime SEC1.ECkeyGen_def SEC1_axioms ell_prime_field.in_carrier_curve ell_prime_field.multEqPowInCurve multGmodn'int prime_gt_0_nat by force text \3.2.2.1 Elliptic Curve Public Key Validation Primitive\ definition ECpublicKeyValid :: "int point \ bool" where "ECpublicKeyValid Q \ (Q \ Infinity) \ (on_curve' Q) \ (point_mult' n Q = Infinity)" text \3.2.3.1 Elliptic Curve Public Key Partial Validation Primitive\ definition ECpublicKeyPartialValid :: "int point \ bool" where "ECpublicKeyPartialValid Q \ (Q \ Infinity) \ (on_curve' Q)" lemma validImpliesPartialValid: "ECpublicKeyValid Q \ ECpublicKeyPartialValid Q" using ECpublicKeyPartialValid_def ECpublicKeyValid_def by blast text \In general, if we have a partially valid public key, it is not necessarily (fully) valid. However, when h = 1, then this is true.\ lemma partValidImpliesValidIFheq1: assumes "h = 1" "ECpublicKeyPartialValid Q" shows "ECpublicKeyValid Q" proof - have 1: "Q \ Infinity" using assms(2) ECpublicKeyPartialValid_def by fast have 2: "on_curve' Q" using assms(2) ECpublicKeyPartialValid_def by fast have 3: "card {P. on_curve a b P} = n" using ECparamsValid ECdomainParametersValid_def assms(1) by simp have 4: "point_mult' n Q = Infinity" using 2 3 CurveOrderPrime EPF.AB_in_carrier EPF.nonsingular_in_bf N.p_prime by blast show ?thesis using 1 2 4 ECpublicKeyValid_def by presburger qed lemma partValidImpliesValidIFheq1': assumes "h = 1" "Q \ Infinity" "on_curve' Q" shows "point_mult' n Q = Infinity" using partValidImpliesValidIFheq1 assms ECpublicKeyPartialValid_def ECpublicKeyValid_def by blast lemma keyPairValidImpliespublicKeyValid: assumes "ECkeyPairValid d Q" shows "ECpublicKeyValid Q" using ECkeyPairNotInf ECkeyPairOnCurve ECkeyPairOrd_n ECpublicKeyValid_def assms by blast lemma ECpublicKeyValid_curve: assumes "ECpublicKeyValid Q" shows "Q \ \\<^bsub>EPF.curve\<^esub> \ Q \ carrier EPF.curve \ Q[^]\<^bsub>EPF.curve\<^esub> n = \\<^bsub>EPF.curve\<^esub>" using assms ECpublicKeyValid_def EPF.in_carrier_curve EPF.one_curve EPF.multEqPowInCurve by auto lemma ECpublicKeyPartValid_curve: assumes "ECpublicKeyPartialValid Q" shows "Q \ \\<^bsub>EPF.curve\<^esub> \ Q \ carrier EPF.curve" using assms ECpublicKeyPartialValid_def EPF.in_carrier_curve EPF.one_curve by auto lemma ECkeyGenValid: assumes "ECprivateKeyValid d" shows "ECpublicKeyValid (ECkeyGen d)" using ECkeyGen_def ECkeyPairValid_def ECprivateKeyValid_def assms keyPairValidImpliespublicKeyValid by auto lemma ECKeyGenValidPair: assumes "ECprivateKeyValid d" shows "ECkeyPairValid d (ECkeyGen d)" using assms ECprivateKeyValid_def ECkeyGen_def ECkeyPairValid_def by simp subsection \3.3 Elliptic Curve Diffie-Hellman Primitives\ text \3.3.1 Elliptic Curve Diffie-Hellman Primitive\ definition ECDHprim :: "nat \ int point \ int option" where "ECDHprim d\<^sub>U Q\<^sub>V = ( let P = point_mult' d\<^sub>U Q\<^sub>V in ( case P of Infinity \ None | Point x\<^sub>P y\<^sub>P \ Some x\<^sub>P ) )" lemma ECDHinCarrier: assumes "ECDHprim d Q = Some z" "on_curve' Q" shows "z \ carrier R" proof - let ?P = "point_mult' d Q" have 1: "?P \ Infinity" using assms(1) ECDHprim_def by fastforce obtain x and y where 2: "?P = Point x y" by (meson 1 point.exhaust) have 3: "z = x" using 1 2 assms(1) ECDHprim_def by auto have 4: "on_curve' ?P" using assms(2) point_mult_closed by auto show ?thesis using 2 3 4 on_curve_def by auto qed lemma ECDH_validKeys: assumes "ECprivateKeyValid d" "ECpublicKeyValid Q" shows "\z. ECDHprim d Q = Some z" proof - have 1: "0 < d \ d < n" using assms(1) ECprivateKeyValid_def by simp have 2: "on_curve' Q \ Q \ Infinity \ point_mult' n Q = Infinity" using assms(2) ECpublicKeyValid_def by blast let ?P = "point_mult' d Q" have 3: "?P \ Infinity" using 1 2 EPF.order_n by blast obtain x and y where 4: "?P = Point x y" using 3 point.exhaust by blast have 5: "ECDHprim d Q = Some x" using 4 ECDHprim_def by force show ?thesis using 5 by fast qed lemma ECDH_curveMult: assumes "on_curve' Q" shows "ECDHprim d Q = ( let P = Q [^]\<^bsub>EPF.curve\<^esub> d in ( case P of Infinity \ None | Point x\<^sub>P y\<^sub>P \ Some x\<^sub>P ))" using assms ECDHprim_def EPF.multEqPowInCurve by presburger text \The important thing about Diffie-Hellman is that two entities can compute the same value using only their own (private) key and the other's public information. This lemma shows that.\ lemma ECDH_2ValidKeyPairs: assumes "ECkeyPairValid d1 P1" "ECkeyPairValid d2 P2" shows "ECDHprim d1 P2 = ECDHprim d2 P1" by (metis (no_types, lifting) assms(1,2) ECDHprim_def ECkeyPairValid_def EPF.AB_in_carrier(1,2) EPF.in_carrier_curve EPF.nonsingular_in_bf GonEPFcurve mult.commute point_mult_mult) text \3.3.2 Elliptic Curve Cofactor Diffie-Hellman Primitive\ definition ECcofDHprim :: "nat \ int point \ int option" where "ECcofDHprim d\<^sub>U Q\<^sub>V = ( let P = point_mult' (h*d\<^sub>U) Q\<^sub>V in ( case P of Infinity \ None | Point x\<^sub>P y\<^sub>P \ Some x\<^sub>P ) )" lemma ECcofDHinCarrier: assumes "ECcofDHprim d Q = Some z" "on_curve' Q" shows "z \ carrier R" proof - let ?P = "point_mult' (h*d) Q" have 1: "?P \ Infinity" using assms(1) ECcofDHprim_def by fastforce obtain x and y where 2: "?P = Point x y" by (meson 1 point.exhaust) have 3: "z = x" using 1 2 assms(1) ECcofDHprim_def by auto have 4: "on_curve' ?P" using assms(2) point_mult_closed by auto show ?thesis using 2 3 4 on_curve_def by auto qed lemma ECcofDH_validKeys: assumes "ECprivateKeyValid d" "ECpublicKeyValid Q" shows "\z. ECcofDHprim d Q = Some z" proof - have 1: "0 < d \ d < n" using assms(1) ECprivateKeyValid_def by simp have 2: "on_curve' Q \ Q \ Infinity \ point_mult' n Q = Infinity" using assms(2) ECpublicKeyValid_def by blast have 3: "h < n" using h_less_n ECparamsValid by auto have 4: "0 < h" using ECdomainParametersValid_def ECparamsValid by presburger have 5: "\ n dvd d" using 1 by fastforce have 6: "\ n dvd h" using 3 4 by fastforce have 7: "\ n dvd (h*d)" by (simp add: 5 6 N.p_coprime_left coprime_dvd_mult_left_iff) let ?m = "(h*d) mod n" have 8: "0 < ?m" using 7 mod_greater_zero_iff_not_dvd by blast have 9: "?m < n" using 3 by fastforce let ?P = "point_mult' (h*d) Q" have 10: "?P = point_mult' ?m Q" using EPF.multQmodn assms(2) ECpublicKeyValid_def by blast have 11: "?P \ Infinity" using 8 9 2 10 EPF.order_n by auto obtain x and y where 12: "?P = Point x y" using 11 point.exhaust by blast have "ECcofDHprim d Q = Some x" using 12 ECcofDHprim_def by force then show ?thesis by fast qed lemma ECcoDH_curveMult: assumes "on_curve' Q" shows "ECcofDHprim d Q = ( let P = Q [^]\<^bsub>EPF.curve\<^esub> (h*d) in ( case P of Infinity \ None | Point x\<^sub>P y\<^sub>P \ Some x\<^sub>P ))" using assms ECcofDHprim_def EPF.multEqPowInCurve by presburger text \The important thing about Diffie-Hellman is that two entities can compute the same value using only their own (private) key and the other's public information. This lemma shows that.\ lemma ECcofDH_2ValidKeyPairs: assumes "ECkeyPairValid k1 P1" "ECkeyPairValid k2 P2" shows "ECcofDHprim k1 P2 = ECcofDHprim k2 P1" proof - have 1: "P1 = point_mult' k1 G" using assms(1) by (simp add: ECkeyPairValid_def) have 2: "P2 = point_mult' k2 G" using assms(2) by (simp add: ECkeyPairValid_def) have 3: "point_mult' (h*k2) P1 = point_mult' (k1*(h*k2)) G" using 1 EPF.AB_in_carrier(1,2) EPF.in_carrier_curve EPF.nonsingular_in_bf GonEPFcurve point_mult_mult by presburger have 4: "point_mult' (h*k1) P2 = point_mult' (k2*(h*k1)) G" using 2 EPF.AB_in_carrier(1,2) EPF.in_carrier_curve EPF.nonsingular_in_bf GonEPFcurve point_mult_mult by presburger show ?thesis by (metis 3 4 mult.commute mult.left_commute ECcofDHprim_def) qed text \3.3.* Choose either "normal" or cofactor Diffie-Hellman. In some of the schemes below, the entities U and V must agree to use either "normal" or cofactor Diffie-Hellman. This primitive is convenient to use in those definitions.\ definition ECDHprimChoose :: "bool \ nat \ int point \ int option" where "ECDHprimChoose useCoDH d\<^sub>U Q\<^sub>V = (if useCoDH then (ECcofDHprim d\<^sub>U Q\<^sub>V) else (ECDHprim d\<^sub>U Q\<^sub>V) )" lemma ECDHchooseinCarrier: assumes "ECDHprimChoose useCo d Q = Some z" "on_curve' Q" shows "z \ carrier R" using ECDHprimChoose_def ECDHinCarrier ECcofDHinCarrier assms by presburger lemma ECDHchoose_validKeys: assumes "ECprivateKeyValid d" "ECpublicKeyValid Q" shows "\z. ECDHprimChoose useCo d Q = Some z" using ECDHprimChoose_def ECDH_validKeys ECcofDH_validKeys assms by presburger text \The important thing is that two entities can compute the same value using only their own (private) key together with the other's public information.\ lemma ECDHch_2ValidKeyPairs: assumes "ECkeyPairValid k1 P1" "ECkeyPairValid k2 P2" shows "ECDHprimChoose useCoDH k1 P2 = ECDHprimChoose useCoDH k2 P1" by (simp add: ECDHprimChoose_def ECDH_2ValidKeyPairs ECcofDH_2ValidKeyPairs assms(1,2)) subsection \3.4 Elliptic Curve MQV Primitive Compared to the Diffie-Hellman primitive, the MQV primitive is involved. So we break the definition into smaller chunks. MQVcomputeQbar corresponds to steps 2 and 4 in the standard. MQVcompute_s is step 3. MQVcomputeP is step 5.\ definition MQVcomputeQbar :: "int point \ nat option" where "MQVcomputeQbar Q = (case Q of Infinity \ None | Point x y \ ( let z = (2::nat)^(nat (\(log 2 n)/2\)); xBar = nat (x mod z) in Some (xBar + z) ))" definition MQVcompute_s :: "nat \ nat option \ nat \ nat option" where "MQVcompute_s d2U Q2Ubar d1U = ( case Q2Ubar of None \ None | Some Q2Ubar' \ Some ((d2U + Q2Ubar'*d1U) mod n))" definition MQVcomputeP :: "nat option \ int point \ nat option \ int point \ int point option" where "MQVcomputeP s Q2V Q2Vbar Q1V = ( case s of None \ None | Some s' \ (case Q2Vbar of None \ None | Some Q2Vbar' \ Some ( point_mult' (h*s') (add' Q2V (point_mult' Q2Vbar' Q1V)) ) ) )" definition ECMQVprim_validInput :: "nat \ int point \ nat \ int point \ int point \ int point \ bool" where "ECMQVprim_validInput d1U Q1U d2U Q2U Q1V Q2V \ (ECkeyPairValid d1U Q1U) \ (ECkeyPairValid d2U Q2U) \ (ECpublicKeyPartialValid Q1V) \ (ECpublicKeyPartialValid Q2V)" text \The standard has (d1U, Q1U) as one of the inputs, but the curve point Q1U does not appear in the computation. We write this function to be consistent with the text of the standard.\ definition ECMQVprim :: "nat \ int point \ nat \ int point \ int point \ int point \ int option" where "ECMQVprim d1U Q1U d2U Q2U Q1V Q2V = ( let Q2Ubar = MQVcomputeQbar Q2U; s = MQVcompute_s d2U Q2Ubar d1U; Q2Vbar = MQVcomputeQbar Q2V; P = MQVcomputeP s Q2V Q2Vbar Q1V in (case P of None \ None | Some Infinity \ None | Some (Point x y) \ Some x ) )" lemma MQVcomputeQbar_notInf: assumes "Q \ Infinity" shows "\Qbar. MQVcomputeQbar Q = Some Qbar" by (metis (no_types, lifting) MQVcomputeQbar_def assms point.case(2) point.exhaust) lemma MQVcomputeQbar_validPair: assumes "ECkeyPairValid d2U Q2U" shows "\Q2Ubar. MQVcomputeQbar Q2U = Some Q2Ubar" using ECkeyPairNotInf MQVcomputeQbar_notInf assms by blast lemma MQVcomputeQbar_validPub: assumes "ECpublicKeyPartialValid Q2V" shows "\Q2Vbar. MQVcomputeQbar Q2V = Some Q2Vbar" using ECpublicKeyPartialValid_def MQVcomputeQbar_notInf assms by presburger lemma MQVcompute_s_validIn: assumes "ECkeyPairValid d2U Q2U" "Q2Ubar = MQVcomputeQbar Q2U" shows "\s. MQVcompute_s d2U Q2Ubar d1U = Some s" using MQVcomputeQbar_validPair MQVcompute_s_def assms(1,2) by fastforce lemma MQVcomputeP_validIn: assumes "ECkeyPairValid d2U Q2U" "ECpublicKeyPartialValid Q2V" "Q2Ubar = MQVcomputeQbar Q2U" "s = MQVcompute_s d2U Q2Ubar d1U" "Q2Vbar = MQVcomputeQbar Q2V" shows "\P. MQVcomputeP s Q2V Q2Vbar Q1V = Some P" using MQVcomputeP_def MQVcomputeQbar_validPair MQVcomputeQbar_validPub MQVcompute_s_def assms by fastforce text \The important thing is that both U and V can compute the same value using only their own key pairs and the other's public keys. This lemma shows that is true.\ lemma MQV_reverseUV: assumes "ECkeyPairValid d1U Q1U" "ECkeyPairValid d2U Q2U" "ECkeyPairValid d1V Q1V" "ECkeyPairValid d2V Q2V" shows "ECMQVprim d1U Q1U d2U Q2U Q1V Q2V = ECMQVprim d1V Q1V d2V Q2V Q1U Q2U" proof - have U11: "Q1U = point_mult' d1U G" using assms(1) by (simp add: ECkeyPairValid_def) have U21: "Q2U = point_mult' d2U G" using assms(2) by (simp add: ECkeyPairValid_def) have V11: "Q1V = point_mult' d1V G" using assms(3) by (simp add: ECkeyPairValid_def) have V21: "Q2V = point_mult' d2V G" using assms(4) by (simp add: ECkeyPairValid_def) obtain Q2Ubar where Q2Ubar: "MQVcomputeQbar Q2U = Some Q2Ubar" using MQVcomputeQbar_validPair assms(2) by presburger obtain Q2Vbar where Q2Vbar: "MQVcomputeQbar Q2V = Some Q2Vbar" using MQVcomputeQbar_validPair assms(4) by presburger obtain sU where sU: "MQVcompute_s d2U (Some Q2Ubar) d1U = Some sU" by (simp add: MQVcompute_s_def) have sU1: "sU = (d2U + Q2Ubar*d1U) mod n" using sU MQVcompute_s_def by simp obtain sV where sV: "MQVcompute_s d2V (Some Q2Vbar) d1V = Some sV" by (simp add: MQVcompute_s_def) have sV1: "sV = (d2V + Q2Vbar*d1V) mod n" using sV MQVcompute_s_def by simp obtain PU where PU: "MQVcomputeP (Some sU) Q2V (Some Q2Vbar) Q1V = Some PU" by (simp add: MQVcomputeP_def) have U1: "PU = point_mult' (h*sU) (add' Q2V (point_mult' Q2Vbar Q1V))" using MQVcomputeP_def PU by auto have U2: "PU = point_mult' (h*sU) (add' (point_mult' d2V G) (point_mult' Q2Vbar (point_mult' d1V G)))" using U1 V11 V21 by fast have U3: "PU = point_mult' (h*sU) (point_mult' (d2V + Q2Vbar*d1V) G)" by (metis U2 point_mult_add EPF.AB_in_carrier(1,2) EPF.in_carrier_curve EPF.nonsingular_in_bf GonEPFcurve V11 mult.commute point_mult_mult) have U4: "PU = point_mult' (h*sU*(d2V + Q2Vbar*d1V)) G" by (metis U3 point_mult_mult EPF.AB_in_carrier(1,2) EPF.in_carrier_curve EPF.nonsingular_in_bf GonEPFcurve mult.commute) let ?xU = "(h*sU*(d2V + Q2Vbar*d1V)) mod n" have U5: "PU = point_mult' ?xU G" using U4 multGmodn by blast obtain PV where PV: "MQVcomputeP (Some sV) Q2U (Some Q2Ubar) Q1U = Some PV" by (simp add: MQVcomputeP_def) have V1: "PV = point_mult' (h*sV) (add' Q2U (point_mult' Q2Ubar Q1U))" using MQVcomputeP_def PV by auto have V2: "PV = point_mult' (h*sV) (add' (point_mult' d2U G) (point_mult' Q2Ubar (point_mult' d1U G)))" using V1 U11 U21 by fast have V3: "PV = point_mult' (h*sV) (point_mult' (d2U + Q2Ubar*d1U) G)" by (metis V2 point_mult_add EPF.AB_in_carrier(1,2) EPF.in_carrier_curve EPF.nonsingular_in_bf GonEPFcurve U11 mult.commute point_mult_mult) have V4: "PV = point_mult' (h*sV*(d2U + Q2Ubar*d1U)) G" by (metis V3 point_mult_mult EPF.AB_in_carrier(1,2) EPF.in_carrier_curve EPF.nonsingular_in_bf GonEPFcurve mult.commute) let ?xV = "(h*sV*(d2U + Q2Ubar*d1U)) mod n" have V5: "PV = point_mult' ?xV G" using V4 multGmodn by blast have x1: "?xU = ?xV" by (metis sU1 sV1 mult.commute mod_mult_right_eq mult.assoc) have x2: "PU = PV" using x1 U5 V5 by argo show ?thesis using x2 PU PV Q2Ubar Q2Vbar sU sV ECMQVprim_def by presburger qed end (* of SEC1 locale *) subsection \3.5 Hash Functions\ text \Section 3.5 of SEC 1 covers the HASH functions that are supported by this standard. The supported hash functions are SHA-1, 224, 256, 384, and 512, all defined in FIPS 180, currently in version 4. We have translated that standard to HOL in FIPS180_4.thy. In this theory, we don't define any hash function. We only need to know the type of a hash function. For that, we use the type synonym below. The key fact we will use for a hash function is that for any input, the length of the output has hLen octets, for some fixed hLen. For example, SHA-256 always outputs hLen = 32 octets (256 bits). There are limits on the length of a message that should be given to a SHA function, however as we show in FIPS180_4.thy, you can still compute a hash output for any length of input message and that output will have the correct length. If Hash is hash_type, we can insist \\x. octets_valid (Hash x)\ and \\x. length (Hash x) = hLen\. \ type_synonym hash_type = "octets \ octets" subsection \3.6 Key Derivation Functions\ text \This sections lists the key derivation functions (KDFs) that are supported by this standard. Those are ANSI-X9.63-KDF, IKEv2-KDF, TLS-KDF, NIST-800-56-Concatenation-KDF. The specification states "The NIST-800-56-Catenation-KDF should be used, except for backwards compatability with implementations already using one of the three other key derivation functions." Again, we don't define any key derivation functions here. We need only a type for key derivation functions. The example in 3.6.1 shows that the input is an octet string Z and a non-negative integer keydatalen. It has an optional input called SharedInfo which is also octets. If the optional SharedInfo is not used, that input will be the empty list []. It produces an octet string of length keydatalen. So we have the type of a key derivation function. We also introduce the type of a function which indicates if the input is valid for the key derivation function. So if KDF is kdf_type and KDF_ValidIn is kdf_validin_type, then we can insist that \\Z l I. KDF_ValidIn Z l I \ length (KDF Z l I) = l \ octets_valid (KDF Z l I)\\ type_synonym kdf_type = "octets \ nat \ octets \ octets" type_synonym kdf_validin_type = "octets \ nat \ octets \ bool" subsection \3.7 MAC Schemes\ text \This section lists the MAC Schemes that are supported by this standard. There are many, based on HMAC (using various SHAs) and CMAC (using AES of various sizes). A MAC scheme will be used by ECIES in Section 5.1 below. Again we need only to define a type for MAC functions. The example given in 3.7 says that the entities need to agree on a MAC scheme plus mackeylen and maclen. Then they establish a shared secret key K, an octet string of length mackeylen. Then the tagging function takes an input an octet string M and outputs an octet string D of length maclen. Since mackeylen = length K, we don't need to list it separately as an input. And maclen is the length of the output of MAC, so it's not an input as much as a feature of the MAC scheme. So our type synonym for a generic MAC function, using the variable names in the standard, follows MAC(K, M) = D. We also give a mac_validin_type. So if MAC is mac_type, and MAC_VI is mac_validin_type, We could insist that \\K M. MAC_VI K M \ length K = mackeylen\ and \\K M. MAC_VI K M \ length (MAC K M) = maclen\ and \\K M. MAC_VI K M \ octets_valid (MAC K M)\. \ type_synonym mac_type = "octets \ octets \ octets" type_synonym mac_validin_type = "octets \ octets \ bool" subsection \3.8 Symmetric Encryption Schemes\ text \This standard supports 3-key TDES, XOR, and AES as symmetric encryption schemes. And encryption scheme, labelled ENC later in this standard, takes a shared private key K and a message M and produces a ciphertext C. All of these are octet strings. The type of the decryption is the same: it takes K and C and produces M. If ENC is enc_type, ENC_VI is enc_validin_type, DEC is dec_type, and DEC_VI is dec_valid_in_type, we might insist \ENC_VI K M \ DEC_VI K (ENC K M) \ DEC K (ENC K M) = M\ and \DEC_VI K C \ ENC_VI K (DEC K C) \ ENC K (DEC K C) = C\ \ type_synonym enc_type = "octets \ octets \ octets" type_synonym enc_validin_type = "octets \ octets \ bool" type_synonym dec_type = "octets \ octets \ octets" type_synonym dec_validin_type = "octets \ octets \ bool" subsection \3.9 Key Wrap Schemes\ text \This subsection specifies that either the NIST AES key wrap algorithm or the CMS TDES key wrap algorithm • must be used as the key wrap scheme in the Wrapped Key Transport Scheme, and • should be used more generally when wrapping an existing symmetric key with another sym- metric key. A key wrap scheme takes as input a key-encryption key K, an octet string of length wrapkeylen, and an octet string C which is the key that needs wrapping. It produces an octet string W, which is the wrapped key. If WRAP is wrap_type, WRAP_VI is wrap_validin_type, UNWRAP is unwrap_type, and UNWRAP_VI is unwrap_validin_type, we might insist \WRAP_VI K C \ UNWRAP_VI K (WRAP K C) \ UNWRAP K (WRAP K C) = C\ and \UNWRAP_VI K W \ WRAP_VI K (UNWRAP K W) \ WRAP K (UNWRAP K W) = W\ \ type_synonym wrap_type = "octets \ octets \ octets" type_synonym wrap_validin_type = "octets \ octets \ bool" type_synonym unwrap_type = "octets \ octets \ octets" type_synonym unwrap_validin_type = "octets \ octets \ bool" subsection \3.10 Random Number Generation\ text \Cryptographic keys must be generated in a way that prevents an adversary from guessing the private key. Keys should be generated with the help of a random number generator. Random number generators must comply with ANS X9.82 [X9.82] or corresponding NIST publi- cation [800-90]. We don't model random number generation in this theory. When the standard says "pick a random number", we simply include that random value as an input to the function.\ subsection \3.11 Security Levels and Protection Lifetimes\ text \"Data protected with cryptography today may continue to need protection in the future. Advances in cryptanalysis can be predicted, at least approximately. Based on current approximations, this document requires that data that needs protection beyond the year 2010 must be protected with 112-bit security or higher. Data that needs protection beyond the year 2030 must be protected with 128-bit security or higher. Data that needs protection beyond the year 2040 should be protected with 192-bit security or higher. Data that needs protection beyond 2080 should be protected with 256-bit security or higher."\ section \4 Signature Schemes\ type_synonym sig_type = "nat \ nat" subsection \4.1 Elliptic Curve Digital Signature Algorithm\ text \In the setup (Section 4.1.1) for the Elliptic Curve Digital Signature Algorithm (ECDSA), the "entities" U and V need to agree on a few things. They have to pick a hash algorithm. Above we discuss that the only approved hash algorithms are found in FIPS 180, The Secure Hash Standard. This hash is used only in the computation of the message digest e. They must also agree on an elliptic curve. Also (Section 4.1.2) Entity U should generate a key pair \(d\<^sub>U, Q\<^sub>U)\. Entity V should obtain \Q\<^sub>U\.\ subsubsection \4.1.3 Signing Operation\ text \Note in the standard, they convert the binary string Ebar to an octet string E and then convert E to an integer e. We do not need to compute E. We can go straight from a bit string to a nat. You can see this in Words.bits_to_octets_to_nat, which shows that converting to a nat from the bits is the same as going to octets then to a nat. As an aside, this "e" is a "message digest". As noted above for p, because n is odd (n is an odd prime), and thus not a power of 2, \\log 2 n\ = bit_length n\. I feel that using \bit_length n\ is more understandable, and for Isabelle to compute these values for given test vectors I must use \bit_length n\.\ context SEC1 begin definition ECDSA_compute_e :: "hash_type \ octets \ nat" where "ECDSA_compute_e Hash M = ( let H = Hash M; Hbar = octets_to_bits_len H; x = bit_length n; Ebar = take x Hbar in bits_to_nat Ebar )" lemma ECDSA_e_bnd: assumes "e = ECDSA_compute_e Hash M" "octets_valid (Hash M)" shows "e < 2^(bit_length n)" proof - let ?H = "Hash M" let ?Hbar = "octets_to_bits_len ?H" let ?x = "bit_length n" let ?Ebar = "take ?x ?Hbar" have 1: "e = bits_to_nat ?Ebar" using assms(1) ECDSA_compute_e_def by presburger have 2: "length ?Ebar \ ?x" by fastforce have 3: "bits_valid ?Hbar" using assms(2) octets_valid_to_bits_len_valid by fast have 4: "bits_valid ?Ebar" using 3 words_valid_take by fast have 5: "e < 2^(length ?Ebar)" by (metis 1 4 words_to_nat_len_bnd words_valid_def less_numeral_extra(1) power_one_right) show ?thesis by (meson 2 5 le_trans linorder_not_less one_le_numeral power_increasing) qed lemma ECDSA_e_bitlen: assumes "e = ECDSA_compute_e Hash M" "octets_valid (Hash M)" shows "bit_length e \ bit_length n" using assms less_bit_len2 ECDSA_e_bnd by blast lemma ECDSA_e_bitlen': assumes "e = ECDSA_compute_e Hash M" "octets_valid (Hash M)" shows "bit_length e \ nat \log 2 n\" by (metis N.p_prime ECparamsValid n_not_2 ECDSA_e_bitlen bitLenLog2Prime assms nat_int) text \It could be that the hash function output has the same number of bits as the order n. For example, the NIST curve P-256 has a 256-bit n. Then if the hash used is SHA-256, the hash output always has 256 bits. In this case, the computation of e is simplified.\ lemma ECDSA_e_hlen_eq: assumes "e = ECDSA_compute_e Hash M" "octets_valid (Hash M)" "length (Hash M) = hlen" "bit_length n = 8*hlen" shows "e = octets_to_nat (Hash M)" by (metis ECDSA_compute_e_def assms octets_to_bits_len_len octets_to_bits_len_to_nat order.refl take_all) lemma ECDSA_e_hlen_eq': assumes "e = ECDSA_compute_e Hash M" "octets_valid (Hash M)" "length (Hash M) = hlen" "nat \log 2 n\ = 8*hlen" shows "e = octets_to_nat (Hash M)" by (metis ECparamsValid n_not_2 N.p_prime bitLenLog2Prime nat_int assms ECDSA_e_hlen_eq) text \Note that in the standard, the ephemeral key pair is denoted (k, R). However the letter R is used in this theory to refer to the ring of integers modulo the prime p. We need to use a different letter, so we use P instead of R. Also, the procedure for singing a message is to first compute the message digest e and then sign that digest. For ease of application, we clearly delineate these two steps.\ definition ECDSA_Sign_e :: "nat \ nat \ nat \ int point \ sig_type option" where "ECDSA_Sign_e d\<^sub>U e k P = ( case P of Infinity \ None | Point x\<^sub>P y\<^sub>P \ ( let r = nat (x\<^sub>P mod n); kinv = inv_mod_n k; s = (kinv*(e + r*d\<^sub>U)) mod n in if r = 0 then None else ( if s = 0 then None else Some (r,s) ) ) )" abbreviation ECDSA_Sign :: "hash_type \ nat \ octets \ nat \ int point \ sig_type option" where "ECDSA_Sign Hash d\<^sub>U M k P \ ECDSA_Sign_e d\<^sub>U (ECDSA_compute_e Hash M) k P" text \The standard says to pick a key pair (k,P) for signing. But as P is a function of k, we can abbreviate these definitions.\ abbreviation ECDSA_Sign_e' :: "nat \ nat \ nat \ sig_type option" where "ECDSA_Sign_e' d\<^sub>U e k \ ECDSA_Sign_e d\<^sub>U e k (ECkeyGen k)" abbreviation ECDSA_Sign' :: "hash_type \ nat \ octets \ nat \ sig_type option" where "ECDSA_Sign' Hash d\<^sub>U M k \ ECDSA_Sign Hash d\<^sub>U M k (ECkeyGen k)" text \Now a few facts about the signing operation. Some of these are helpful in the public key recovery operation which is defined in 4.1.6\ lemma ECDSA_Sign_e_Inf: assumes "P = Infinity" shows "ECDSA_Sign_e d\<^sub>U e k P = None" by (simp add: ECDSA_Sign_e_def assms) lemma ECDSA_Sign_Inf: assumes "P = Infinity" shows "ECDSA_Sign Hash d\<^sub>U M k P = None" using assms ECDSA_Sign_e_Inf ECDSA_Sign_e_def by simp lemma ECDSA_Sign_e_Some: assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" shows "P \ Infinity" using ECDSA_Sign_e_Inf assms by force lemma ECDSA_Sign_Some: assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" shows "P \ Infinity" using ECDSA_Sign_e_Some assms ECDSA_Sign_e_def by simp lemma ECDSA_Sign_e_SomeOut: assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" "P = Point x y" "r = nat (x mod n)" "kinv = inv_mod_n k" "s = (kinv*(e + r*d\<^sub>U)) mod n" shows "S = (r, s)" by (smt (z3) assms ECDSA_Sign_e_def option.distinct(1) option.inject point.simps(5)) lemma ECDSA_Sign_SomeOut: assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "P = Point x y" "r = nat (x mod n)" "kinv = inv_mod_n k" "s = (kinv*(e + r*d\<^sub>U)) mod n" "e = ECDSA_compute_e Hash M" shows "S = (r, s)" using ECDSA_Sign_e_SomeOut assms ECDSA_Sign_e_def by algebra lemma ECDSA_Sign_e_inRn: assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" "r = fst S" "s = snd S" shows "0 < r \ r < n \ r \ carrier Rn \ 0 < s \ s < n \ s \ carrier Rn" proof - obtain x y where P: "P = Point x y" by (metis assms(1) ECDSA_Sign_e_Some point.exhaust) let ?r = "nat (x mod n)" let ?kinv = "inv_mod_n k" let ?s = "(?kinv*(e + ?r*d\<^sub>U)) mod n" have r1: "0 < ?r" using assms(1) ECDSA_Sign_e_def P gr_zeroI by fastforce have r2: "?r < n" using nat_less_iff prime_gt_1_nat r1 by fastforce have r3: "?r \ carrier Rn" using r2 by fastforce have s1: "0 < ?s" using assms(1) ECDSA_Sign_e_def P gr_zeroI r1 by fastforce have s2: "?s < n" using nat_less_iff prime_gt_1_nat by fastforce have s3: "?s \ carrier Rn" by (metis N.mod_in_carrier of_nat_mod) have S1: "S = (?r, ?s)" using ECDSA_Sign_e_SomeOut P assms(1) by blast show ?thesis using S1 r1 r2 r3 s1 s2 s3 assms(2,3) by force qed lemma ECDSA_Sign_inRn: assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "r = fst S" "s = snd S" shows "0 < r \ r < n \ r \ carrier Rn \ 0 < s \ s < n \ s \ carrier Rn" using ECDSA_Sign_e_inRn assms ECDSA_Sign_e_def by algebra lemma ECDSA_Sign_e_invRn: assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" "r = fst S" "s = snd S" shows "inv_mod_n r = inv\<^bsub>Rn\<^esub>r \ r\\<^bsub>Rn\<^esub>inv\<^bsub>Rn\<^esub>r = \\<^bsub>Rn\<^esub> \ inv\<^bsub>Rn\<^esub>r \ carrier Rn \ inv_mod_n s = inv\<^bsub>Rn\<^esub>s \ s\\<^bsub>Rn\<^esub>inv\<^bsub>Rn\<^esub>s = \\<^bsub>Rn\<^esub> \ inv\<^bsub>Rn\<^esub>s \ carrier Rn" using ECDSA_Sign_e_inRn ECkeyPair_invRn' ECkeyPair_inv_d' ECprivateKeyValid_def assms by presburger lemma ECDSA_Sign_invRn: assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "r = fst S" "s = snd S" shows "inv_mod_n r = inv\<^bsub>Rn\<^esub>r \ r\\<^bsub>Rn\<^esub>inv\<^bsub>Rn\<^esub>r = \\<^bsub>Rn\<^esub> \ inv\<^bsub>Rn\<^esub>r \ carrier Rn \ inv_mod_n s = inv\<^bsub>Rn\<^esub>s \ s\\<^bsub>Rn\<^esub>inv\<^bsub>Rn\<^esub>s = \\<^bsub>Rn\<^esub> \ inv\<^bsub>Rn\<^esub>s \ carrier Rn" using ECDSA_Sign_e_invRn assms ECDSA_Sign_e_def by algebra lemma ECDSA_Sign_e_SomeOut_Curve: assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" "P = Point x y" "r = nat (x mod n)" "kinv = inv\<^bsub>Rn\<^esub> k" "c = (e + r*d\<^sub>U) mod n" "s = kinv \\<^bsub>Rn\<^esub> c" "ECkeyPairValid k P" shows "S = (r, s)" by (smt (verit) assms ECDSA_Sign_e_SomeOut ECDSA_Sign_e_invRn ECkeyPair_inv_d split_beta N.res_mult_eq fst_eqD mod_mult_right_eq of_nat_mod semiring_1_class.of_nat_mult snd_eqD) lemma ECDSA_Sign_e_SomeOut_Curve': assumes "ECDSA_Sign_e' d\<^sub>U e k = Some S" "P = Point x y" "r = nat (x mod n)" "kinv = inv\<^bsub>Rn\<^esub> k" "c = (e + r*d\<^sub>U) mod n" "s = kinv \\<^bsub>Rn\<^esub> c" "ECkeyGen k = P" "ECprivateKeyValid k" shows "S = (r, s)" using assms ECkeyPairEqKeyGen ECDSA_Sign_e_SomeOut_Curve by blast lemma ECDSA_Sign_SomeOut_Curve: assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "P = Point x y" "r = nat (x mod n)" "kinv = inv\<^bsub>Rn\<^esub> k" "c = (e + r*d\<^sub>U) mod n" "s = kinv \\<^bsub>Rn\<^esub> c" "ECkeyPairValid k P" "e = ECDSA_compute_e Hash M" shows "S = (r, s)" using ECDSA_Sign_e_SomeOut_Curve assms ECDSA_Sign_e_def by algebra lemma ECDSA_Sign_SomeOut_Curve': assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "P = Point x y" "r = nat (x mod n)" "kinv = inv\<^bsub>Rn\<^esub> k" "c = (e + r*d\<^sub>U) mod n" "s = kinv \\<^bsub>Rn\<^esub> c" "ECkeyGen k = P" "e = ECDSA_compute_e Hash M" "ECprivateKeyValid k" shows "S = (r, s)" using assms ECkeyPairEqKeyGen ECDSA_Sign_SomeOut_Curve by presburger lemma ECDSA_Sign_e_SomeOut_r: assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" "P = Point x y" "r = fst S" shows "r = nat (x mod n)" using assms ECDSA_Sign_e_SomeOut by fastforce lemma ECDSA_Sign_SomeOut_r: assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "P = Point x y" "r = fst S" shows "r = nat (x mod n)" using ECDSA_Sign_e_SomeOut_r assms ECDSA_Sign_e_def by blast lemma ECDSA_Sign_e_SomeOut_s: assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" "P = Point x y" "s = snd S" "r = nat (x mod n)" "kinv = inv_mod_n k" shows "s = (kinv*(e + r*d\<^sub>U)) mod n" using assms ECDSA_Sign_e_SomeOut by fastforce lemma ECDSA_Sign_SomeOut_s: assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "P = Point x y" "s = snd S" "r = nat (x mod n)" "kinv = inv_mod_n k" "e = ECDSA_compute_e Hash M" shows "s = (kinv*(e + r*d\<^sub>U)) mod n" using ECDSA_Sign_e_SomeOut_s assms ECDSA_Sign_e_def by blast lemma ECDSA_Sign_e_SomeOut_curve_s: fixes x :: int assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" "P = Point x y" "s = snd S" "r = nat (x mod n)" "kinv = inv\<^bsub>Rn\<^esub> k" "c = (e + r*d\<^sub>U) mod n" "ECkeyPairValid k P" shows "s = kinv \\<^bsub>Rn\<^esub> c" by (smt (verit) assms ECDSA_Sign_e_SomeOut_s ECkeyPair_inv_d N.mult_cong N.res_mult_eq mod_mod_trivial of_nat_mod semiring_1_class.of_nat_mult) lemma ECDSA_Sign_e_SomeOut_curve_s': fixes x :: int assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" "P = Point x y" "s = snd S" "r = nat (x mod n)" "kinv = inv\<^bsub>Rn\<^esub> k" "c = (e + r*d\<^sub>U) mod n" "ECkeyGen k = P" "ECprivateKeyValid k" shows "s = kinv \\<^bsub>Rn\<^esub> c" using assms ECkeyPairEqKeyGen ECDSA_Sign_e_SomeOut_curve_s by algebra lemma ECDSA_Sign_SomeOut_curve_s: fixes x :: int assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "P = Point x y" "s = snd S" "r = nat (x mod n)" "kinv = inv\<^bsub>Rn\<^esub> k" "c = (e + r*d\<^sub>U) mod n" "ECkeyPairValid k P" "e = ECDSA_compute_e Hash M" shows "s = kinv \\<^bsub>Rn\<^esub> c" using ECDSA_Sign_e_SomeOut_curve_s assms ECDSA_Sign_e_def by algebra lemma ECDSA_Sign_SomeOut_curve_s': fixes x :: int assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "P = Point x y" "s = snd S" "r = nat (x mod n)" "kinv = inv\<^bsub>Rn\<^esub> k" "c = (e + r*d\<^sub>U) mod n" "ECkeyGen k = P" "ECprivateKeyValid k" "e = ECDSA_compute_e Hash M" shows "s = kinv \\<^bsub>Rn\<^esub> c" using assms ECkeyPairEqKeyGen ECDSA_Sign_SomeOut_curve_s by algebra text \Now a few facts about signing with (k mod n) as the private ephemeral key versus k.\ lemma ECDSA_Sign_e_kmodn: "ECDSA_Sign_e d\<^sub>U e k P = ECDSA_Sign_e d\<^sub>U e (k mod n) P" using inv_mod_mod ECDSA_Sign_e_def by presburger lemma ECDSA_Sign_kmodn: "ECDSA_Sign Hash d\<^sub>U M k P = ECDSA_Sign Hash d\<^sub>U M (k mod n) P" using ECDSA_Sign_e_kmodn ECDSA_Sign_e_def by fast lemma ECDSA_Sign_e'_kmodn: "ECDSA_Sign_e' d\<^sub>U e k = ECDSA_Sign_e' d\<^sub>U e (k mod n)" using ECDSA_Sign_e_kmodn ECkeyGen_mod_n by presburger lemma ECDSA_Sign'_kmodn: "ECDSA_Sign' Hash d\<^sub>U M k = ECDSA_Sign' Hash d\<^sub>U M (k mod n)" using ECDSA_Sign_e'_kmodn ECDSA_Sign_e_def by fast subsubsection \4.1.4 Verifying Operation\ text \Again, the standard uses the letter R to represent a curve point. Here R is used to denote the ring of integers mod p. So use a different letter (P) to avoid confusion. As above, the way to verify that a signature S matches a message M is to first compute the message digest e and verify that S matches e. So we again delineate the two steps of calculating the digest and verifying S against e. Recall that \(d\<^sub>U,Q\<^sub>U)\ is the (long-term) key pair for the signer U.\ definition ECDSA_Verify_e :: "int point \ nat \ sig_type \ bool" where "ECDSA_Verify_e Q\<^sub>U e S = ( let r = fst S; s = snd S; sinv = inv_mod_n s; u1 = (e*sinv) mod n; u2 = (r*sinv) mod n; P = add' (point_mult' u1 G) (point_mult' u2 Q\<^sub>U) in (case P of Infinity \ False | Point x y \ ( (0 < r) \ (r < n) \ (0 < s) \ (s < n) \ (nat (x mod n) = r) ) ) )" abbreviation ECDSA_Verify :: "hash_type \ int point \ octets \ sig_type \ bool" where "ECDSA_Verify Hash Q\<^sub>U M S \ ECDSA_Verify_e Q\<^sub>U (ECDSA_compute_e Hash M) S" text \If M is signed with a valid ephemeral key pair (k, P), then the resulting signature S will pass the verification routine for M.\ lemma ECDSA_Sign_e_Verify: assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" "ECkeyPairValid k P" "ECkeyPairValid d\<^sub>U Q\<^sub>U" shows "ECDSA_Verify_e Q\<^sub>U e S" proof - obtain x y where P: "P = Point x y" by (metis assms(1) ECDSA_Sign_e_Some point.exhaust) let ?r = "nat (x mod n)" let ?kinv = "inv_mod_n k" let ?s = "(?kinv*(e + ?r*d\<^sub>U)) mod n" have S1: "S = (?r, ?s)" using ECDSA_Sign_e_SomeOut P assms(1) by blast have S2: "fst S = ?r \ snd S = ?s" using S1 by force have r1: "0 < ?r" by (metis assms(1) S2 ECDSA_Sign_e_inRn) have r2: "?r < n" by (metis assms(1) S2 ECDSA_Sign_e_inRn) have s1: "0 < ?s" by (metis assms(1) S2 ECDSA_Sign_e_inRn) have s2: "?s < n" by (metis assms(1) S2 ECDSA_Sign_e_inRn) let ?x = "(e + ?r*d\<^sub>U) mod n" have x0: "?s = ?kinv \\<^bsub>Rn\<^esub> ?x" by (metis ECDSA_Sign_e_SomeOut_curve_s ECkeyPair_inv_d P S2 assms(1,2)) have x1: "0 < ?x" by (metis s1 mod_mod_trivial mod_mult_right_eq mult_0_right not_gr0) have x2: "?x \ carrier Rn" by (metis N.mod_in_carrier of_nat_mod) have k1: "k \ carrier Rn" using assms(2) ECkeyPair_dInRn by fast have k2: "k \ 0" using ECkeyPairValid_def assms(2) by blast have k3: "?kinv = inv\<^bsub>Rn\<^esub>k" using ECkeyPair_inv_d assms(2) by auto have k4: "inv\<^bsub>Rn\<^esub>?kinv = k" using N.nonzero_inverse_inverse_eq N.zero_cong k1 k2 k3 by force let ?sinv = "inv_mod_n ?s" have s3: "?sinv = inv\<^bsub>Rn\<^esub>?s" by (metis s2 s1 inv_mod_n_inv') have s4: "?sinv = inv\<^bsub>Rn\<^esub>?kinv \\<^bsub>Rn\<^esub> inv\<^bsub>Rn\<^esub>?x" using N.m_comm N.nonzero_imp_inverse_nonzero N.nonzero_inverse_mult_distrib N.zero_cong k1 k2 k3 s3 x0 x1 x2 by auto have s5: "?sinv = k \\<^bsub>Rn\<^esub> inv\<^bsub>Rn\<^esub>?x" using s4 k4 by presburger let ?u1 = " (e*?sinv) mod n" let ?u2 = "(?r*?sinv) mod n" let ?P = "add' (point_mult' ?u1 G) (point_mult' ?u2 Q\<^sub>U)" have kP2: "G [^]\<^bsub>EPF.curve\<^esub>k = P" using assms(2) ECkeyPair_curveMult by fast have P1: "?P = G [^]\<^bsub>EPF.curve\<^esub>?u1 \\<^bsub>EPF.curve\<^esub> Q\<^sub>U [^]\<^bsub>EPF.curve\<^esub>?u2" using ECkeyPairOnCurve EPF.add_curve EPF.in_carrier_curve GonEPFcurve assms(3) EPF.multEqPowInCurve by presburger have P2: "Q\<^sub>U [^]\<^bsub>EPF.curve\<^esub>?u2 = G [^]\<^bsub>EPF.curve\<^esub>(d\<^sub>U*?u2)" using ECkeyPair_curveMult EPF.curve.nat_pow_pow GonEPFcurve assms(3) by blast have P3: "?P = G [^]\<^bsub>EPF.curve\<^esub>(?u1 + d\<^sub>U*?u2)" using P1 P2 EPF.curve.nat_pow_mult GonEPFcurve by presburger have k5: "(?u1 + d\<^sub>U*?u2) mod n = ?x*?sinv mod n" by (smt (z3) add_mult_distrib mod_add_cong mod_add_left_eq mod_mult_right_eq mult.assoc mult.commute) have k6: "?x*?sinv mod n = ?x\\<^bsub>Rn\<^esub>?sinv" by (simp add: N.res_mult_eq of_nat_mod) have k7: "?x\\<^bsub>Rn\<^esub>?sinv = ?x \\<^bsub>Rn\<^esub> k \\<^bsub>Rn\<^esub> inv\<^bsub>Rn\<^esub>?x" by (metis N.inv_closed N.m_assoc N.zero_cong gr_implies_not0 k1 of_nat_eq_0_iff x1 x2 k6 s5) have k8: "?x \\<^bsub>Rn\<^esub> k \\<^bsub>Rn\<^esub> inv\<^bsub>Rn\<^esub>?x = k" by (metis N.inv_closed N.m_lcomm N.r_inv N.r_one N.zero_cong k1 k7 less_not_refl2 of_nat_eq_0_iff s5 x1 x2) have k9: "(?u1 + d\<^sub>U*?u2) mod n = k" by (metis k5 k6 k7 k8 nat_int) have P4: "?P = P" by (metis P3 k9 kP2 multGmodn') show ?thesis using P4 S2 P r1 r2 s1 s2 ECDSA_Verify_e_def by auto qed lemma ECDSA_Sign_e_Verify': assumes "ECDSA_Sign_e' d\<^sub>U e k = Some S" "ECprivateKeyValid k" "ECkeyPairValid d\<^sub>U Q\<^sub>U" shows "ECDSA_Verify_e Q\<^sub>U e S" using assms ECDSA_Sign_e_Verify ECkeyPairEqKeyGen by blast lemma ECDSA_Sign_Verify: assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "ECkeyPairValid k P" "ECkeyPairValid d\<^sub>U Q\<^sub>U" shows "ECDSA_Verify Hash Q\<^sub>U M S" using ECDSA_Sign_e_def ECDSA_Sign_e_Verify ECDSA_Verify_e_def assms by fast lemma ECDSA_Sign_Verify': assumes "ECDSA_Sign' Hash d\<^sub>U M k = Some S" "ECprivateKeyValid k" "ECkeyPairValid d\<^sub>U Q\<^sub>U" shows "ECDSA_Verify Hash Q\<^sub>U M S" using assms ECDSA_Sign_Verify ECkeyPairEqKeyGen by blast subsubsection \4.1.5 Alternative Verifying Operation\ text \If the verifier knows U's private key \d\<^sub>U\, they can use this alternate form of the verification routine. In short, if the verifier knows \d\<^sub>U\, they can compute the (private) ephemeral signing k labeled k above. Note that this alternate verifying operation needs only one elliptic curve operation, as opposed to the two needed above.\ definition ECDSA_Verify_e_Alt :: "nat \ nat \ sig_type \ bool" where "ECDSA_Verify_e_Alt d\<^sub>U e S = ( let r = fst S; s = snd S; sinv = inv_mod_n s; u1 = nat ((e*sinv) mod n); u2 = nat ((r*sinv) mod n); P = point_mult' (u1 + u2*d\<^sub>U) G in (case P of Infinity \ False | Point x y \ ( (0 < r) \ (r < n) \ (0 < s) \ (s < n) \ (nat (x mod n) = r) ) ) )" definition ECDSA_Verify_Alt :: "hash_type \ nat \ octets \ sig_type \ bool" where "ECDSA_Verify_Alt Hash d\<^sub>U M S = ECDSA_Verify_e_Alt d\<^sub>U (ECDSA_compute_e Hash M) S" lemma ECDSA_Verify_e_eq: assumes "ECkeyPairValid d\<^sub>U Q\<^sub>U" shows "ECDSA_Verify_e_Alt d\<^sub>U e S = ECDSA_Verify_e Q\<^sub>U e S" by (smt (z3) ECDSA_Verify_e_Alt_def ECDSA_Verify_e_def ECkeyPairOnCurve ECkeyPairValid_def EPF.add_curve EPF.curve.nat_pow_mult EPF.curve.nat_pow_pow EPF.in_carrier_curve GonEPFcurve assms(1) mult.commute EPF.multEqPowInCurve nat_int) lemma ECDSA_Verify_eq: assumes "ECkeyPairValid d\<^sub>U Q\<^sub>U" shows "ECDSA_Verify_Alt Hash d\<^sub>U M S = ECDSA_Verify Hash Q\<^sub>U M S" using ECDSA_Verify_Alt_def ECDSA_Verify_e_eq assms by presburger text \If you can recover the ephemeral signing key (because you know the long-term secret key) from a paired message (digest) and verified signature, then if you sign that message with that recovered signing key, you will get back that signature. This is the inverse of the lemma above which says that if you sign a message and get S, then S will be verified as a correct signature for the message.\ definition ECDSA_Verify_Alt_recover_k :: "nat \ nat \ sig_type \ nat" where "ECDSA_Verify_Alt_recover_k d\<^sub>U e S = ( let r = fst S; s = snd S; sinv = inv_mod_n s; u1 = nat ((e*sinv) mod n); u2 = nat ((r*sinv) mod n) in (u1 + u2*d\<^sub>U) mod n )" lemma recovered_k_less_n: "ECDSA_Verify_Alt_recover_k d\<^sub>U e S < n" by (metis ECDSA_Verify_Alt_recover_k_def N.p_prime mod_less_divisor prime_gt_0_nat) lemma ECDSA_Verify_e_Sign: assumes "ECDSA_Verify_e Q\<^sub>U e S" "P = point_mult' k G" "ECkeyPairValid d\<^sub>U Q\<^sub>U" "k = ECDSA_Verify_Alt_recover_k d\<^sub>U e S" shows "ECDSA_Sign_e d\<^sub>U e k P = Some S" proof - have 1: "ECDSA_Verify_e_Alt d\<^sub>U e S" using assms(1,3) ECDSA_Verify_e_eq by fast let ?r = "fst S" let ?s = "snd S" let ?sinv = "inv_mod_n ?s" let ?u1 = "nat (( e*?sinv) mod n)" let ?u2 = "nat ((?r*?sinv) mod n)" let ?P = "point_mult' (?u1 + ?u2*d\<^sub>U) G" have P1: "?P = point_mult' ((?u1 + ?u2*d\<^sub>U) mod n) G" using multGmodn by presburger have P2: "?P = point_mult' k G" by (metis P1 assms(4) ECDSA_Verify_Alt_recover_k_def) have P3: "?P = P" using assms(2) ECkeyPairValid_def P2 by algebra have P4: "?P \ Infinity" using 1 ECDSA_Verify_e_Alt_def by fastforce obtain x and y where P5: "?P = Point x y" using P4 point.exhaust by blast have 2: "(0 < ?r) \ (?r < n) \ (0 < ?s) \ (?s < n) \ (nat (x mod n) = ?r)" by (smt (verit) 1 ECDSA_Verify_e_Alt_def P5 point.case(2)) have s1: "?sinv = inv\<^bsub>Rn\<^esub>?s" using 2 inv_mod_n_inv' by fastforce have s2: "inv\<^bsub>Rn\<^esub>?sinv = ?s" by (simp add: 2 s1 ECkeyPair_dInRn' ECprivateKeyValid_def N.nonzero_inverse_inverse_eq N.zero_cong) have k1: "0 < k" by (metis P2 P4 assms(2) gr0I point_mult.simps(1)) have k2: "k < n" using assms(4) recovered_k_less_n by simp let ?c = "(e + ?r*d\<^sub>U) mod n" have k3: "k = ?c*?sinv mod n" by (metis (no_types, lifting) ECDSA_Verify_Alt_recover_k_def assms(4) distrib_right mod_add_left_eq mod_add_right_eq mod_mult_left_eq mult.assoc mult.commute nat_int) have c1: "0 < ?c \ ?c < n" by (metis bot_nat_0.not_eq_extremum k1 k2 k3 less_nat_zero_code mod_0 mod_less_divisor mult_0) let ?kinv = "inv_mod_n k" have k4: "?kinv = inv\<^bsub>Rn\<^esub> k" using k1 k2 inv_mod_n_inv' by fastforce have k5: "k = ?c \\<^bsub>Rn\<^esub> inv\<^bsub>Rn\<^esub>?s" by (simp add: N.res_mult_eq k3 of_nat_mod s1) have k6: "inv\<^bsub>Rn\<^esub> k = (inv\<^bsub>Rn\<^esub>?c) \\<^bsub>Rn\<^esub> ?s" by (metis 2 k5 s1 s2 c1 ECkeyGen_valid_mod_n ECkeyPair_dInRn' ECkeyPair_invRn' N.m_comm N.nonzero_imp_inverse_nonzero N.nonzero_inverse_mult_distrib N.zero_cong bot_nat_0.not_eq_extremum mod_if of_nat_eq_0_iff) let ?s' = "(?kinv*(e + ?r*d\<^sub>U)) mod n" have s3: "?s' = (inv\<^bsub>Rn\<^esub> k) \\<^bsub>Rn\<^esub> ?c" by (simp add: N.res_mult_eq k4 mod_mult_right_eq of_nat_mod) have s4: "?s' = ((inv\<^bsub>Rn\<^esub>?c) \\<^bsub>Rn\<^esub> ?c) \\<^bsub>Rn\<^esub> ?s" by (metis (no_types, lifting) 2 k4 k6 s3 ECkeyPair_invRn' ECprivateKeyValid_def ECkeyPair_dInRn' N.cring_simprules(11) N.res_mult_eq c1 mult.commute) have s5: "?s = ?s'" by (metis 2 s4 ECkeyGen_valid_mod_n ECkeyPair_dInRn' ECkeyPair_invRn' EPF.nat_int_eq N.m_closed N.m_comm N.r_one c1 mod_if nat_neq_iff) have A1: "ECDSA_Sign_e d\<^sub>U e k P = ( let r = nat (x mod n); kinv = inv_mod_n k; s = (kinv*(e + r*d\<^sub>U)) mod n in if r = 0 then None else (if s = 0 then None else Some (r,s)))" using ECDSA_Sign_e_def P3 P4 P5 by force have A2: "ECDSA_Sign_e d\<^sub>U e k P = (if ?r = 0 then None else (if ?s = 0 then None else Some (?r,?s)))" by (smt (verit) 2 A1 s5) show ?thesis using A2 2 by auto qed lemma ECDSA_Verify_Sign: assumes "ECDSA_Verify Hash Q\<^sub>U M S" "P = point_mult' k G" "ECkeyPairValid d\<^sub>U Q\<^sub>U" "k = ECDSA_Verify_Alt_recover_k d\<^sub>U (ECDSA_compute_e Hash M) S" shows "ECDSA_Sign Hash d\<^sub>U M k P = Some S" by (metis ECDSA_Verify_e_Sign assms) text \One last variation on this theme. Suppose that the cofactor h = 1. Then you know that every point on the curve is a power of the generator G. So if you know that \Q\<^sub>U\ is a partially valid public key, then we know that there is a corresponding private key \d\<^sub>U\. If you could find that private key (which, in practice, you can't), then you can use it to recover the ephemeral signing key.\ lemma ECDSA_Verify_Sign_h1: assumes "ECpublicKeyPartialValid Q\<^sub>U" "ECDSA_Verify_e Q\<^sub>U e S" "h = 1" shows "\d\<^sub>U. (Q\<^sub>U = point_mult' d\<^sub>U G \ ECDSA_Sign_e' d\<^sub>U e (ECDSA_Verify_Alt_recover_k d\<^sub>U e S) = Some S)" proof - have 1: "on_curve' Q\<^sub>U" using assms(1) ECpublicKeyPartialValid_def by blast obtain d\<^sub>U where 2: "d\<^sub>U Q\<^sub>U = point_mult' d\<^sub>U G" using assms(3) 1 EC_h1_cyclic' by blast have 3: "Q\<^sub>U \ Infinity" using assms(1) ECpublicKeyPartialValid_def by blast have 4: "0 < d\<^sub>U" by (metis 2 3 not_gr0 point_mult.simps(1)) have 5: "ECkeyPairValid d\<^sub>U Q\<^sub>U" using 2 4 ECkeyPairValid_def by blast have 6: "ECDSA_Sign_e' d\<^sub>U e (ECDSA_Verify_Alt_recover_k d\<^sub>U e S) = Some S" using 5 ECDSA_Verify_e_Sign ECkeyGen_def assms(2) by blast show ?thesis using 2 6 by fast qed subsubsection \4.1.6 Public Key Recovery Operation\ text \The idea is that if you have a signature, you can generate a small list of putative values for the public key. Then if you have an oracle, you can test each of these putative values. This is a convenient option in bandwidth-constrained environments. What might this oracle be? For example, if you have a second message (M2) and the signature for that message (S2), then you can check that (ECDSA_Verify Hash Q M2 S2) returns True when the putative public key Q is used. We don't have access to the oracle needed in step 1.6.2 of the public key recovery operation. So here we simply call it the ValidationOracle, which takes a point and returns True or False. Note that the public key recovery operation loops over j. In HOL/Isabelle, we translate the loop as a recursive function ECDSA_PublicKeyRecovery_rec. The "meat" of each iteration is contained in the function ECDSA_PublicKeyRecovery_1. ECDSA_PublicKeyRecovery first computes e (because it is the same in every loop iteration) and kicks off the loop by calling ECDSA_PublicKeyRecovery_rec. Lastly, we note again that the standard writes \(log 2 p)/8\. We know because p is odd (and thus not a power of two) that octet_length p = \(log 2 p)/8\. The proof of this may be found in Words.thy/octetLenLog2Prime.\ definition ECDSA_PublicKeyRecovery_1 :: "(int point \ bool) \ nat \ sig_type \ nat \ int point option" where "ECDSA_PublicKeyRecovery_1 ValidationOracle e Sig j = ( let r = fst Sig; s = snd Sig; x = r + j*n; mlen = octet_length p; X = nat_to_octets_len x mlen; P = octets_to_point' (2 # X) in ( case P of None \ None | Some R' \ ( let nR = point_mult' n R'; rInv = inv_mod_n r; eG = opp (point_mult' e G); sR1 = point_mult' s R'; sR2 = point_mult' s (opp R'); Q1 = point_mult' rInv (add' sR1 eG); Q2 = point_mult' rInv (add' sR2 eG) in if (nR = Infinity) then ( if (ValidationOracle Q1) then (Some Q1) else if (ValidationOracle Q2) then (Some Q2) else None ) else None ) ) )" text \It is a little more convenient to have a loop counter that decreases until it reaches 0. So the loop index i here satisfies j = h + 1 - i, where j is the loop counter from step 1 in the standard. i starts at h+1, which corresponds to j = 0. When i = 1, j = h, so when i = 0, there are no more iterations to perform. If the signature is valid, the public key should be recovered so that i = 0 is never reached.\ fun ECDSA_PublicKeyRecovery_rec :: "(int point \ bool) \ nat \ sig_type \ nat \ int point option" where "ECDSA_PublicKeyRecovery_rec ValidationOracle e Sig 0 = None" | "ECDSA_PublicKeyRecovery_rec ValidationOracle e Sig i = ( let P = ECDSA_PublicKeyRecovery_1 ValidationOracle e Sig (h+1-i) in ( if (P = None) then (ECDSA_PublicKeyRecovery_rec ValidationOracle e Sig (i-1)) else P ) )" definition ECDSA_PublicKeyRecovery_e :: "(int point \ bool) \ nat \ sig_type \ int point option" where "ECDSA_PublicKeyRecovery_e ValidationOracle e S = ECDSA_PublicKeyRecovery_rec ValidationOracle e S (h+1)" definition ECDSA_PublicKeyRecovery :: "hash_type \ (int point \ bool) \ octets \ sig_type \ int point option" where "ECDSA_PublicKeyRecovery Hash ValidationOracle M S = ECDSA_PublicKeyRecovery_e ValidationOracle (ECDSA_compute_e Hash M) S" text \Knowing nothing about the hash function or the oracle, at the very least we know that if the key recovery operation returns a point, then that point is on the elliptic curve.\ lemma ECDSA_PublicKeyRecovery_1_onCurve: assumes "ECDSA_PublicKeyRecovery_1 ValidationOracle e Sig j = Some Q" shows "on_curve' Q" proof - let ?r = "fst Sig" let ?s = "snd Sig" let ?x = "?r + j*n" let ?mlen = "octet_length p" let ?X = "nat_to_octets_len ?x ?mlen" let ?P = "octets_to_point' (2 # ?X)" have 1: "?P \ None" by (metis (no_types, lifting) ECDSA_PublicKeyRecovery_1_def assms option.case(1) option.distinct(1)) obtain R' where 2: "?P = Some R'" using 1 by blast have 3: "on_curve' R'" using 2 octets2PointOnCurve by fast let ?rInv = "inv_mod_n ?r" let ?eG = "opp (point_mult' e G)" let ?sR1 = "point_mult' ?s R'" let ?sR2 = "point_mult' ?s (opp R')" let ?Q1 = "point_mult' ?rInv (add' ?sR1 ?eG)" let ?Q2 = "point_mult' ?rInv (add' ?sR2 ?eG)" have 4: "Q = ?Q1 \ Q = ?Q2" by (smt (verit, best) 2 assms ECDSA_PublicKeyRecovery_1_def not_None_eq not_Some_eq option.distinct(1) option.inject option.simps(5)) have 5: "on_curve' ?eG \ on_curve' ?sR1 \ on_curve' ?sR2" using 3 ECparamsValid ECdomainParametersValid_def opp_closed point_mult_closed by simp have 6: "on_curve' ?Q1 \ on_curve' ?Q2" using 5 point_mult_closed add_closed by auto show ?thesis using 4 6 by fast qed lemma ECDSA_PublicKeyRecovery_rec_onCurve: assumes "ECDSA_PublicKeyRecovery_rec ValidationOracle e Sig i = Some Q" shows "on_curve' Q" using assms apply (induct i) apply simp using ECDSA_PublicKeyRecovery_1_onCurve ECDSA_PublicKeyRecovery_rec.simps(2) diff_Suc_1 by (smt (verit, best)) text \The public key recovery operation is guaranteed to return the correct public key \Q\<^sub>U\, given a correct validation oracle. In the next several lemmas, we work toward proving this. In the ECDSA locale, U has signed a message with its key pair. So the correct oracle will only return "true" when the putative public key matches U's public key.\ definition UOracle :: "int point \ int point \ bool" where "UOracle Q\<^sub>U possibleQ = (Q\<^sub>U = possibleQ)" lemma OracleCorrect_1: assumes "ECDSA_PublicKeyRecovery_1 (UOracle Q\<^sub>U) e Sig j = Some P" shows "P = Q\<^sub>U" proof - let ?r = "fst Sig" let ?s = "snd Sig" let ?x = "?r + j*n" let ?mlen = "octet_length p" let ?X = "nat_to_octets_len ?x ?mlen" let ?P = "octets_to_point' (2 # ?X)" have 1: "?P \ None" by (metis (no_types, lifting) ECDSA_PublicKeyRecovery_1_def assms option.case(1) option.distinct(1)) obtain R' where 2: "?P = Some R'" using 1 by blast show ?thesis by (smt (verit) 2 ECDSA_PublicKeyRecovery_1_def UOracle_def assms option.distinct(1) option.inject option.simps(5)) qed lemma OracleCorrect_rec: assumes "ECDSA_PublicKeyRecovery_rec (UOracle Q\<^sub>U) e Sig i = Some P" shows "P = Q\<^sub>U" using assms apply (induct i) apply simp by (smt (verit, del_insts) ECDSA_PublicKeyRecovery_rec.simps(2) OracleCorrect_1 diff_Suc_1) lemma KeyRecoveryWorks_1: assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" "ECkeyPairValid k P" "P = Point x y" "j = (nat x) div n" "ECkeyPairValid d\<^sub>U Q\<^sub>U" shows "ECDSA_PublicKeyRecovery_1 (UOracle Q\<^sub>U) e S j = Some Q\<^sub>U" proof - have x0: "0 \ x \ x < p \ x \ carrier R" using ECkeyPairOnCurve assms(2,3) inCarrier onCurveEq2 by blast have j0: "0 \ x div n" using x0 div_int_pos_iff of_nat_0_le_iff by blast have j1: "j = nat (x div n)" using assms(4) nat_div_distrib nat_int x0 by presburger let ?r = "fst S" let ?s = "snd S" have r1: "?r = nat (x mod n)" by (metis ECDSA_Sign_e_SomeOut_r assms(1,3)) have r2: "?r \ carrier Rn" using ECDSA_Sign_e_inRn assms(1) by blast have r3: "0 < ?r" using ECDSA_Sign_e_inRn assms(1) by blast let ?x = "?r + j*n" let ?kinv = "inv\<^bsub>Rn\<^esub> k" have k0: "0 < k" using assms(2) ECkeyPairValid_def by simp have k1: "k \ carrier Rn" using assms(2) ECkeyPair_dInRn by blast have k2: "?kinv \ carrier Rn" by (metis k1 k0 N.inv_closed N.zero_cong bot_nat_0.not_eq_extremum dual_order.refl of_nat_le_0_iff) have k3: "k \\<^bsub>Rn\<^esub> ?kinv = \\<^bsub>Rn\<^esub>" using ECkeyPairValid_def N.zero_cong assms(2) k1 by force let ?c = "(e + ?r*d\<^sub>U) mod n" have c1: "?c \ carrier Rn" by (simp add: of_nat_mod) have s1: "?s = ?kinv \\<^bsub>Rn\<^esub> ?c" using ECDSA_Sign_e_SomeOut_curve_s assms(1,2,3,5) r1 by blast have x1: "?x = x" using j0 j1 r1 r3 by auto let ?l = "octet_length p" let ?X = "nat_to_octets_len ?x ?l" let ?P = "octets_to_point' (2 # ?X)" have P1: "on_curve' P" using assms(2) ECkeyPairOnCurve by blast have P2: "?P = Some P \ ?P = Some (opp P)" by (metis P1 assms(3) x1 point2Octets2PointComp_PoppP' nat_int) have P3: "point_mult' n P = Infinity" using assms(2) by (simp add: ECkeyPairOrd_n) have P4: "P \ carrier EPF.curve" by (simp add: EPF.in_carrier_curve P1) have P5: "opp P = inv\<^bsub>EPF.curve\<^esub> P" using EPF.inv_curve P4 by presburger have P6: "point_mult' n (opp P) = Infinity" by (metis EPF.curve.inv_one EPF.curve.nat_pow_inv EPF.one_curve P1 P3 P4 P5 EPF.multEqPowInCurve opp_closed) have P7: "point_mult' k G = P" using assms(2) ECkeyPairValid_def by blast have P8: "G[^]\<^bsub>EPF.curve\<^esub>k = P" using P7 EPF.in_carrier_curve GonEPFcurve EPF.multEqPowInCurve by force obtain R' where R1: "?P = Some R'" using P2 by blast have R2: "R' = P \ R' = opp P" using P2 R1 by auto have R3: "(R' = opp P) = (opp R' = P)" using P1 R2 opp_opp by auto let ?nR = "point_mult' n R'" have nR1: "?nR = point_mult' n P \ ?nR = point_mult' n (opp P)" using R2 by fast have nR2: "?nR = Infinity" using nR1 P3 P6 by argo let ?rInv = "inv_mod_n ?r" have rInv: "?rInv = inv\<^bsub>Rn\<^esub> ?r" using ECDSA_Sign_e_inRn assms(1) bot_nat_0.not_eq_extremum inv_mod_n_inv by presburger have rInv2: "?r \\<^bsub>Rn\<^esub> ?rInv = \\<^bsub>Rn\<^esub>" by (simp add: N.zero_cong r2 r3 rInv) let ?eG = "opp (point_mult' e G)" have eG1: "?eG = G[^]\<^bsub>EPF.curve\<^esub>(-(int e))" using EPF.curve.int_pow_neg_int EPF.in_carrier_curve EPF.inv_curve GonEPFcurve dGonEPFcurve EPF.multEqPowInCurve by auto have eG2: "?eG = G[^]\<^bsub>EPF.curve\<^esub>(-(int e) mod n)" using eG1 multGmodn'int by presburger let ?em = "(-(int e) mod n)" have em1: "?em = \\<^bsub>Rn\<^esub>e" using N.res_neg_eq by presburger let ?sR1 = "point_mult' ?s R'" let ?sR2 = "point_mult' ?s (opp R')" have sR1: "{?sR1, ?sR2} = {point_mult' ?s P, point_mult' ?s (opp P)}" using R2 R3 by fastforce have sR2: "{?sR1, ?sR2} = {P[^]\<^bsub>EPF.curve\<^esub>?s, P[^]\<^bsub>EPF.curve\<^esub>(-?s)}" by (metis sR1 P1 P4 P5 EPF.curve.int_pow_neg_int EPF.curve.nat_pow_inv opp_closed EPF.multEqPowInCurve) have sR3: "(R' = opp P) \ (?sR1 = P[^]\<^bsub>EPF.curve\<^esub>(-?s) \ ?sR2 = P[^]\<^bsub>EPF.curve\<^esub>?s)" by (metis P1 doubleton_eq_iff EPF.multEqPowInCurve sR2 opp_opp) let ?Q1 = "point_mult' ?rInv (add' ?sR1 ?eG)" let ?Q2 = "point_mult' ?rInv (add' ?sR2 ?eG)" let ?Q = "((P[^]\<^bsub>EPF.curve\<^esub>?s) \\<^bsub>EPF.curve\<^esub> (G [^]\<^bsub>EPF.curve\<^esub>?em))[^]\<^bsub>EPF.curve\<^esub>?rInv" have Q1: "(R' = P) \ ?Q1 = ?Q" using P4 eG2 EPF.add_curve GonEPFcurve EPF.curve.int_pow_closed EPF.curve.m_closed EPF.curve.nat_pow_closed EPF.in_carrier_curve EPF.multEqPowInCurve by auto have Q2: "(R' = opp P) \ ?Q2 = ?Q" using P4 sR3 eG2 GonEPFcurve EPF.add_curve EPF.curve.int_pow_closed EPF.curve.m_closed EPF.curve.nat_pow_closed EPF.in_carrier_curve EPF.multEqPowInCurve by auto have Q3: "P[^]\<^bsub>EPF.curve\<^esub>?s = G[^]\<^bsub>EPF.curve\<^esub>(k*?s)" using P8 EPF.curve.nat_pow_pow GonEPFcurve by blast have Q4: "(G[^]\<^bsub>EPF.curve\<^esub>(k*?s) \\<^bsub>EPF.curve\<^esub> (G [^]\<^bsub>EPF.curve\<^esub>?em)) = G[^]\<^bsub>EPF.curve\<^esub>(k*?s+?em)" by (metis EPF.curve.is_group GonEPFcurve group.int_pow_mult int_pow_int) have Q5: "?Q = G[^]\<^bsub>EPF.curve\<^esub>((k*?s+?em)*?rInv)" by (metis Q3 Q4 EPF.curve.int_pow_pow GonEPFcurve int_pow_int) have Q6: "?Q = G[^]\<^bsub>EPF.curve\<^esub>((k*?s+?em)*?rInv mod n)" using Q5 multGmodn'int by algebra have ex1: "(k*?s+?em)*?rInv mod n = (k\\<^bsub>Rn\<^esub>?s \\<^bsub>Rn\<^esub> ?em)\\<^bsub>Rn\<^esub>?rInv" by (simp add: N.res_add_eq N.res_mult_eq mod_add_left_eq mod_mult_right_eq mult.commute) have ex2: "(k*?s+?em)*?rInv mod n = (k\\<^bsub>Rn\<^esub>(?kinv \\<^bsub>Rn\<^esub> ?c) \\<^bsub>Rn\<^esub>e)\\<^bsub>Rn\<^esub>(inv\<^bsub>Rn\<^esub> ?r)" using ex1 s1 em1 rInv by algebra have ex3: "k\\<^bsub>Rn\<^esub>(?kinv \\<^bsub>Rn\<^esub> ?c) = k\\<^bsub>Rn\<^esub>?kinv \\<^bsub>Rn\<^esub> ?c" using m_assoc k1 k2 c1 by algebra have ex4: "k\\<^bsub>Rn\<^esub>(?kinv \\<^bsub>Rn\<^esub> ?c) = ?c" using ex3 k3 N.l_one c1 by presburger have ex5: "?c \\<^bsub>Rn\<^esub>e = e \\<^bsub>Rn\<^esub> (?r\\<^bsub>Rn\<^esub>d\<^sub>U) \\<^bsub>Rn\<^esub>e" by (simp add: N.res_add_eq N.res_mult_eq mod_add_right_eq of_nat_mod) have ex6: "?c \\<^bsub>Rn\<^esub>e = ?r\\<^bsub>Rn\<^esub>d\<^sub>U" using ex5 by (simp add: N.res_add_eq N.res_diff_eq N.res_mult_eq mod_diff_left_eq) have ex7: "(k*?s+?em)*?rInv mod n = ?r\\<^bsub>Rn\<^esub>d\<^sub>U\\<^bsub>Rn\<^esub>?rInv" using ex2 ex4 ex6 rInv by argo have ex8: "(k*?s+?em)*?rInv mod n = d\<^sub>U mod n" by (metis (mono_tags, opaque_lifting) ex7 rInv2 N.res_mult_eq N.res_one_eq mod_mult_right_eq mult.left_commute mult.right_neutral of_nat_mod semiring_1_class.of_nat_mult) have ex9: "(k*?s+?em)*?rInv mod n = d\<^sub>U" using ex8 assms(5) ECkeyPairValid_def by force have Q7: "?Q = G[^]\<^bsub>EPF.curve\<^esub>d\<^sub>U" by (metis Q6 ex9 int_pow_int) have Q8: "?Q = Q\<^sub>U" using Q7 ECkeyPair_curveMult assms(5) by presburger have A1: "(R' = P) \ ?Q1 = Q\<^sub>U" using Q8 Q1 by argo have A11: "(R' = P) \ UOracle Q\<^sub>U ?Q1" using A1 UOracle_def by blast have A12: "(R' = P) \ ECDSA_PublicKeyRecovery_1 (UOracle Q\<^sub>U) e S j = Some Q\<^sub>U" using A11 ECDSA_PublicKeyRecovery_1_def P2 nR2 A1 R1 by auto have A2: "(R' = opp P) \ ?Q2 = Q\<^sub>U" using Q8 Q2 by argo have A21: "(R' = opp P) \ UOracle Q\<^sub>U ?Q2" using A2 UOracle_def by blast have A22: "(R' = opp P) \ ECDSA_PublicKeyRecovery_1 (UOracle Q\<^sub>U) e S j = Some Q\<^sub>U" by (smt (verit) nR2 A21 R1 UOracle_def ECDSA_PublicKeyRecovery_1_def option.simps(5)) show ?thesis using A12 A22 R2 by fast qed lemma KeyRecoveryWorks_1': assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "ECkeyPairValid k P" "P = Point x y" "j = (nat x) div n" "ECkeyPairValid d\<^sub>U Q\<^sub>U" "e = ECDSA_compute_e Hash M" shows "ECDSA_PublicKeyRecovery_1 (UOracle Q\<^sub>U) e S j = Some Q\<^sub>U" using KeyRecoveryWorks_1 assms ECDSA_Sign_e_def by algebra lemma KeyRecoveryWorks_rec_h1: assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" "ECkeyPairValid k P" "P = Point x y" "j = (nat x) div n" "ECkeyPairValid d\<^sub>U Q\<^sub>U" shows "ECDSA_PublicKeyRecovery_rec (UOracle Q\<^sub>U) e S (h+1-j) = Some Q\<^sub>U" proof - have 1: "x \ carrier R" using assms(2,3) ECkeyPairOnCurve onCurveEq2 by blast have 2: "j \ h" by (metis 1 assms(4) less_p_div_n' ECparamsValid inCarrierNatInt) have 3: "h+1-(h+1-j) = j" using 2 by fastforce have 4: "ECDSA_PublicKeyRecovery_1 (UOracle Q\<^sub>U) e S j = Some Q\<^sub>U" using KeyRecoveryWorks_1 assms(1,2,3,4,5) by blast show ?thesis by (smt (verit, ccfv_SIG) 2 3 4 ECDSA_PublicKeyRecovery_rec.simps(2) Suc_diff_le Suc_eq_plus1 option.distinct(1)) qed lemma KeyRecoveryWorks_rec_h1': assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "ECkeyPairValid k P" "P = Point x y" "j = (nat x) div n" "ECkeyPairValid d\<^sub>U Q\<^sub>U" "e = ECDSA_compute_e Hash M" shows "ECDSA_PublicKeyRecovery_rec (UOracle Q\<^sub>U) e S (h+1-j) = Some Q\<^sub>U" using KeyRecoveryWorks_rec_h1 assms ECDSA_Sign_e_def by algebra lemma KeyRecoveryWorks_rec_h2: assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" "ECkeyPairValid k P" "P = Point x y" "j = (nat x) div n" "h+1-j \ L" "ECkeyPairValid d\<^sub>U Q\<^sub>U" shows "ECDSA_PublicKeyRecovery_rec (UOracle Q\<^sub>U) e S L = Some Q\<^sub>U" using assms proof (induct L) case 0 then show ?case by (metis KeyRecoveryWorks_rec_h1 bot_nat_0.extremum_uniqueI) next case (Suc L) then show ?case proof (cases "ECDSA_PublicKeyRecovery_1 (UOracle Q\<^sub>U) e S (h+1-(Suc L)) = None") case True then have "ECDSA_PublicKeyRecovery_rec (UOracle Q\<^sub>U) e S (Suc L) = ECDSA_PublicKeyRecovery_rec (UOracle Q\<^sub>U) e S L" using ECDSA_PublicKeyRecovery_rec.simps(2) diff_Suc_1 by presburger then show ?thesis by (metis KeyRecoveryWorks_rec_h1 Suc.hyps Suc.prems(1,5,6) assms(2,3,4) le_eq_less_or_eq less_Suc_eq_le) next case F0: False have F1: "ECDSA_PublicKeyRecovery_rec (UOracle Q\<^sub>U) e S (Suc L) = ECDSA_PublicKeyRecovery_1 (UOracle Q\<^sub>U) e S (h+1-(Suc L))" using ECDSA_PublicKeyRecovery_rec.simps(2) F0 by presburger have F2: "ECDSA_PublicKeyRecovery_1 (UOracle Q\<^sub>U) e S (h+1-(Suc L)) = Some Q\<^sub>U" using F0 OracleCorrect_1 by blast then show ?thesis using F1 F2 by presburger qed qed lemma KeyRecoveryWorks_rec_h2': assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "ECkeyPairValid k P" "P = Point x y" "j = (nat x) div n" "h+1-j \ L" "ECkeyPairValid d\<^sub>U Q\<^sub>U" "e = ECDSA_compute_e Hash M" shows "ECDSA_PublicKeyRecovery_rec (UOracle Q\<^sub>U) e S L = Some Q\<^sub>U" using KeyRecoveryWorks_rec_h2 assms ECDSA_Sign_e_def by algebra lemma KeyRecoveryWorks_rec: assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" "ECkeyPairValid k P" "ECkeyPairValid d\<^sub>U Q\<^sub>U" shows "ECDSA_PublicKeyRecovery_rec (UOracle Q\<^sub>U) e S (h+1) = Some Q\<^sub>U" by (metis assms(1,2,3) ECDSA_Sign_e_Some KeyRecoveryWorks_rec_h2 diff_le_self point.exhaust) lemma KeyRecoveryWorks_rec': assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "ECkeyPairValid k P" "ECkeyPairValid d\<^sub>U Q\<^sub>U" "e = ECDSA_compute_e Hash M" shows "ECDSA_PublicKeyRecovery_rec (UOracle Q\<^sub>U) e S (h+1) = Some Q\<^sub>U" using KeyRecoveryWorks_rec assms ECDSA_Sign_e_def by algebra text \Here finally is the lemma that we have been working toward. The entity U has signed the message M with the ephemeral key pair (k,P). Given a correct oracle for \Q\<^sub>U\, the key recovery operation will recover \Q\<^sub>U\ given the message M and the signature S. Again, what could the oracle be? Typically, the person trying to recover \Q\<^sub>U\ will have at least 2 messages signed by U. They can run the key recovery operation with one (message, signature) pair and use a second (message, key) pair to to check any putative \Q\<^sub>U\ that the key recovery operation presents.\ lemma KeyRecoveryWorks: assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" "ECkeyPairValid k P" "ECkeyPairValid d\<^sub>U Q\<^sub>U" shows "ECDSA_PublicKeyRecovery_e (UOracle Q\<^sub>U) e S = Some Q\<^sub>U" using ECDSA_PublicKeyRecovery_e_def KeyRecoveryWorks_rec assms(1,2,3) by presburger lemma KeyRecoveryWorks': assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "ECkeyPairValid k P" "ECkeyPairValid d\<^sub>U Q\<^sub>U" shows "ECDSA_PublicKeyRecovery Hash (UOracle Q\<^sub>U) M S = Some Q\<^sub>U" using ECDSA_PublicKeyRecovery_def ECDSA_Sign_e_def KeyRecoveryWorks assms by presburger subsubsection \4.1.7 Self-Signing Operation\ text \The statement of the self-signing operation in the standard is vague in two ways and overly complex in another. First, it takes as input "information", labeled I. The format of this "information" is not specified. It is only said that I "should include the identity of the signer", without indicating what "identity" means. Second, step 4 in the standard says "Form a message M containing both I and (r,s)." It does not indicate how such a message should be formed. Isabelle is a typed language so we need to specify the type of I. We specify the type of I as 'a, which is the generic term for "some type". And we assume the existence of a "form a message" function that takes "information" type and a signature type and produces octets. Note that the standard lists as input the curve parameters and the "information" I. But also needed as input are an ephemeral key, which we label (k, P), and a random integer s in the range [1,n-1]. As noted above, we cannot use (k, R) to represent the ephemeral key because R is already used in this theory to refer to the integer ring R of integers modulo p. Finally, let's examine step 5 in the standard. It says to "Use the Public Key Recovery Operation" to recover a public key Q. In that operation, we know only (x mod n), where x is the first coordinate of the ephemeral key point P. So we must loop over all possible x and try to reconstruct P. In the self-signing operation, we know the ephemeral key is (k,P). No need to loop to recover P. Then we get the public key is Q = r^(-1)*(s*P - e*G). Because (k,P) is a valid key pair, we know that P = k*G, so Q = r^(-1)(s*k - e)*G. But we already know this if we look at step 6 in the standard which tells us that the private key is d = r^(-1)(s*k - e). If (d,Q) is to be a valid key pair, we must have Q = d*G. In summary, there is no need to actually use the public key recovery operation. We can simply defined d as shown in step 6 and set Q = d*G. Special note for computing d: As mentioned above, Isabelle/HOL is a typed language. We need to make sure that the difference computed for d is done on integers, in case that s*k < e.\ text \These are the conditions in the standard for the user of the self-signing operation. They must pick a valid ephemeral key (k,P) and a value s in [1, n-1].\ definition ECDSA_SelfSign_validIn' :: "nat \ int point \ nat \ bool" where "ECDSA_SelfSign_validIn' k P s = (ECkeyPairValid k P \ 0 < s \ s < n)" text \These are the conditions that guarantee that the self-signing operation will produce some output.\ definition ECDSA_SelfSign_validIn :: "hash_type \ ('a \ sig_type \ octets) \ 'a \ nat \ int point \ nat \ bool" where "ECDSA_SelfSign_validIn Hash FormMessage I k P s = ( case P of Infinity \ False | Point x y \ ( let r = nat (x mod n); M = FormMessage I (r,s); e = ECDSA_compute_e Hash M; rInv = inv_mod_n r; d = nat ((rInv*(s*(int k) - (int e)) ) mod n) in 0 < r \ 0 < d ) )" definition ECDSA_SelfSign :: "hash_type \ ('a \ sig_type \ octets) \ 'a \ nat \ int point \ nat \ (nat \ int point \ octets) option" where "ECDSA_SelfSign Hash FormMessage I k P s = ( case P of Infinity \ None | Point x y \ ( let r = nat (x mod n); M = FormMessage I (r,s); e = ECDSA_compute_e Hash M; rInv = inv_mod_n r; d = nat ((rInv*(s*(int k) - (int e)) ) mod n); Q = point_mult' d G in ( if (r=0) then None else ( if (d=0) then None else Some (d, Q, M) )) ) )" lemma ECDSA_SelfSign_ValidIn: assumes "ECDSA_SelfSign_validIn Hash FormMessage I k P s" shows "\d Q M. ECDSA_SelfSign Hash FormMessage I k P s = Some (d, Q, M)" proof - have 1: "P \ Infinity" using assms ECDSA_SelfSign_validIn_def by force obtain x and y where 2: "P = Point x y" by (meson 1 point.exhaust) let ?r = "nat (x mod n)" let ?M = "FormMessage I (?r,s)" let ?e = "ECDSA_compute_e Hash ?M" let ?rInv = "inv_mod_n ?r" let ?d = "nat ((?rInv*(s*(int k) - (int ?e)) ) mod n)" have 3: "0 < ?r" by (smt (verit) assms 1 2 ECDSA_SelfSign_validIn_def point.simps(5)) have 4: "0 < ?d" by (smt (verit) assms 1 2 ECDSA_SelfSign_validIn_def point.simps(5)) have 5: "?r \ 0" using 3 by fast have 6: "?d \ 0" using 4 by fast show ?thesis using 1 2 5 6 ECDSA_SelfSign_def option.discI point.simps(5) by (smt (z3)) qed lemma ECDSA_SelfSign_PnotInf: assumes "ECDSA_SelfSign Hash FormMessage I k P s = Some (d, Q, M)" shows "P \ Infinity" by (metis (no_types, lifting) assms(1) ECDSA_SelfSign_def option.discI point.simps(4)) lemma ECDSA_SelfSign_Some: fixes x :: int assumes "ECDSA_SelfSign Hash FormMessage I k P s = Some (d, Q, M)" "P = Point x y" "r = nat (x mod n)" "M' = FormMessage I (r,s)" "e = ECDSA_compute_e Hash M'" "rInv = inv_mod_n r" "d' = nat ((rInv*(s*(int k) - (int e)) ) mod n)" "Q' = point_mult' d' G" shows "d = d' \ Q = Q' \ M = M' \ r \0 \ d \ 0" by (smt (verit, del_insts) assms ECDSA_SelfSign_def option.simps(1,3) point.case(2) prod.sel(1,2)) lemma ECDSA_SelfSign_SomeValid: assumes "ECDSA_SelfSign Hash FormMessage I k P s = Some (d, Q, M)" shows "ECDSA_SelfSign_validIn Hash FormMessage I k P s" proof - have 1: "P \ Infinity" by (metis ECDSA_SelfSign_PnotInf assms) obtain x and y where 2: "P = Point x y" by (meson 1 point.exhaust) let ?r = "nat (x mod n)" let ?M = "FormMessage I (?r,s)" let ?e = "ECDSA_compute_e Hash ?M" let ?rInv = "inv_mod_n ?r" let ?d = "nat ((?rInv*(s*(int k) - (int ?e)) ) mod n)" let ?Q = "point_mult' ?d G" have 3: "d = ?d \ Q = ?Q \ M = ?M \ ?r \0 \ d \ 0" by (metis 2 ECDSA_SelfSign_Some assms) have 4: "0 < ?r \ 0 < ?d" using 3 by force show ?thesis using 1 2 4 ECDSA_SelfSign_validIn_def by fastforce qed text \The output (d,Q) of the self-signing operation is a valid key pair.\ lemma ECDSA_SelfSign_ValidKeyPair: assumes "ECDSA_SelfSign Hash FormMessage I k P s = Some (d, Q, M)" shows "ECkeyPairValid d Q" proof - have 1: "P \ Infinity" by (metis ECDSA_SelfSign_PnotInf assms) obtain x and y where 2: "P = Point x y" by (meson 1 point.exhaust) let ?r = "nat (x mod n)" let ?M = "FormMessage I (?r,s)" let ?e = "ECDSA_compute_e Hash ?M" let ?rInv = "inv_mod_n ?r" let ?d = "nat ((?rInv*(s*(int k) - (int ?e)) ) mod n)" let ?Q = "point_mult' ?d G" have 3: "d = ?d \ Q = ?Q \ M = ?M \ ?r \0 \ d \ 0" by (metis 2 ECDSA_SelfSign_Some assms) have 4: "d < n" by (metis 3 Euclidean_Rings.pos_mod_bound N.residues_prime_axioms nat_int prime_gt_1_nat residues_prime.p_prime zless_nat_conj) show ?thesis using 3 4 ECkeyPairValid_def by blast qed text \Here is the main result about the self-signing operation. Namely, if you follow the self-signing operation and get (d, Q, M), then the message M is formed from the information I and (r,s) and when M is signed using (d,Q) as the signer's key pair (together with the ephemeral key (k,P), then the resulting signature is (r,s).\ lemma ECDSA_SelfSign_SignWorks: assumes "ECDSA_SelfSign Hash FormMessage I k P s = Some (d, Q, M)" "P = Point x y" "r = nat (x mod n)" "ECkeyPairValid k P" "s ECDSA_Sign Hash d M k P = Some (r,s)" proof - let ?M = "FormMessage I (r,s)" let ?e = "ECDSA_compute_e Hash M" let ?rInv = "inv_mod_n r" let ?d = "nat ((?rInv*(s*(int k) - (int ?e)) ) mod n)" let ?Q = "point_mult' d G" have 1: "d = ?d \ Q = ?Q \ M = ?M \ r \0 \ d \ 0" by (metis ECDSA_SelfSign_Some assms(1,2,3)) have r1: "r \ carrier Rn" by (metis 1 N.mod_in_carrier assms(3) nat_eq_iff2) have r2: "?rInv = inv\<^bsub>Rn\<^esub> r" using 1 r1 inv_mod_n_inv by presburger have d1: "?d = inv\<^bsub>Rn\<^esub>r \\<^bsub>Rn\<^esub> (s \\<^bsub>Rn\<^esub> k \\<^bsub>Rn\<^esub> ?e)" by (metis 1 r2 N.res_diff_eq N.res_mult_eq int_nat_eq int_ops(1) mod_diff_left_eq mod_mult_right_eq nat_int) let ?kinv = "inv_mod_n k" have k1: "0 < k \ k < n" using assms(4) ECkeyPairValid_def by blast have k2: "?kinv = inv\<^bsub>Rn\<^esub> k" using inv_mod_n_inv' k1 by presburger have k3: "k \ carrier Rn" using k1 ECkeyPair_dInRn' ECprivateKeyValid_def by blast let ?s = "(?kinv*(?e + r*?d)) mod n" have s1: "?s = inv\<^bsub>Rn\<^esub> k \\<^bsub>Rn\<^esub> (?e \\<^bsub>Rn\<^esub> r \\<^bsub>Rn\<^esub> ?d)" by (simp add: k2 N.res_add_eq N.res_mult_eq mod_add_right_eq mod_mult_right_eq of_nat_mod) have d2: "r \\<^bsub>Rn\<^esub> ?d = s \\<^bsub>Rn\<^esub> k \\<^bsub>Rn\<^esub> ?e" by (metis 1 r1 d1 EPF.nat_int_eq N.inv_closed N.l_one N.m_assoc N.mod_in_carrier N.r_inv N.res_diff_eq N.zero_cong int_ops(1)) have s2: "(?e \\<^bsub>Rn\<^esub> r \\<^bsub>Rn\<^esub> ?d) = s \\<^bsub>Rn\<^esub> k" using d2 by (simp add: N.res_add_eq N.res_diff_eq N.res_mult_eq mod_add_right_eq) have s3: "?s = inv\<^bsub>Rn\<^esub> k \\<^bsub>Rn\<^esub> (s \\<^bsub>Rn\<^esub> k)" using s1 s2 by argo have s4: "?s = inv\<^bsub>Rn\<^esub> k \\<^bsub>Rn\<^esub> k \\<^bsub>Rn\<^esub> s" using s3 by (simp add: N.res_mult_eq mod_mult_right_eq mult.commute mult.left_commute) have s5: "?s = s" by (metis s4 k3 assms(4,5,6) ECkeyPair_dInRn' ECkeyPair_invRn ECprivateKeyValid_def N.l_one N.m_comm nat_int) show ?thesis by (smt (verit) 1 s5 ECDSA_Sign_e_def assms(2,3,6) bot_nat_0.not_eq_extremum point.simps(5)) qed text \This is a restatement of the lemma above where we use "ECDSA_SelfSign_validIn k P s", which again means that (k,P) is a valid key pair and s is in [1,n-1].\ lemma ECDSA_SelfSign_SignWorks': assumes "ECDSA_SelfSign Hash FormMessage I k P s = Some (d, Q, M)" "P = Point x y" "r = nat (x mod n)" "ECDSA_SelfSign_validIn' k P s" shows "M = FormMessage I (r,s) \ ECDSA_Sign Hash d M k P = Some (r,s)" using ECDSA_SelfSign_SignWorks assms ECDSA_SelfSign_validIn'_def by fastforce end (* SEC1 locale *) section \5 Encryption and Key Transport Schemes\ text \"This section specifies the public-key encryption and key transport schemes based on ECC supported in this document. Public-key encryption schemes are designed to be used by two entities - a sender U and a recipient V - when U wants to send a message M to V confidentially, and V wants to recover M. Key transport schemes are a special class of public-key encryption schemes where the message M is restricted to be a cryptographic key, usually a symmetric key. Except for this restriction, most of the discussion below about public-key encryption schemes also applies to key transport schemes. Here, public-key encryption schemes are described in terms of an encryption operation, a decryption operation, and associated setup and key deployment procedures. Entities U and V should use the scheme as follows when they want to communicate. First U and V should use the setup procedure to establish which options to use the scheme with, then V should use the key deployment procedure to select a key pair and U should obtain V's public key - U will use V's public key to control the encryption procedure, and V will use its key pair to control the decryption operation. Then each time U wants to send a message M to V, U should apply the encryption operation to M under V's public key to compute an encryption or ciphertext C of M , and convey C to V . Finally when V receives C, entity V should apply the decryption operation to C under its key pair to recover the message M. Loosely speaking, public-key encryption schemes are designed so that it is hard for an adversary who does not possess V 's secret key to recover messages from their ciphertexts. Thereby, public-key encryption schemes provide data confidentiality."\ subsection \5.1 Elliptic Curve Integrated Encryption Scheme\ locale ECIES = SEC1 + fixes KDF :: kdf_type and KDF_VI :: kdf_validin_type and MAC :: mac_type and MAC_VI :: mac_validin_type and mackeylen maclen :: nat and ENC :: enc_type and ENC_VI :: enc_validin_type and DEC :: dec_type and DEC_VI :: dec_validin_type and enckeylen :: nat and useCoDH useCompress :: bool and d\<^sub>V :: nat and Q\<^sub>V :: "int point" assumes VkeyPairValid : "ECkeyPairValid d\<^sub>V Q\<^sub>V" and ENC_valid : "\K M. ENC_VI K M \ DEC K (ENC K M) = M" begin subsubsection \5.1.3 Encryption Operation\ text\Note for Step 6: This definition doesn't cover the case when ENC = XOR and backward compatibility mode is not selected.\ definition ECIES_Encryption :: "octets \ octets \ octets \ nat \ int point \ (octets \ octets \ octets) option" where "ECIES_Encryption M SharedInfo1 SharedInfo2 k P = ( case (ECDHprimChoose useCoDH k Q\<^sub>V) of None \ None | Some z \ ( let Pbar = point_to_octets P useCompress; mlen = octet_length p; Z = nat_to_octets_len (nat z) mlen; K = KDF Z (enckeylen + mackeylen) SharedInfo1; EK = take enckeylen K; MK = drop enckeylen K; EM = ENC EK M; D = MAC MK (EM @ SharedInfo2) in Some (Pbar, EM, D) ) )" definition ECIES_Encryption_validIn :: "octets \ octets \ octets \ nat \ int point \ bool" where "ECIES_Encryption_validIn M SharedInfo1 SharedInfo2 k P = ( case (ECDHprimChoose useCoDH k Q\<^sub>V) of None \ False | Some z \ ( let mlen = octet_length p; Z = nat_to_octets_len (nat z) mlen; K = KDF Z (enckeylen + mackeylen) SharedInfo1; EK = take enckeylen K; MK = drop enckeylen K; EM = ENC EK M in (ECkeyPairValid k P) \ (ENC_VI EK M) \ (KDF_VI Z (enckeylen + mackeylen) SharedInfo1) \ (MAC_VI MK (EM @ SharedInfo2)) ) )" lemma ECIES_Encryption_validIn_someOut: assumes "ECIES_Encryption_validIn M SharedInfo1 SharedInfo2 k P" shows "ECIES_Encryption M SharedInfo1 SharedInfo2 k P \ None" by (smt (z3) assms ECIES_Encryption_validIn_def ECIES_Encryption_def case_optionE option.discI option.simps(5)) lemma ECIES_Encryption_validIn_someZ: assumes "ECIES_Encryption_validIn M SharedInfo1 SharedInfo2 k P" shows "\z. ECDHprimChoose useCoDH k Q\<^sub>V = Some z" using assms ECIES_Encryption_validIn_def case_optionE by blast lemma ECIES_Encryption_someZout: assumes "ECDHprimChoose useCoDH k Q\<^sub>V = Some z" shows "\Pbar EM D. ECIES_Encryption M SharedInfo1 SharedInfo2 k P = Some (Pbar, EM, D)" by (metis (no_types, lifting) assms ECIES_Encryption_def option.simps(5)) lemma ECIES_Encryption_SomeOutEq: assumes "ECDHprimChoose useCoDH k Q\<^sub>V = Some z" "Pbar = point_to_octets P useCompress" "mlen = octet_length p" "Z = nat_to_octets_len (nat z) mlen" "K = KDF Z (enckeylen + mackeylen) SharedInfo1" "EK = take enckeylen K" "MK = drop enckeylen K" "EM = ENC EK M" "D = MAC MK (EM @ SharedInfo2)" shows "ECIES_Encryption M SharedInfo1 SharedInfo2 k P = Some (Pbar, EM, D)" by (metis (mono_tags, lifting) assms ECIES_Encryption_def option.simps(5)) subsubsection \5.1.4 Decryption Operation\ definition ECIES_Decryption :: "(octets \ octets \ octets) \ octets \ octets \ octets option" where "ECIES_Decryption C SharedInfo1 SharedInfo2 = ( let Pbar = fst C; EM = fst (snd C); D = snd (snd C) in ( case (octets_to_point' Pbar) of None \ None | Some P \ ( let T = (if useCoDH then (ECpublicKeyPartialValid P) else (ECpublicKeyValid P) ) in ( case (ECDHprimChoose useCoDH d\<^sub>V P) of None \ None | Some z \ ( let mlen = octet_length p; Z = nat_to_octets_len (nat z) mlen; K = KDF Z (enckeylen + mackeylen) SharedInfo1; EK = take enckeylen K; MK = drop enckeylen K; D' = MAC MK (EM @ SharedInfo2); M = DEC EK EM in if (T \ D = D') then (Some M) else None ) ) ) ) )" text \This is the main thing we would like to know. If you encrypt a message with ECIES_Encryption, then you can get back that message with ECIES_Decryption.\ lemma ECIES_Decryption_validIn: assumes "ECIES_Encryption_validIn M SharedInfo1 SharedInfo2 k P" "ECIES_Encryption M SharedInfo1 SharedInfo2 k P = Some (Pbar, EM, D)" shows "ECIES_Decryption (Pbar, EM, D) SharedInfo1 SharedInfo2 = Some M" proof - obtain z where z1: "ECDHprimChoose useCoDH k Q\<^sub>V = Some z" using ECIES_Encryption_validIn_someZ assms(1) by blast let ?mlen = "octet_length p" let ?Z = "nat_to_octets_len (nat z) ?mlen" let ?K = "KDF ?Z (enckeylen + mackeylen) SharedInfo1" let ?EK = "take enckeylen ?K" let ?MK = "drop enckeylen ?K" have P1: "Pbar = point_to_octets P useCompress" by (metis (mono_tags, lifting) assms(2) z1 ECIES_Encryption_SomeOutEq option.simps(1) prod.sel(1)) have P2: "ECkeyPairValid k P" by (metis (no_types, lifting) assms(1) z1 ECIES_Encryption_validIn_def option.simps(5)) have P3: "on_curve' P" using P2 ECkeyPairOnCurve by blast have P4: "octets_to_point' Pbar = Some P" using P1 P3 point2Octets2Point by fast have P5: "ECpublicKeyValid P \ ECpublicKeyPartialValid P" using P2 keyPairValidImpliespublicKeyValid validImpliesPartialValid by blast let ?T = "(if useCoDH then (ECpublicKeyPartialValid P) else (ECpublicKeyValid P) )" have T1: "?T" using P5 by presburger have M1: "EM = ENC ?EK M" using assms(2) z1 ECIES_Encryption_SomeOutEq option.simps(1) prod.sel(1) by simp let ?M = "DEC ?EK EM" have M2: "ENC_VI ?EK M" using assms(1) ECIES_Encryption_validIn_def by (smt (verit, best) assms(1) ECIES_Encryption_validIn_def z1 option.simps(5)) have M3: "?M = M" using M1 M2 ENC_valid by blast have D1: "D = MAC ?MK (EM @ SharedInfo2)" using assms(2) z1 ECIES_Encryption_SomeOutEq option.simps(1) prod.sel(1) by force have z2: "ECDHprimChoose useCoDH k Q\<^sub>V = ECDHprimChoose useCoDH d\<^sub>V P" using P2 VkeyPairValid ECDHch_2ValidKeyPairs by blast show ?thesis using z1 z2 P4 T1 M3 D1 ECIES_Decryption_def by force qed subsection \5.2 Wrapped Key Transport Scheme\ text \"The wrapped key transport scheme uses a combination of a key wrap scheme and a key agreement scheme." This section of the standard discusses how a wrapped key transport scheme may work without making any specific definitions.\ end (* of ECIES locale *) section \6 Key Agreement Schemes\ text \"Key agreement schemes are designed to be used by two entities --- an entity U and an entity V --- when U and V want to establish shared keying data that they can later use to control the operation of a symmetric cryptographic scheme. ... Note that this document does not address how specific keys should be derived from keying data established using a key agreement scheme. This detail is left to be determined on an application by application basis." \ subsection \6.1 Elliptic Curve Diffie-Hellman Scheme\ locale ECDH = SEC1 + fixes KDF :: kdf_type and KDF_VI :: kdf_validin_type and useCoDH :: bool and d\<^sub>V d\<^sub>U :: nat and Q\<^sub>V Q\<^sub>U :: "int point" assumes VkeyPairValid : "ECkeyPairValid d\<^sub>V Q\<^sub>V" and UkeyPairValid : "ECkeyPairValid d\<^sub>U Q\<^sub>U" begin subsubsection \6.1.3 Key Agreement Operation\ definition ECDH_KeyAgreement :: "nat \ octets \ octets option" where "ECDH_KeyAgreement keydatalen SharedInfo = ( let z = ECDHprimChoose useCoDH d\<^sub>U Q\<^sub>V in case z of None \ None | Some z' \ ( let mlen = octet_length p; Z = nat_to_octets_len (nat z') mlen in Some (KDF Z keydatalen SharedInfo) ) )" text \Remember ECDHchoose_validKeys: Because these are both valid key pairs, we know that \\z. ECDHprimChoose useCo d Q = Some z\.\ definition ECDH_KeyAgreement_validIn :: "nat \ octets \ bool" where "ECDH_KeyAgreement_validIn keydatalen SharedInfo \ \z. (ECDHprimChoose useCoDH d\<^sub>U Q\<^sub>V = Some z) \ (KDF_VI (nat_to_octets_len (nat z) (octet_length p)) keydatalen SharedInfo)" text \This lemma shows that entity V can get the same key that U produced. U produces the key with U's private key and V's public key. V produces the same key with V's private key and U's public key.\ lemma ECDH_KeyAgreement_eq: "ECDH_KeyAgreement keydatalen SharedInfo = ( let z = ECDHprimChoose useCoDH d\<^sub>V Q\<^sub>U in case z of None \ None | Some z' \ ( let mlen = octet_length p; Z = nat_to_octets_len (nat z') mlen in Some (KDF Z keydatalen SharedInfo) ) )" by (metis ECDH_KeyAgreement_def ECDHch_2ValidKeyPairs UkeyPairValid VkeyPairValid option.simps(5) ECDHchoose_validKeys ECkeyPairImpliesPrivateKeyValid keyPairValidImpliespublicKeyValid) end (* of ECDH locale *) subsection \6.2 Elliptic Curve MQV Scheme\ locale ECMQV = SEC1 + fixes KDF :: kdf_type and d1U d2U d1V d2V :: nat and Q1U Q2U Q1V Q2V :: "int point" assumes VkeyPairValid : "ECkeyPairValid d1V Q1V" "ECkeyPairValid d2V Q2V" and UkeyPairValid : "ECkeyPairValid d1U Q1U" "ECkeyPairValid d2U Q2U" begin subsubsection \6.2.3 Key Agreement Operation\ definition ECMQV_KeyAgreement :: "nat \ octets \ octets option" where "ECMQV_KeyAgreement keydatalen SharedInfo = ( let z = ECMQVprim d1U Q1U d2U Q2U Q1V Q2V in case z of None \ None | Some z' \ ( let mlen = octet_length p; Z = nat_to_octets_len (nat z') mlen in Some (KDF Z keydatalen SharedInfo) ) )" text \As with the Diffie-Hellman Key Agreement operation, the important thing to note with the ECMQV Key Agreement operation is that both U and V can compute the same key using their own keys together with the public keys of the other party. This follows from MQV_reverseUV above.\ lemma ECMQV_KeyAgreement_eq: "ECMQV_KeyAgreement keydatalen SharedInfo = ( let z = ECMQVprim d1V Q1V d2V Q2V Q1U Q2U in case z of None \ None | Some z' \ ( let mlen = octet_length p; Z = nat_to_octets_len (nat z') mlen in Some (KDF Z keydatalen SharedInfo) ) )" by (metis ECMQV_KeyAgreement_def MQV_reverseUV UkeyPairValid VkeyPairValid) end (* of ECMQV locale *) end \ No newline at end of file diff --git a/thys/Crypto_Standards/Words.thy b/thys/Crypto_Standards/Words.thy --- a/thys/Crypto_Standards/Words.thy +++ b/thys/Crypto_Standards/Words.thy @@ -1,2176 +1,2174 @@ theory Words imports More_Bit_Operations_Nat "HOL.Transcendental" "HOL-Number_Theory.Residues" begin text \This theory implements the conversion from non-negative integers to a string of "octets" (eight-bit bytes) and back again as defined in PKCS #1 v2.2. Moreover, this theory generalizes these conversions to be between non-negative integers and w-bit words. The following is an excerpt from the standard where we have replaced "octet" with "word" and any instance of 8 with w: "Two data conversion primitives are employed in the schemes defined in this document: •I2OSP – Integer-to-Word-String primitive •OS2IP – Word-String-to-Integer primitive For the purposes of this document, and consistent with ASN.1 syntax, a word string is an ordered sequence of words (w-bit values). The sequence is indexed from first (conventionally, leftmost) to last (rightmost). For purposes of conversion to and from integers, the first word is considered the most significant in the following conversion primitives. I2OSP – Integer-to-Word-String primitive: Write the integer x in its unique xLen-digit representation in base (2^w): x = x_{xLen–1} (2^w)^{xLen–1} + x_{xLen–2} (2^w)^{xLen–2} + ... + x_1 (2^w) + x_0, where 0 \ x_i < (2^w) (note that one or more leading digits will be zero if x is less than (2^w)^{xLen–1}). Let the word X_i have the integer value x_{xLen–i} for 1 \ i \xLen. Output the word string X = X_1 X_2 ... X_{xLen}. OS2IP – Word-String-to-Integer primitive: Let X_1 X_2 ...X_{xLen} be the words of X from first to last, and let x_{xLen–i} be the integer value of the word X_i for 1 \ i \ xLen. Let x = x_{xLen–1} (2^w)^{xLen–1} + x_{xLen–2} (2^w)^{xLen–2} + ... + x_1 (2^w) + x_0. Output x." For this theory, we want to be agnostic to the variable type that might be used to store the word in any particular implementation of the standard. Thus we concern ourselves only with the (non- negative) integer value of the words. Thus our words are nats and our word strings are lists of nats. Then for a w-bit word to be valid, it must be less than (2^w). Also, the standard only has an opinion about what I2OSP should be when the input x can be written in xLen or fewer words. We are not constrained by the standard when defining nat_to_words_len when x is too large. We choose to define nat_to_words_len so that it converts the natural number x to as many words as required, even if it is more than xLen. The other choice would be to truncate the list of words at xLen. Either choice is valid and only affects the sorts of lemmas we must prove. \ type_synonym words = "nat list" lemma n_div_2tow_cases: assumes zero: "(n::nat) = 0 \ R" and div : "\ n div 2^w < n ; 0 < n \ \ R" and wpos: "0 < w" shows "R" proof (cases "n = 0") assume "n = 0" thus R by (rule zero) next assume "n \ 0" hence "0 < n" by simp hence "n div 2^w < n" by (metis wpos One_nat_def Suc_leI div_less_dividend le_less_trans less_exp) from this and \0 < n\ show R by (rule div) qed section\Words Valid\ text \A nat is a valid w-bit word exactly when it is < 2^w. So a list of nats is valid when all the elements are < 2^w.\ definition words_valid :: "nat \ words \ bool" where "words_valid w os = (\b\set os. b < 2^w)" lemma words_valid_nil: "words_valid w []" using words_valid_def by force lemma words_valid_zeros: "words_valid w (replicate n 0)" using words_valid_def by simp lemma words_valid_concat: assumes "words_valid w a" "words_valid w b" shows "words_valid w (a @ b)" using assms words_valid_def by fastforce lemma words_valid_cons: assumes "words_valid w as" "a < 2^w" shows "words_valid w (a # as)" using assms words_valid_def by auto lemma words_valid_trunc: assumes "words_valid w (as @ bs)" shows "words_valid w as \ words_valid w bs" using assms words_valid_def by force lemma words_valid_take: assumes "words_valid w as" shows "words_valid w (take n as)" by (metis append_take_drop_id assms words_valid_trunc) lemma words_valid_drop: assumes "words_valid w as" shows "words_valid w (drop n as)" by (meson assms in_set_dropD words_valid_def) lemma words_valid_hd: assumes "words_valid w as" "as \ []" shows "(hd as) < 2^w" using assms list.set_sel(1) words_valid_def by blast lemma words_valid_tl: assumes "words_valid w as" shows "words_valid w (tl as)" by (metis (no_types) assms drop0 drop_Suc words_valid_drop) lemma words_valid_last: assumes "words_valid w as" "as \ []" shows "last as < 2^w" using assms last_in_set words_valid_def by blast lemma words_valid_butlast: assumes "words_valid w as" shows "words_valid w (butlast as)" by (meson assms words_valid_def in_set_butlastD) lemma words_valid_ith: assumes "words_valid w as" "i < length as" shows "as ! i < 2^w" using assms nth_mem words_valid_def by blast text \In FIPS 180-4, addition is done modulo the wordsize. The following is convenient to have.\ lemma words_valid_sum_mod: "words_valid w (map2 (\x y. (x+y) mod (2^w)) X Y)" proof (induct "length (map2 (\x y. (x+y) mod (2^w)) X Y)" arbitrary: X Y) case 0 then show ?case by (metis length_0_conv words_valid_nil) next case C: (Suc xa) let ?Z = "map2 (\x y. (x+y) mod (2^w)) X Y" let ?hZ = "hd ?Z" let ?tZ = "tl ?Z" have 1: "?hZ = ((hd X) + (hd Y)) mod 2^w" by (metis (no_types, lifting) Nil_is_map_conv C.hyps(2) hd_zip length_0_conv list.map_sel(1) nat.simps(3) prod.simps(2) zip_eq_Nil_iff) have 2: "?hZ < 2^w" using 1 by simp have 3: "length ?tZ = xa" using C(2) by auto have 4: "?Z = ?hZ # ?tZ" by (metis C(2) list.exhaust_sel list.size(3) old.nat.distinct(1)) have 5: "?tZ = map2 (\x y. (x+y) mod (2^w)) (tl X) (tl Y)" by (smt (z3) 4 Nil_is_map_conv list.exhaust_sel list.sel(3) map_tl zip_Cons_Cons zip_eq_Nil_iff) have 6: "words_valid w ?tZ" by (metis C(1) 3 5) show ?case by (metis 2 4 6 words_valid_cons) qed section\Strip Words\ text\In the PKCS #1 standard, we sometimes pad a word string with 0s and sometimes with non-0s. Then when decoding, we must strip an word string of leading 0s (resp. non-0s).\ fun words_strip_zero :: "words \ words" where "words_strip_zero [] = []" | "words_strip_zero os = (if (hd os = 0) then words_strip_zero (tl os) else os)" fun words_strip_nonzero :: "words \ words" where "words_strip_nonzero [] = []" | "words_strip_nonzero os = (if (hd os = 0) then os else words_strip_nonzero (tl os) )" lemma strip_prepend_zeros: "words_strip_zero ((replicate x 0) @ bs) = words_strip_zero bs" apply (induction x) by auto lemma strip_zero_length: "length (words_strip_zero bs) \ length bs" apply (induction bs) by auto lemma strip_zero_leading_zeros: assumes "ls = length (words_strip_zero bs)" "lb = length bs" shows "\i < (lb - ls). bs ! i = 0" proof - let ?l = "lb - ls" show ?thesis using assms proof (induction ?l arbitrary: bs lb ls) case 0 then show ?case by fastforce next case 1: (Suc l) let ?h = "hd bs" let ?t = "tl bs" let ?s = "words_strip_zero bs" have 3: "lb = l + 1 + ls" using 1(2) by simp have 4: "0 < lb" using 1(2) by simp have 5: "bs = ?h # ?t" using 1(4) 4 by fastforce have 6: "?h = 0" by (metis 1(2,3,4) 5 Zero_not_Suc cancel_comm_monoid_add_class.diff_cancel words_strip_zero.simps(2)) have 7: "?s = words_strip_zero ?t" using 5 6 words_strip_zero.simps(2) by metis show ?case by (metis 1(1,2,3,4) 5 6 7 cancel_ab_semigroup_add_class.diff_right_commute diff_Suc_1 length_tl less_Suc_eq_0_disj nth_Cons') qed qed lemma strip_zero_head: assumes "words_strip_zero bs \ []" shows "0 < (words_strip_zero bs) ! 0" using assms proof (induction bs) case Nil then show ?case using words_strip_zero.simps(1) by blast next case (Cons a bs) then show ?case by auto qed lemma strip_zero_drop: "words_strip_zero bs = drop ((length bs) - (length (words_strip_zero bs))) bs" proof (induction bs) case Nil then show ?case by force next case (Cons a bs) then show ?case by (metis (no_types, opaque_lifting) add.right_neutral add_Suc_right add_diff_cancel_left' diff_le_self drop_Cons' le_add_diff_inverse length_Cons length_drop list.sel(3) nat.simps(3) words_strip_zero.simps(2)) qed lemma strip_zero_concat: assumes "bl = length bs" "s = words_strip_zero bs" "sl = length s" "z = (replicate (bl-sl) 0)" shows "bs = z @ s" by (metis append_take_drop_id assms diff_le_self le_eq_less_or_eq length_replicate nth_replicate nth_take_lemma strip_zero_drop strip_zero_leading_zeros take_all) lemma strip_non0_nil: "words_strip_nonzero [] = []" by simp lemma strip_non0_nil2: assumes "words_strip_nonzero as = []" shows "\a \ set as. 0 < a" using assms proof (induct as) case Nil then show ?case by force next case (Cons a as) then show ?case by (metis list.distinct(1) list.sel(1) list.sel(3) neq0_conv words_strip_nonzero.simps(2) set_ConsD) qed lemma strip_non0_nil3: assumes "\a \ set as. 0 < a" shows "words_strip_nonzero as = []" using assms proof (induct as) case Nil then show ?case by auto next case (Cons a as) then show ?case by simp qed lemma strip_non0_cons: assumes "0 < a" shows "words_strip_nonzero (a # as) = words_strip_nonzero as" using assms by auto lemma strip_non0_cons2: "words_strip_nonzero (0 # as) = (0 # as)" by simp lemma strip_non0_concat1: assumes "words_strip_nonzero as \ []" shows "words_strip_nonzero (as @ bs) = (words_strip_nonzero as) @ bs" using assms proof (induct as arbitrary: bs) case Nil then show ?case by force next case (Cons a as) then show ?case by auto qed lemma strip_non0_concat: assumes "\a \ set as. 0 < a" shows "words_strip_nonzero (as @ [0] @ bs) = [0] @ bs" using assms proof (induct as arbitrary: bs) case Nil then show ?case by simp next case I: (Cons a as) have 1: "0 < a" using I by simp have 2: "words_strip_nonzero (a # as @ [0] @ bs) = words_strip_nonzero (as @ [0] @ bs)" using 1 strip_non0_cons by blast then show ?case using 2 I by auto qed lemma strip_non0_len: assumes "as = words_strip_nonzero bs" shows "(length as) \ (length bs)" using assms proof (induct bs arbitrary: as) case Nil then show ?case by simp next case (Cons a bs) then show ?case by force qed lemma strip_non0_drop: "words_strip_nonzero bs = drop ((length bs) - (length (words_strip_nonzero bs))) bs" proof (induction bs) case Nil then show ?case by fastforce next case (Cons a bs) then show ?case using Suc_diff_le strip_non0_len by force qed lemma strip_non0_head: assumes "words_strip_nonzero bs \ []" shows "(words_strip_nonzero bs) ! 0 = 0" using assms proof (induction bs) case Nil then show ?case by auto next case (Cons a bs) then show ?case by force qed lemma strip_non0_idem: "words_strip_nonzero (words_strip_nonzero bs) = words_strip_nonzero bs" by (metis hd_conv_nth strip_non0_head strip_non0_nil words_strip_nonzero.elims) lemma strip_non0_leading_pos: assumes "ls = length (words_strip_nonzero bs)" "lb = length bs" shows "\i < (lb - ls). 0 < bs ! i" proof - let ?l = "lb - ls" show ?thesis using assms proof (induction ?l arbitrary: bs lb ls) case 0 then show ?case by fastforce next case 1: (Suc l) let ?h = "hd bs" let ?t = "tl bs" let ?s = "words_strip_nonzero bs" have 3: "lb = l + 1 + ls" using 1(2) by simp have 4: "0 < lb" using 1(2) by simp have 5: "bs = ?h # ?t" using 1(4) 4 by fastforce have 6: "0 < ?h" by (metis 1(2,3,4) 5 Zero_not_Suc cancel_comm_monoid_add_class.diff_cancel words_strip_nonzero.simps(2) gr0I) have 7: "?s = words_strip_nonzero ?t" by (metis strip_non0_cons 5 6) show ?case by (metis 1(1,2,3,4) 5 6 7 cancel_ab_semigroup_add_class.diff_right_commute diff_Suc_1 length_tl less_Suc_eq_0_disj nth_Cons') qed qed section\Nat to Words\ text\Convert any natural number to a word string. Also, we define a function that converts a natural number to a specified minimum number of words by padding with zeros if necessary.\ function nat_to_words_rec :: "nat \ nat \ words \ words" where "nat_to_words_rec w n nl = (if w = 0 then nl else (if n = 0 then nl else nat_to_words_rec w (n div (2^w)) ((n mod (2^w))#nl) ) )" by auto termination nat_to_words_rec apply (relation "measure (\(w, n, nl). n)") apply auto using n_div_2tow_cases by blast definition nat_to_words :: "nat \ nat \ words" where "nat_to_words w n = nat_to_words_rec w n []" definition nat_to_words_len :: "nat \ nat \ nat \ words" where "nat_to_words_len w n len = (let l = length (nat_to_words w n) in (replicate (len-l) 0) @ (nat_to_words w n))" definition word_length :: "nat \ nat \ nat" where "word_length w n = length (nat_to_words w n)" lemma zero_to_word [simp]: "nat_to_words w 0 = []" by (simp add: nat_to_words_def) lemma zero_to_words_len [simp]: "nat_to_words_len w 0 l = replicate l 0" using zero_to_word nat_to_words_len_def by auto lemma nat_to_words_len_len: "l \ length (nat_to_words_len w n l)" by (simp add: nat_to_words_len_def) lemma nat_to_zero_bit_words [simp]: "nat_to_words 0 n = []" by (simp add: nat_to_words_def) lemma nat_to_zero_bit_words_len [simp]: "nat_to_words_len 0 n l = replicate l 0" by (simp add: nat_to_words_len_def) lemma nat_to_zero_bit_words_len2 [simp]: "length (nat_to_words_len 0 n l) = l" by (simp add: nat_to_words_len_def) lemma nat_to_zero_bit_word_len [simp]: "word_length 0 n = 0" by (simp add: word_length_def) lemma wpos: assumes "(n::nat) < 2^w" "0 < n" shows "0 < w" using assms less_one by fastforce lemma nat_to_one_word [simp]: assumes "n < 2^w" "0 < n" shows "nat_to_words w n = [n]" using wpos assms nat_to_words_def by force lemma nat_to_two_words: assumes "n < 2^(2*w)" "2^w \ n" shows "nat_to_words w n = [n div 2^w, n mod 2^w]" proof - have 0: "0 < w" using assms leD by fastforce have 1: "n div 2^w < 2^w" by (metis assms(1) less_mult_imp_div_less power2_eq_square power_even_eq) have 2: "0 < n div 2^w" by (simp add: assms(2) div_greater_zero_iff) have 3: "nat_to_words w n = nat_to_words_rec w (n div 2^w) [n mod 2^w]" by (metis 2 0 div_0 less_not_refl nat_to_words_def nat_to_words_rec.simps) show ?thesis using 0 1 2 3 by force qed lemma unfold_nat_to_words_rec: "nat_to_words_rec w n l = (nat_to_words_rec w n []) @ l" proof (cases "w=0") case True then show ?thesis by simp next case False then have [simp]: "0 < w" by blast have "\l. nat_to_words_rec w n l = nat_to_words_rec w n [] @ l" proof (induct n rule: less_induct) fix m assume ind: "!!j. j < m \ \ l. nat_to_words_rec w j l = nat_to_words_rec w j [] @ l" show "\l. nat_to_words_rec w m l = nat_to_words_rec w m [] @ l" proof fix l show "nat_to_words_rec w m l = nat_to_words_rec w m [] @ l" proof (cases "m < 0") assume "m < 0" thus ?thesis by blast next assume "~m < 0" show ?thesis proof (rule n_div_2tow_cases [of m]) assume [simp]: "m = 0" show ?thesis by simp next assume n2n: "m div 2^w < m" assume [simp]: "0 < m" hence n20: "0 \ m div 2^w" by arith from ind [of "m div 2^w"] and n2n n20 have ind': "\l. nat_to_words_rec w (m div 2^w) l = nat_to_words_rec w (m div 2^w) [] @ l" by (meson \0 < m\ div_less_dividend ind one_less_numeral_iff semiring_norm(76)) show ?thesis by (metis (no_types) append.assoc append_Cons append_Nil ind' nat_to_words_rec.simps) next show "0 < w" by simp qed qed qed qed thus ?thesis .. qed lemma nat_to_words_h: assumes "0 < n" "0 < w" shows "nat_to_words w n = (nat_to_words w (n div 2^w))@[n mod 2^w]" by (metis assms nat_to_words_def nat_to_words_rec.simps neq0_conv unfold_nat_to_words_rec) lemma nat_to_words_len_h: assumes "0 < n" "0 < w" shows "nat_to_words_len w n l = (nat_to_words_len w (n div 2^w) (l-1))@[n mod 2^w]" using assms nat_to_words_len_def nat_to_words_h by simp lemma nat_to_words_len_h2: assumes "0 < l" "0 < w" shows "nat_to_words_len w n l = (nat_to_words_len w (n div 2^w) (l-1))@[n mod 2^w]" proof (cases "n=0") case True then show ?thesis by (metis Suc_pred' append_Nil2 assms(1) div_mult2_eq le_0_eq mod_less_eq_dividend nat_mult_div_cancel_disj replicate.simps(2) replicate_app_Cons_same zero_to_words_len) next case False then show ?thesis by (meson assms(2) nat_to_words_len_h neq0_conv) qed lemma nat_to_words_div_len: assumes "0 < n" "0 < w" shows "length (nat_to_words w (n div 2^w)) + 1 = length (nat_to_words w n)" using assms nat_to_words_h by force lemma nat_bnd_word_len: assumes "n < (2^w)^m" shows "length (nat_to_words w n) \ m" proof (cases "w = 0") case True then show ?thesis by simp next case False then show ?thesis using assms proof (induct m arbitrary: n w) case 0 then have "n = 0" by simp then show ?case by auto next case 1: (Suc m) then show ?case proof (cases) assume "n = 0" then show ?case by simp next assume "\ n = 0" then have 2: "0 < n" by blast let ?nd = "n div 2^w" let ?nm = "n mod 2^w" have 3: "?nd < (2^w)^m" by (metis 1(3) less_mult_imp_div_less mult.commute power_Suc) have 4: "length (nat_to_words w ?nd) \ m" using 1(1,2) 3 by fastforce have 5: "nat_to_words w n = (nat_to_words w ?nd)@[?nm]" using 2 nat_to_words_h 1(2) by blast show ?case by (simp add: 4 5) qed qed qed lemma nat_bnd_word_len2: assumes "n < (2^w)^m" shows "word_length w n \ m" using assms word_length_def nat_bnd_word_len by presburger lemma nat_bnd_word_len_inv: assumes "m < length (nat_to_words w n)" shows "(2^w)^m \ n" using assms nat_bnd_word_len not_less by blast lemma nat_bnd_word_len_inv2: assumes "m < word_length w n" shows "(2^w)^m \ n" using assms word_length_def nat_bnd_word_len_inv by simp lemma nat_to_words_nil_iff_zero: assumes "0 < w" shows "length (nat_to_words w n) = 0 \ n = 0" using assms zero_to_word nat_to_words_h by fastforce lemma nat_to_words_nil_iff_zero2: assumes "0 < w" shows "word_length w n = 0 \ n = 0" using assms word_length_def nat_to_words_nil_iff_zero by presburger lemma nat_lowbnd_word_len: assumes "(2^w)^m \ n" "0 < w" shows "m < length (nat_to_words w n)" using assms proof (induction m arbitrary: n) case 0 then have "0 < n" by simp then show ?case using assms(2) nat_to_words_nil_iff_zero by simp next case 1: (Suc m) have 2: "0 < n" using 1 by fastforce let ?d = "n div 2^w" let ?m = "n mod 2^w" have 3: "(2^w)^m \ ?d" by (metis 1(2) div_greater_zero_iff div_mult2_eq power_Suc zero_less_numeral zero_less_power) have 4: "m < length (nat_to_words w ?d)" using 1(1) 3 assms(2) by blast have 5: "nat_to_words w n = (nat_to_words w ?d)@[?m]" using 2 nat_to_words_h assms(2) by blast show ?case using 4 5 by force qed lemma nat_lowbnd_word_len2: assumes "(2^w)^m \ n" "0 < w" shows "m < word_length w n" using assms word_length_def nat_lowbnd_word_len by presburger lemma word_length_eq: assumes "0 < n" "0 < w" shows "(word_length w n = l) = ((n < ((2^w)^l)) \ (((2^w)^(l-1)) \ n))" by (smt (verit, ccfv_threshold) One_nat_def Suc_eq_plus1 Suc_leI add.commute assms diff_less le_add_diff_inverse le_less less_le_trans nat_bnd_word_len2 nat_lowbnd_word_len2 not_le power_0 zero_less_one) lemma word_length_eq2: assumes "word_length w n = l" "0 < w" shows "n < (2^w)^l" by (metis assms less_nat_zero_code nat_neq_iff word_length_eq zero_less_numeral zero_less_power) lemma word_length_eq3: assumes "word_length w n = l" "0 < n" "0 < w" shows "(2^w)^(l-1) \ n" using assms word_length_eq by blast lemma word_length_not2pow: assumes "word_length w n = l" "0 < n" "0 < w" "\ (\ m. n = 2^m)" shows "(2^w)^(l-1) < n" by (metis assms nat_less_le power_mult word_length_eq3) lemma log2iffPow_h1: fixes n :: nat assumes "(log 2 n = \log 2 n\)" "0 < n" shows "(\ m. n = 2^m)" proof - let ?m = "nat \log 2 n\" have 0: "0 \ \log 2 n\" by (metis (no_types) assms(1,2) le_log2_of_power linorder_not_less of_int_0_le_iff of_nat_less_0_iff order_le_less_trans power_mult word_length_eq3) have 1: "\log 2 n\ = real ?m" using 0 by simp have 2: "log 2 n = real ?m" using assms(1) 1 by presburger have 3: "n = 2 ^ ?m" by (metis (full_types) assms(2) 2 less_log2_of_power log2_of_power_less nat_neq_iff of_nat_less_iff) show ?thesis using 3 by fast qed lemma log2iffPow_h2: assumes "(\ m. n = 2^m)" shows "(log 2 n = \log 2 n\)" using assms by force lemma log2iffPow: assumes "0 < (n::nat)" shows "(log 2 n = \log 2 n\) = (\ m. n = 2^m)" using assms log2iffPow_h1 by auto lemma word_length_not2pow': assumes "word_length w n = l" "0 < n" "0 < w" "log 2 n < \log 2 n\" shows "(2^w)^(l-1) < n" using assms log2iffPow[of n] word_length_not2pow[of w n l] by argo lemma oddNotPow2: assumes "odd (n::nat)" "1 < n" shows "\ (\ m. n = 2^m)" using assms(1,2) by fastforce lemma oddLog2Ceil: assumes "odd (n::nat)" "1 < n" shows "(log 2 n) < \(log 2 n)\" proof - have "log 2 n \ \log 2 n\" using assms oddNotPow2 log2iffPow by simp then show ?thesis by linarith qed lemma word_length_eq_odd: assumes "word_length w n = l" "0 < w" "odd n" "1 < n" shows "(2^w)^(l-1) < n" by (metis assms bot_nat_0.not_eq_extremum dvd_power nat_less_le odd_pos power_0 word_length_not2pow) text\We can write the word_length as a function of the log base 2. We only need that p is not a power of two. The lemmas for odd p or "p prime greater than 2" follow from this.\ lemma wordLenLog2NotPow: assumes "0 < w" "1 < p" "(log 2 p) < \(log 2 p)\" shows "word_length w p = \(log 2 p)/w\" proof - let ?l = "word_length w p" have r1: "p < (2^w)^?l" using word_length_eq2 assms(1) by blast have r2: "p < 2^(w*?l)" using r1 power_mult by metis have r3: "log 2 p < w*?l" using r2 assms(2) log2_of_power_less by blast have r4: "(log 2 p) / w < real(?l)" using r3 assms(1) by (simp add: mult.commute pos_divide_less_eq) have l1: "(2^w)^(?l-1) < p" using assms(1,2,3) word_length_not2pow' by auto have l2: "2^(w*(?l-1)) < p" using l1 power_mult by metis have l3: "w*(?l-1) < log 2 p" using l2 less_log2_of_power by fast have l4: "real (?l-1) < (log 2 p) / w" by (metis assms(1) l3 mult.commute mult_imp_less_div_pos of_nat_0_less_iff of_nat_mult) show ?thesis using r4 l4 by linarith qed lemma wordLenLog2Odd: assumes "odd p" "1 < p" "0 < w" shows "word_length w p = \(log 2 p)/w\" using assms(1,2,3) oddLog2Ceil wordLenLog2NotPow by blast lemma wordLenLog2Prime: assumes "prime p" "2 < p" "0 < w" shows "word_length w p = \(log 2 p)/w\" using assms(1,2,3) wordLenLog2Odd prime_gt_1_nat prime_odd_nat by blast lemma nth_word: assumes "i < length (nat_to_words w n)" "0 < w" shows "(nat_to_words w n) ! ((length (nat_to_words w n))-1-i) = (n div (2^w)^i) mod (2^w)" using assms proof (induction i arbitrary: n w) case 0 then show ?case proof (cases) assume "n = 0" then have "length (nat_to_words w n) = 0" using zero_to_word by blast then show ?thesis using 0 by blast next assume "\ n = 0" then have "0 < n" by blast then have "nat_to_words w n = (nat_to_words w (n div (2^w)))@[n mod (2^w)]" using nat_to_words_h 0(2) by blast then show ?thesis by fastforce qed next case 1: (Suc i) let ?l = "length (nat_to_words w n)" let ?d = "n div (2^w)" let ?m = "n mod (2^w)" have 11: "1 < ?l" using 1(2) by fastforce have 12: "0 < n" using 11 nat_bnd_word_len_inv by force have 13: "nat_to_words w n = nat_to_words w ?d @ [?m]" using 12 1(3) nat_to_words_h by blast let ?ld = "length (nat_to_words w ?d)" have 14: "?ld = ?l - 1" using 13 by fastforce have 15: "0 < ?ld" using 14 11 by force have 16: "i < ?ld" using 14 1(2) by linarith have 17: "?ld - 1 - i = ?l - 1 - (Suc i)" using 14 by fastforce have 18: "(nat_to_words w ?d) ! (?ld - 1 - i) = (?d div (2^w)^i) mod (2^w)" using 1(1,3) 16 by blast have 19: "nat_to_words w n ! (?l - 1 - Suc i) = (nat_to_words w ?d) ! (?ld - 1 - i)" by (metis 13 14 15 17 diff_less nth_append zero_less_Suc) have "?d div (2^w)^i = n div (2^w)^(Suc i)" by (simp add: div_mult2_eq) then show ?case using 18 19 by argo qed lemma nth_word2: assumes "j < length (nat_to_words w n)" "0 < w" shows "(nat_to_words w n) ! j = (n div (2^w)^(((length (nat_to_words w n))-1-j))) mod (2^w)" by (metis (no_types, lifting) One_nat_def Suc_pred add_diff_cancel_right' diff_less_Suc le_add_diff_inverse le_eq_less_or_eq length_greater_0_conv less_Suc_eq list.size(3) not_less_zero assms nth_word) lemma nth_word_len: assumes "i < length (nat_to_words_len w n l)" "0 < w" shows "(nat_to_words_len w n l) ! ((length (nat_to_words_len w n l))-1-i) = (n div (2^w)^i) mod (2^w)" proof (cases) assume "l \ word_length w n" then have "nat_to_words_len w n l = nat_to_words w n" by (simp add: nat_to_words_len_def word_length_def) then show ?thesis using assms nth_word by presburger next assume "\ (l \ word_length w n)" then have 1: "word_length w n < l" by simp let ?os = "nat_to_words w n" let ?ol = "length ?os" have 2: "?ol < l" using 1 word_length_def by auto have 3: "nat_to_words_len w n l = (replicate (l-?ol) 0) @ (nat_to_words w n)" using nat_to_words_len_def by presburger then show ?thesis proof (cases) assume 4: "i < ?ol" have "(nat_to_words_len w n l) ! ((length (nat_to_words_len w n l))-1-i) = (nat_to_words w n) ! ((length (nat_to_words w n))-1-i)" by (smt (z3) 2 3 4 One_nat_def Suc_pred add.commute assms diff_diff_cancel diff_right_commute le_add_diff_inverse length_append length_replicate less_imp_le_nat not_less_eq nth_append zero_less_diff diff_diff_left) then show ?thesis using nth_word 4 assms(2) by presburger next assume 5: "\(i < ?ol)" have 6: "(nat_to_words_len w n l) ! ((length (nat_to_words_len w n l))-1-i) = 0" by (metis 2 3 5 add.commute diff_Suc_eq_diff_pred diff_less_mono2 le_add_diff_inverse length_append length_replicate less_imp_le_nat not_less_eq nth_append nth_replicate) have 7: "n div (2^w)^i = 0" by (meson 5 assms(2) div_less le_less_linear nat_lowbnd_word_len) show ?thesis using 6 7 by presburger qed qed lemma nth_word2_len: assumes "j < length (nat_to_words_len w n l)" "0 < w" shows "(nat_to_words_len w n l) ! j = (n div (2^w)^(((length (nat_to_words_len w n l))-1-j))) mod (2^w)" by (smt (z3) assms nth_word_len One_nat_def Suc_leI Suc_pred add_diff_cancel_right' diff_less_Suc le_add_diff_inverse length_greater_0_conv less_nat_zero_code list.size(3) not_le) lemma words_upper_bound: assumes "os = nat_to_words w n" "b \ set os" shows "b < (2^w)" proof (cases "w = 0") case True then show ?thesis using assms by auto next case False obtain i where 0: "b = os ! i \ 0 \ i \ i < length os" by (metis le0 assms(2) in_set_conv_nth) let ?l = "length os" have "b = (n div (2^w)^(?l-1-i)) mod (2^w)" using 0 nth_word2 assms(1) False by blast then show ?thesis using mod_less_divisor False by simp qed lemma high_word_pos: assumes "0 < n" "0 < w" shows "0 < (nat_to_words w n) ! 0" proof - let ?os = "nat_to_words w n" let ?l = "length ?os" have 1: "0 < ?l" using assms nat_to_words_nil_iff_zero by auto have 2: "(nat_to_words w n) ! 0 = ( n div (2^w)^(?l-1) ) mod (2^w)" using 1 nth_word2 assms(2) by auto have 3: "n < (2^w)^?l" using word_length_def word_length_eq2 assms(2) by blast have 4: "(2^w)^(?l-1) \ n" using assms word_length_def word_length_eq by blast - have 5: "1 \ n div (2^w)^(?l-1)" - by (metis 4 bits_div_by_1 div_greater_zero_iff nat_neq_iff not_less_zero power_eq_0_iff - zero_neq_numeral) + have 5: "1 \ n div (2^w)^(?l-1)" + using 4 by (simp add: Suc_le_eq div_greater_zero_iff) have 6: "n div (2^w)^(?l-1) < (2^w)" by (metis 1 3 One_nat_def Suc_leI le_add_diff_inverse less_mult_imp_div_less power_add power_one_right) have 7: "( n div (2^w)^(?l-1) ) mod (2^w) = n div (2^w)^(?l-1)" using 5 6 by force show ?thesis using 2 5 7 by linarith qed lemma high_word_pos2: assumes "0 < n" "0 < w" shows "0 < hd (nat_to_words w n)" by (metis assms high_word_pos hd_Cons_tl less_not_refl3 list.size(3) nat_to_words_nil_iff_zero2 nth_Cons_0 word_length_def) lemma nat_to_words_valid: "words_valid w (nat_to_words w n)" by (metis words_upper_bound words_valid_def) lemma nat_to_words_valid_le: assumes "w1 \ w2" shows "words_valid w2 (nat_to_words w1 n)" by (smt (z3) assms le_less_trans linorder_not_less nat_power_less_imp_less nat_to_words_valid numeral_2_eq_2 words_valid_def zero_less_Suc) lemma nat_to_words_len_lowbnd: assumes "(2^w)^l \ n" "0 < w" shows "nat_to_words_len w n l = nat_to_words w n" by (simp add: assms less_imp_le_nat nat_lowbnd_word_len nat_to_words_len_def) lemma nat_to_words_len_upbnd: assumes "n < (2^w)^l" shows "length (nat_to_words_len w n l) = l" using assms nat_bnd_word_len nat_to_words_len_def by fastforce lemma word_len_mono: assumes "a \ b" shows "word_length w a \ word_length w b" by (metis assms gr_implies_not0 le_less_trans nat_bnd_word_len2 nat_to_zero_bit_word_len not_le not_less_iff_gr_or_eq word_length_eq2) lemma nat_to_word_len_mono: assumes "a < b" shows "length (nat_to_words_len w a l) \ length (nat_to_words_len w b l)" by (smt (z3) assms diff_Suc_1 le_eq_less_or_eq le_less_trans nat_bnd_word_len nat_to_words_len_len nat_to_words_len_lowbnd nat_to_words_len_upbnd nat_to_zero_bit_words_len2 neq0_conv not_le word_length_def word_length_eq2) lemma nat_to_word_len_mono': assumes "a < b" "word_length w b = l" "0 < w" shows "length (nat_to_words_len w a l) = l" by (metis assms less_imp_le_nat nat_to_words_len_upbnd order_le_less_trans word_length_eq2) lemma nat_to_words_len_valid: "words_valid w (nat_to_words_len w n l)" using nat_to_words_len_def nat_to_words_valid words_valid_concat words_valid_zeros by presburger lemma sum_to_words: assumes "y < (2^w)^l" "0 < x" "0 < w" shows "nat_to_words w (x*((2^w)^l) + y) = (nat_to_words w x) @ (nat_to_words_len w y l)" using assms proof (induction l arbitrary: y) case 0 then show ?case by force next case C: (Suc l) let ?yd = "y div (2^w)" let ?ym = "y mod (2^w)" have 1: "?yd < (2 ^ w) ^ l" by (metis C(2) less_mult_imp_div_less mult.commute power_Suc) let ?n = "x * (2 ^ w) ^ Suc l + y" let ?nd = "?n div 2^w" let ?nm = "?n mod 2^w" have 2: "nat_to_words w ?n = (nat_to_words w ?nd)@[?nm]" using nat_to_words_h assms(2,3) by auto have 3: "?nd = x * (2 ^ w) ^ l + ?yd" by (smt (z3) C.prems(1) div_mult_self3 gr_implies_not0 mult.commute mult.left_commute power_Suc power_eq_0_iff zero_less_Suc) have 4: "?nm = ?ym" by (metis add.commute mod_mult_self2 mult.left_commute power_Suc) have 5: "nat_to_words w ?nd = nat_to_words w x @ nat_to_words_len w ?yd l" using C(1) 3 assms(2,3) 1 by presburger have 6: "(nat_to_words_len w ?yd l)@[?nm] = nat_to_words_len w y (Suc l)" using 4 assms(3) diff_Suc_1 nat_to_words_len_h2 zero_less_Suc by presburger then show ?case by (metis 2 5 6 append_assoc) qed lemma sum_to_words_len: assumes "y < (2^w)^l" "0 < x" "0 < w" shows "nat_to_words_len w (x*((2^w)^l) + y) (k+l) = (nat_to_words_len w x k) @ (nat_to_words_len w y l)" using assms nat_to_words_len_def nat_to_words_len_upbnd sum_to_words by force lemma nat_to_words_len_bnd: assumes "y < (2^w)^l" shows "nat_to_words_len w y (k+l) = (replicate k 0) @ (nat_to_words_len w y l)" by (smt (z3) add_diff_cancel_right' append_assoc assms length_append length_replicate nat_to_words_len_def nat_to_words_len_upbnd replicate_add) lemma sum_to_words_len2: assumes "y < (2^w)^l" "0 < w" shows "nat_to_words_len w (x*((2^w)^l) + y) (k+l) = (nat_to_words_len w x k) @ (nat_to_words_len w y l)" proof (cases "x = 0") case T: True then show ?thesis by (simp add: assms(1) nat_to_words_len_bnd) next case False then have "0 < x" by fast then show ?thesis using assms sum_to_words_len by blast qed section\Words to Nat\ text\Convert any word string to a natural number. When the words are valid (< (2^w)), this operation is the inverse of nat_to_words.\ definition words_to_nat :: "nat \ words \ nat" where "words_to_nat w = foldl (%b a. (2^w)*b + a) 0" lemma one_word_to_nat [simp]: "words_to_nat w [n] = n" by (simp add: words_to_nat_def) lemma words_to_nat_empty [simp]: "words_to_nat w [] = 0" by (simp add: words_to_nat_def) lemma words_to_nat_cons [simp]: "words_to_nat w (b # bs) = b*((2^w)^ length bs) + words_to_nat w bs" proof - let ?words_to_nat' = "foldl (\bn b. (2^w)*bn + b)" have helper: "\base. ?words_to_nat' base bs = base * (2^w) ^ length bs + ?words_to_nat' 0 bs" proof (induct bs) case Nil show ?case by simp next case (Cons x xs base) show ?case apply (simp only: foldl_Cons) apply (subst Cons [of "(2^w)*base + x"]) apply simp apply (subst Cons [of "x"]) apply (simp add:add_mult_distrib) done qed show ?thesis by (simp add: words_to_nat_def) (rule helper) qed lemma words_to_nat_concat: "words_to_nat w (as @ bs) = (words_to_nat w as)*((2^w)^ length bs) + words_to_nat w bs" using words_to_nat_cons words_to_nat_def apply (induct as) apply (metis add.left_neutral append_Nil mult_0 words_to_nat_empty) by (metis (no_types, lifting) diff_zero foldl_Cons foldl_append le_add_diff_inverse mult_0_right zero_le) lemma words_to_nat_append: "words_to_nat w (as @ [a]) = (words_to_nat w as)*(2^w) + a" using words_to_nat_concat by force lemma words_to_nat_cons_zero: "words_to_nat w (0#bs) = words_to_nat w bs" using words_to_nat_cons by auto lemma words_to_nat_prepend_zeros: "words_to_nat w ((replicate z 0) @ bs) = words_to_nat w bs" apply (induction z) apply simp by simp lemma words_to_zero_intro: assumes "\iii (\ib\set bs. b < (2^w)" "0 < nth bs 0" "0 < w" shows "nat_to_words w (words_to_nat w bs) = bs" proof - let ?l = "length bs" show ?thesis using assms proof (induction ?l arbitrary: bs w) case 0 then have "bs = []" by auto then show ?case by simp next case 1: (Suc x) then show ?case proof (cases) assume 11: "Suc x = 1" have 12: "0 < length bs" using 1(2) 11 by force let ?c = "bs ! 0" have 13: "bs = [?c]" by (metis 1(2) 11 12 Cons_nth_drop_Suc One_nat_def Suc_leI drop0 drop_all) have 14: "?c \ set bs" using 12 in_set_conv_nth by auto have 15: "?c < (2^w)" using 14 1(3) by fast have 16: "words_to_nat w bs = ?c" using 13 one_word_to_nat by metis have 17: "nat_to_words w ?c = [?c]" using 1(4) 15 nat_to_one_word by presburger show ?case using 16 17 13 by auto next assume "\ Suc x = 1" then have 20: "1 < Suc x" by auto have 200: "0 < length bs" using 20 1(2) by auto let ?c = "last bs" let ?cs = "butlast bs" have 21: "bs = ?cs@[?c]" by (metis 1(2) append_butlast_last_id list.size(3) nat.simps(3)) have 22: "words_to_nat w bs = (words_to_nat w ?cs)*(2^w) + ?c" by (metis 21 words_to_nat_append) have 23: "x = length ?cs" using 1(2) 22 by auto have 24: "\b \ set ?cs. b < (2^w)" by (metis 1(3) in_set_butlastD) have 25: "0 < length ?cs" using 20 1(2) by simp have 26: "?cs ! 0 = bs ! 0" using 25 nth_butlast by blast have 27: "words_to_nat w bs = (words_to_nat w ?cs)*(2^w) + ?c" by (metis 21 words_to_nat_append) have 28: "nat_to_words w (words_to_nat w ?cs) = ?cs" using 1(1,4,5) 23 24 26 by presburger have 29: "?c \ set bs" using 200 by simp have 30: "?c < (2^w)" using 1(3) 29 by blast have 31: "(words_to_nat w bs) mod (2^w) = ?c" using 27 30 by simp have 32: "(words_to_nat w bs) div (2^w) = (words_to_nat w ?cs)" using 27 30 by simp have 33: "0 < words_to_nat w bs" by (metis 1(4) 200 words_to_zero gr0I) show ?case using 21 31 32 33 28 nat_to_words_h 1(5) by presburger qed qed qed lemma words_to_nat_len_bnd: assumes "\b\set bs. b < (2^w)" "0 < w" shows "words_to_nat w bs < (2^w)^(length bs)" using assms proof (induction bs arbitrary: w) case Nil then show ?case by fastforce next case 1: (Cons a bs) let ?l = "length bs" have 2: "\b\set bs. b < (2^w)" using 1(2) by simp have 3: "words_to_nat w bs < (2^w)^?l" using 1(1,3) 2 by blast show ?case by (metis 1(2,3) 3 diff_Suc_Suc diff_diff_cancel diff_le_self length_Cons linorder_not_less nat_bnd_word_len nat_lowbnd_word_len neq0_conv nth_Cons_0 words_to_nat_cons_zero words_to_nat_to_words verit_comp_simplify1(1)) qed lemma trunc_words_to_nat: assumes "as = nat_to_words w n" "as = bs@[b]" "0 < w" shows "words_to_nat w bs = (n div (2^w))" by (metis assms Nil_is_append_conv append1_eq_conv n_div_2tow_cases nat_to_words_h nat_to_words_to_nat not_Cons_self zero_to_word) lemma high_words_to_nat: assumes "as = nat_to_words w n" "as = bs@cs" "0 < w" shows "words_to_nat w bs = (n div (2^w)^(length cs))" proof - have 1: "words_to_nat w as = n" using assms(1,3) nat_to_words_to_nat by blast have 2: "words_to_nat w as = (words_to_nat w bs)*((2^w)^(length cs)) + words_to_nat w cs" using assms(2) words_to_nat_concat by simp have 3: "\c\set cs. c < (2^w)" by (metis assms(1,2) words_upper_bound append.assoc in_set_conv_decomp) have 4: "words_to_nat w cs < (2^w)^(length cs)" using 3 assms(3) words_to_nat_len_bnd by presburger have 5: "n = (words_to_nat w bs)*((2^w)^(length cs)) + words_to_nat w cs" using 1 2 by meson show ?thesis using 4 5 by simp qed lemma low_words_to_nat: assumes "as = nat_to_words w n" "as = bs@cs" "0 < w" shows "words_to_nat w cs = (n mod (2^w)^(length cs))" by (metis (no_types, lifting) add_left_cancel assms div_mult_mod_eq high_words_to_nat nat_to_words_to_nat words_to_nat_concat) lemma trunc_words_to_nat_len: assumes "as = nat_to_words_len w n l" "as = bs@[b]" "0 < w" shows "words_to_nat w bs = (n div (2^w))" proof (cases "n=0") case T: True then show ?thesis proof (cases "l=0") case True then show ?thesis using assms nat_to_words_len_def trunc_words_to_nat by force next case F0: False then have F1: "0 < l" by blast have F2: "as = replicate l 0" using assms(1) T zero_to_words_len by presburger have F3: "b = 0" by (metis F0 F2 assms(2) last_replicate last_snoc) have F4: "bs = replicate (l-1) 0" by (metis Cons_replicate_eq F1 F2 F3 append_eq_append_conv assms(2) replicate_append_same) have F5: "n div (2^w) = 0" by (simp add: T) show ?thesis by (metis F4 F5 append_Nil2 words_to_nat_empty words_to_nat_prepend_zeros) qed next case False then show ?thesis using assms nat_to_words_h nat_to_words_len_def trunc_words_to_nat words_to_nat_prepend_zeros by auto qed lemma low_words_to_nat2: assumes "as = nat_to_words_len w n l" "as = bs@cs" "0 < w" shows "words_to_nat w cs = (n mod (2^w)^(length cs))" proof (cases "n=0") case T: True then show ?thesis proof (cases "l=0") case True then show ?thesis using T assms(1,2) by force next case False then show ?thesis - by (metis T assms(1,2,3) bits_mod_0 length_0_conv nat_to_words_len_def - nat_to_words_nil_iff_zero2 word_length_def words_to_nat_concat words_to_nat_empty - words_to_nat_prepend_zeros zero_eq_add_iff_both_eq_0) + using T assms(1,2,3) by (auto intro!: words_to_zero_intro) + (metis append_assoc in_set_conv_decomp in_set_replicate list_update_id set_update_memI) qed next case F: False let ?xs = "nat_to_words w n" let ?xl = "length ?xs" let ?zl = "l - ?xl" let ?zs = "replicate ?zl 0" have 1: "as = ?zs @ ?xs" using assms(1) nat_to_words_len_def by meson have 2: "bs @ cs = ?zs @ ?xs" using 1 assms(2) by argo let ?cl = "length cs" let ?bl = "length bs" have 3: "?bl + ?cl = ?zl + ?xl" by (metis 2 length_append length_replicate) then show ?thesis proof (cases "?zl = 0") case True then show ?thesis using 1 assms(2,3) low_words_to_nat by force next case False then show ?thesis proof (cases "length cs \ ?xl") case T1: True let ?ds = "take (?xl - ?cl) ?xs" have "?xs = ?ds @ cs" by (smt (verit, ccfv_threshold) 1 assms(2) T1 append_assoc append_eq_append_conv append_take_drop_id diff_diff_cancel length_drop) then show ?thesis using assms(3) low_words_to_nat by blast next case F1: False have F2: "?xl < ?cl" using F1 by linarith have F3: "?bl < ?zl" using F2 3 by force have F4: "cs = drop ?bl as" using assms(2) by force let ?ds = "replicate (?cl - ?xl) 0" have "cs = ?ds @ ?xs" using 2 F3 F4 assms(2) by force then show ?thesis by (metis Euclidean_Rings.mod_less F1 assms(3) less_or_eq_imp_le nat_lowbnd_word_len nat_to_words_to_nat not_less words_to_nat_prepend_zeros) qed qed qed lemma words_to_nat_to_strip_words: assumes "\b\set bs. b < (2^w)" shows "nat_to_words w (words_to_nat w bs) = words_strip_zero bs" proof (cases "w = 0") case T: True have T1: "\b\set bs. b < 1" using T assms(1) by fastforce have T2: "\b\set bs. b = 0" using T1 by fastforce let ?bl = "length bs" have T3: "bs = replicate ?bl 0" by (simp add: T2 replicate_length_same) show ?thesis by (metis (no_types) T T3 append_Nil2 nat_to_zero_bit_words strip_prepend_zeros words_strip_zero.simps(1)) next case F: False let ?bl = "length bs" let ?s = "words_strip_zero bs" let ?sl = "length ?s" let ?z = "(replicate (?bl-?sl) 0)" let ?zl = "length ?z" have 1: "bs = ?z @ ?s" using strip_zero_concat by blast have 2: "words_to_nat w bs = words_to_nat w ?s" using 1 words_to_nat_prepend_zeros by metis have 3: "\b\set ?s. b\set bs" by (metis 1 append_assoc in_set_conv_decomp) have 4: "\b\set ?s. b<(2^w)" using assms(1) 3 by blast show ?thesis proof (cases) assume 50: "?s = []" then show ?thesis using 2 zero_to_word words_to_nat_empty by auto next assume 60:"?s \ []" have 61: "0 < ?s ! 0" using 60 strip_zero_head by fastforce have 62: "nat_to_words w (words_to_nat w ?s) = ?s" using 61 4 F words_to_nat_to_words by blast show ?thesis using 2 62 by simp qed qed lemma nat_to_words_len_to_nat: assumes "0 < w" shows "words_to_nat w (nat_to_words_len w x xLen) = x" using assms nat_to_words_len_def nat_to_words_to_nat words_to_nat_prepend_zeros by presburger lemma words_to_nat_to_words_len: assumes "l = length os" "\b\set os. b < (2^w)" shows "nat_to_words_len w (words_to_nat w os) l = os" using assms strip_zero_concat words_to_nat_to_strip_words nat_to_words_len_def by presburger lemma words_to_nat_to_words_len2: assumes "l = length os" "words_valid w os" shows "nat_to_words_len w (words_to_nat w os) l = os" using assms words_valid_def words_to_nat_to_words_len by auto lemma words_strip0_to_nat_len_bnd: assumes "\b\set bs. b < (2^w)" shows "words_to_nat w bs < (2^w)^(length (words_strip_zero bs))" proof (cases "w = 0") case T: True have T1: "\b\set bs. b < 1" using T assms(1) by fastforce have T2: "\b\set bs. b = 0" using T1 by fastforce let ?bl = "length bs" have T3: "bs = replicate ?bl 0" by (simp add: T2 replicate_length_same) then show ?thesis by (metis One_nat_def T less_Suc0 nth_replicate power_0 power_one words_to_zero) next case False then show ?thesis by (simp add: assms word_length_def word_length_eq2 words_to_nat_to_strip_words) qed lemma words_strip0_to_nat_len_bnd2: assumes "\b\set bs. b < (2^w)" "length (words_strip_zero bs) \ x" shows "words_to_nat w bs < (2^w)^x" by (meson assms leD leI le_less_trans nat_power_less_imp_less words_strip0_to_nat_len_bnd zero_less_numeral zero_less_power) lemma word_len_comp: assumes "word_length w a < word_length w b" shows "a < b" by (metis assms less_le_trans nat_bnd_word_len_inv2 word_length_eq2 le0 le_less nat_to_zero_bit_word_len) lemma word_len_shift: assumes "word_length w X = l" "0 < X" "0 < w" shows "word_length w (X*((2^w)^n)) = l + n" proof - have 0: "0 < l" using assms nat_to_words_nil_iff_zero2 neq0_conv by blast have l1: "(2^w)^(l-1) \ X" using assms word_length_eq by blast have l2: "(2^w)^(l-1)*(2^w)^n \ X*(2^w)^n" using l1 by force have l3: "(2^w)^(l-1+n) \ X*(2^w)^n" by (metis l2 power_add) have l4: "(2^w)^(l+n-1) \ X*(2^w)^n" using l3 0 by force have u1: "X < (2^w)^l" by (simp add: assms(1,3) word_length_eq2) have u2: "X*(2^w)^n < (2^w)^(n+l)" by (simp add: u1 power_add) show ?thesis by (metis assms(2,3) l4 u2 add.commute mult_pos_pos word_length_eq zero_less_numeral zero_less_power) qed lemma word_len_le_shift: assumes "word_length w X \ l" "0 < w" shows "word_length w (X*((2^w)^n)) \ l + n" by (metis add_le_mono1 assms mult_is_0 neq0_conv trans_le_add1 word_len_shift) lemma word_len_shift_add: assumes "word_length w X = l" "0 < X" "a < (2^w)^n" "0 < w" shows "word_length w (X*((2^w)^n) + a) = l + n" proof - have 1: "X*((2^w)^n) + a = X*((2^w)^n) OR a" by (metis assms(3) OR_sum_nat_hilo_2 mult.commute power_mult) have 2: "word_length w (X*((2^w)^n)) = l + n" using assms(1,2,4) word_len_shift by presburger have 3: "X*((2^w)^n) < (2^w)^(l+n)" by (simp add: 2 assms(4) word_length_eq2) have 4: "a < (2^w)^(n+l)" by (metis assms(3,4) less_le_trans add.right_neutral add_less_cancel_left less_nat_zero_code nat_bnd_word_len2 nat_lowbnd_word_len2 not_le) have 5: "X*((2^w)^n) OR a < (2^w)^(n+l)" by (metis 3 4 add.commute nat_OR_upper power_mult) have 6: "(2^w)^(l+n-1) \ X*((2^w)^n)" by (metis 2 assms(2,4) word_length_eq nat_0_less_mult_iff zero_less_numeral zero_less_power) have 7: "(2^w)^(l+n-1) \ X*((2^w)^n) + a" using 6 by linarith have 8: "X*((2^w)^n) + a < (2^w)^(l+n)" by (metis 1 5 add.commute) have 9: "0 < X*((2^w)^n) + a" using assms(2) by simp show ?thesis using 7 8 9 assms(4) word_length_eq add.commute by presburger qed lemma word_len_mult1: assumes "word_length w A = l" "word_length w B = k" "0 < A" shows "word_length w (A*B) \ l + k" by (smt (verit, ccfv_SIG) assms less_or_eq_imp_le mult_less_cancel1 nat_to_zero_bit_word_len neq0_conv not_less_iff_gr_or_eq word_len_comp word_len_shift word_length_eq2) lemma word_len_mult2: assumes "word_length w A = l" "word_length w B = k" "0 < A" "0 < B" "0 < w" shows "l + k - 1 \ word_length w (A*B)" proof - have A: "(2^w)^(l-1) \ A" using assms(1,3,5) word_length_eq by blast have B: "(2^w)^(k-1) \ B" using assms(2,4,5) word_length_eq by blast have k: "0 < k" using assms(2,4,5) nat_to_words_nil_iff_zero2 neq0_conv by blast have 1: "(2^w)^(l - 1 + k - 1) \ A*B" by (metis A B k power_add Nat.add_diff_assoc2 One_nat_def Suc_leI add.commute mult_le_mono) show ?thesis by (metis 1 A k Nat.add_diff_assoc2 One_nat_def Suc_leI Suc_pred add_gr_0 assms(1,5) diff_is_0_eq' nat_le_linear nat_lowbnd_word_len2) qed text \These conversions between natural numbers and words are injections. Mostly.\ lemma nat_to_words_inj: assumes "nat_to_words w n = nat_to_words w m" "0 < w" shows "n = m" by (metis assms nat_to_words_to_nat) lemma nat_to_words_len_inj: assumes "nat_to_words_len w n l = nat_to_words_len w m l" "0 < w" shows "n = m" by (metis assms nat_to_words_len_to_nat) lemma words_to_nat_inj: assumes "words_to_nat w as = words_to_nat w bs" "0 < w" "0 < hd as" "0 < hd bs" "words_valid w as" "words_valid w bs" shows "as = bs" by (metis assms words_to_nat_to_words words_valid_def hd_conv_nth words_strip_zero.simps(1) words_to_nat_to_strip_words) lemma words_to_nat_inj': assumes "words_to_nat w as = words_to_nat w bs" "words_valid w as" "words_valid w bs" shows "words_strip_zero as = words_strip_zero bs" by (metis assms nat_to_words_len_def strip_prepend_zeros words_to_nat_to_words_len2) section\XOR of Words\ text \In the schemes contained in the PKCS #1 standard, word strings are formed by converting from integers and/or concatenating other word strings. Then word strings of the same length may be XORed together. Here we define the XOR of words and prove the fundamental property that we will need in PKCS #1, which is that if A XOR B = C, then B XOR C = A. Note that we could define xor_words so that it pads the shorter input with zeroes so that the output is the length of the longer input. Since the PKCS1 standard only XORs words of the same length, this is not something that we concern ourselves with here.\ definition xor_words :: "words \ words \ words" where "xor_words xs ys = map2 (XOR) xs ys" lemma xor_words_length: assumes "length a \ length b" shows "length (xor_words a b) = length a" using assms xor_words_def by simp lemma xor_words_length2: "length (xor_words a b) = min (length a) (length b)" using xor_words_def by auto lemma xor_valid_words: assumes "words_valid w xs" "words_valid w ys" shows "words_valid w (xor_words xs ys)" proof - { fix i assume 0: "i < length (xor_words xs ys)" have 1: "(xor_words xs ys) ! i = (xs ! i) XOR (ys ! i)" using 0 xor_words_def by simp have 2: "i < length xs" using 0 xor_words_length2 by auto have 3: "i < length ys" using 0 xor_words_length2 by auto have 4: "xs ! i < (2^w)" using 2 assms(1) words_valid_def by force have 5: "ys ! i < (2^w)" using 3 assms(2) words_valid_def by force have 6: "(xs ! i) XOR (ys ! i) < (2^w)" using 4 5 nat_XOR_upper by blast have 7: "(xor_words xs ys) ! i < (2^w)" using 1 6 by auto } then show ?thesis by (metis in_set_conv_nth words_valid_def) qed lemma xor_words_symm: "xor_words xs ys = xor_words ys xs" using case_prod_unfold map2_map_map map_eq_conv zip_commute zip_map_fst_snd Bit_Operations.semiring_bit_operations_class.xor.commute xor_words_def by (smt (z3)) lemma xor_words_cons: "xor_words (a#as) (b#bs) = (a XOR b)#(xor_words as bs)" using xor_words_def by auto lemma xor_words_append: assumes "length as = length bs" shows "xor_words (as@[a]) (bs@[b]) = (xor_words as bs)@[a XOR b]" using assms xor_words_def by simp lemma nat_xor_inv: assumes "(a::nat) XOR b = c" shows "b XOR c = a" by (metis assms bit.xor_self xor.left_commute xor.right_neutral xor_nat_def) text\This is the main fact about the XOR of word strings that we will need to show that encoding and then decoding a message (for example in section 7 of PKCS #1) will get back to where you started.\ lemma words_xor_inv: assumes "xor_words as bs = cs" "length as = length bs" "length as = n" shows "xor_words bs cs = as" using assms xor_words_def apply (induction n arbitrary: as bs cs) apply auto by (smt (z3) Nil_is_map_conv hd_Cons_tl hd_zip length_Suc_conv length_map list.map(2) list.map_sel(1) list.sel(3) map_fst_zip map_snd_zip nat_xor_inv split_conv zip_map_fst_snd) lemma words_xor_inv2: assumes "length as = length bs" shows "xor_words (xor_words as bs) bs = as" using assms words_xor_inv xor_words_symm by presburger section \Abbreviations\ type_synonym octets = words type_synonym bits = words abbreviation bit_length :: "nat \ nat" where "bit_length n \ word_length 1 n" abbreviation nat_to_bits :: "nat \ words" where "nat_to_bits n \ nat_to_words 1 n" abbreviation nat_to_bits_len :: "nat \ nat \ words" where "nat_to_bits_len n l \ nat_to_words_len 1 n l" abbreviation bits_to_nat :: "words \ nat" where "bits_to_nat n \ words_to_nat 1 n" abbreviation bits_valid :: "words \ bool" where "bits_valid bs \ words_valid 1 bs" lemma bitLenLog2Pow2: assumes "n = 2^m" shows "bit_length n = \log 2 n\ + 1" proof - let ?b = "bit_length n" let ?l = "\log 2 (real n)\" have l1: "log 2 (real n) = m" using assms by simp have l2: "?l = m" using l1 by auto have b1: "?b = m + 1" by (simp add: assms word_length_eq) show ?thesis using l2 b1 by simp qed lemma bitLenLog2NotPow2: assumes "0 < n" "b = bit_length n" "2^(b-1) < n" shows "b = \log 2 n\" by (smt (z3) One_nat_def Suc_pred add.commute assms(1,2,3) ceiling_log_nat_eq_if less_or_eq_imp_le nat_to_words_nil_iff_zero2 neq0_conv of_nat_1 of_nat_add one_add_one plus_1_eq_Suc power_one_right word_length_eq2) lemma bitLenLog2Odd: assumes "odd p" "1 < p" shows "bit_length p = \log 2 p\" using assms(1,2) wordLenLog2Odd by force lemma bitLenLog2Prime: assumes "2 < p" "prime p" shows "bit_length p = \log 2 p\" using assms(1,2) bitLenLog2Odd prime_gt_1_nat prime_odd_nat by blast text \I use this fact in many places, so it's convenient to give it a name.\ lemma Twoto8 [simp]: "(2::nat)^8 = 256" by simp abbreviation octet_length :: "nat \ nat" where "octet_length n \ word_length 8 n" abbreviation nat_to_octets :: "nat \ words" where "nat_to_octets n \ nat_to_words 8 n" abbreviation nat_to_octets_len :: "nat \ nat \ words" where "nat_to_octets_len n l \ nat_to_words_len 8 n l" abbreviation octets_to_nat :: "words \ nat" where "octets_to_nat n \ words_to_nat 8 n" abbreviation octets_valid :: "words \ bool" where "octets_valid bs \ words_valid 8 bs" abbreviation xor_octets :: "words \ words \ words" where "xor_octets \ xor_words" lemma octetLenLog2Odd: assumes "odd p" "1 < p" shows "octet_length p = \(log 2 p)/8\" using assms(1,2) wordLenLog2Odd by auto lemma octetLenLog2Prime: assumes "prime p" "2 < p" shows "octet_length p = \(log 2 p)/8\" by (simp add: assms(1,2) wordLenLog2Prime) abbreviation word32_length :: "nat \ nat" where "word32_length n \ word_length 32 n" abbreviation nat_to_word32s :: "nat \ words" where "nat_to_word32s n \ nat_to_words 32 n" abbreviation word32s_to_nat :: "words \ nat" where "word32s_to_nat n \ words_to_nat 32 n" abbreviation word32s_valid :: "words \ bool" where "word32s_valid bs \ words_valid 32 bs" abbreviation word64_length :: "nat \ nat" where "word64_length n \ word_length 64 n" abbreviation nat_to_64words :: "nat \ words" where "nat_to_64words n \ nat_to_words 64 n" abbreviation word64s_to_nat :: "words \ nat" where "word64s_to_nat n \ words_to_nat 64 n" abbreviation word64s_valid :: "words \ bool" where "word64s_valid bs \ words_valid 64 bs" text\Some short-hands for converting between octets and bits.\ definition octets_to_bits :: "octets \ bits" where "octets_to_bits os = nat_to_bits (octets_to_nat os)" definition octets_to_bits_len :: "octets \ bits" where "octets_to_bits_len os = ( let l = 8*(length os) in nat_to_bits_len (octets_to_nat os) l )" definition bits_to_octets :: "bits \ octets" where "bits_to_octets bs = nat_to_octets (bits_to_nat bs)" definition bits_to_octets_len :: "bits \ octets" where "bits_to_octets_len bs = ( let l = length bs; x = (if l mod 8 = 0 then (l div 8) else (l div 8 + 1)) in nat_to_octets_len (bits_to_nat bs) x )" lemma octets_to_bits_to_nat: "(bits_to_nat (octets_to_bits os)) = (octets_to_nat os)" by (simp add: nat_to_words_to_nat octets_to_bits_def) lemma octets_to_bits_len_to_nat: "(bits_to_nat (octets_to_bits_len os)) = (octets_to_nat os)" by (simp add: nat_to_words_len_to_nat octets_to_bits_len_def) lemma bits_to_octets_to_nat: "(octets_to_nat (bits_to_octets bs)) = (bits_to_nat bs)" by (simp add: nat_to_words_to_nat bits_to_octets_def) lemma bits_to_octets_len_to_nat: "(octets_to_nat (bits_to_octets_len bs)) = (bits_to_nat bs)" by (metis bits_to_octets_len_def nat_to_words_len_to_nat zero_less_numeral) lemma octets_to_bits_len_len: assumes "octets_valid os" shows "length (octets_to_bits_len os) = 8 * length os" by (smt (verit, ccfv_SIG) assms nat_to_words_len_upbnd octets_to_bits_len_def power_mult power_one_right words_to_nat_len_bnd words_valid_def zero_less_numeral) lemma bits_to_octets_len_len: assumes "bits_valid bs" "l = length bs" "x = (if l mod 8 = 0 then (l div 8) else (l div 8 + 1))" shows "length (bits_to_octets_len bs) = x" proof - let ?n = "bits_to_nat bs" have 1: "?n < 2^l" by (metis assms(1,2) words_to_nat_len_bnd words_valid_def power_one_right zero_less_one) let ?m = "l mod 8" let ?d = "l div 8" have 2: "l = ?d*8 + ?m" using mod_div_decomp by presburger have 3: "?m < 8" by (meson mod_less_divisor zero_less_numeral) show ?thesis proof (cases "?m = 0") case True then show ?thesis by (smt (z3) 1 2 Nat.add_0_right assms(2,3) bits_to_octets_len_def mult.commute nat_to_words_len_upbnd power_mult) next case F0: False have F1: "x = ?d + 1" using F0 assms(3) by presburger have F2: "l < 8*x" using F1 2 3 by linarith have F3: "?n < (2^8)^x" by (metis (no_types, lifting) 1 F2 power_mult le_less_trans less_imp_le_nat one_less_numeral_iff power_strict_increasing_iff semiring_norm(76)) show ?thesis using F0 F1 F3 nat_to_words_len_upbnd assms(2) bits_to_octets_len_def by auto qed qed lemma octets_valid_to_bits_valid: assumes "octets_valid os" shows "bits_valid (octets_to_bits os)" using nat_to_words_valid octets_to_bits_def by presburger lemma octets_valid_to_bits_len_valid: assumes "octets_valid os" shows "bits_valid (octets_to_bits_len os)" using nat_to_words_len_valid octets_to_bits_len_def by auto lemma bits_valid_to_octets_valid: assumes "bits_valid bs" shows "octets_valid (bits_to_octets bs)" using bits_to_octets_def nat_to_words_valid by presburger lemma bits_valid_to_octets_len_valid: assumes "bits_valid bs" shows "octets_valid (bits_to_octets_len bs)" by (metis bits_to_octets_len_def nat_to_words_len_valid) section \Bit-Level Operations\ text\While most of the PKCS #1 standard only deals with octet-level (byte-level) operations, the signature scheme does some bit-level manipulations. We include the appropriate definitions and lemmas for those here.\ subsection \Set/Test High Bits\ definition numBits_to_numOctets :: "nat \ nat" where "numBits_to_numOctets numBits = (if (numBits mod 8 = 0) then (numBits div 8) else (numBits div 8 + 1))" lemma emLen_emBits: assumes "emLen = numBits_to_numOctets emBits" shows "(8*emLen - emBits < 8) \ (0 \ 8*emLen - emBits)" proof (cases "emBits mod 8 = 0") case True then have "emLen = emBits div 8" using assms(1) numBits_to_numOctets_def by presburger then have "8*emLen = emBits" using True by fastforce then show ?thesis by simp next case 1: False let ?m = "emBits mod 8" let ?d = "emBits div 8" have 2: "?m < 8 \ 0 < ?m" using 1 by fastforce have 3: "emBits = ?d*8 + ?m" by presburger have 4: "emLen = ?d + 1" using 1 assms(1) numBits_to_numOctets_def by presburger have 5: "8*emLen - emBits = (8*?d + 8) - (8*?d + ?m)" using 3 4 by fastforce have 6: "8*emLen - emBits = 8 - ?m" using 5 by linarith show ?thesis using 2 6 by simp qed definition SetLeftmostBits :: "nat \ nat \ octets \ octets" where "SetLeftmostBits emLen emBits OctetsIn = ( if OctetsIn = [] then [] else ( let x = 8*emLen - emBits; leftmostOctetIn = hd OctetsIn; leftmostOctetOut = leftmostOctetIn mod 2^(8-x) in leftmostOctetOut # (tl OctetsIn) ))" definition TestLeftmostBits:: "nat \ nat \ octets \ bool" where "TestLeftmostBits emLen emBits OctetsIn = ( if OctetsIn = [] then True else ( let leftmostOctet = hd OctetsIn; x = 8*emLen - emBits in leftmostOctet < 2^(8-x) ))" lemma SetLeftmostBits_valid: assumes "octets_valid X" shows "octets_valid (SetLeftmostBits a b X)" using assms proof (induction X) case Nil then show ?case using words_valid_nil SetLeftmostBits_def by force next case (Cons a X) have 1: "octets_valid X" by (metis list.sel(3) Cons(2) words_valid_tl) show ?case by (metis 1 Cons.prems SetLeftmostBits_def less_imp_diff_less list.sel(3) modulo_nat_def words_valid_cons words_valid_hd) qed lemma SetLeftmostBits_idem: assumes "Y = SetLeftmostBits a b X" shows "SetLeftmostBits a b Y = Y" using SetLeftmostBits_def assms by force lemma SetTestLeftmostBits: assumes "TestLeftmostBits a b Y" shows "SetLeftmostBits a b Y = Y" proof (cases "Y = []") case True then show ?thesis by (simp add: SetLeftmostBits_def) next case 1: False let ?h = "hd Y" let ?t = "tl Y" have 2: "Y = ?h # ?t" using 1 by simp let ?y = "8 - (8*a - b)" have 3: "?h < 2^?y" using assms 1 TestLeftmostBits_def by metis have 4: "?h mod 2^?y = ?h" using 3 by fastforce show ?thesis using 2 4 SetLeftmostBits_def by presburger qed lemma SetLeftmostBits_len: "length (SetLeftmostBits a b X) = length X" by (simp add: SetLeftmostBits_def) lemma SetLeftmostBits_hd: assumes "X \ []" shows "hd (SetLeftmostBits a b X) < 2^(8-(8*a - b))" using SetLeftmostBits_def assms by force lemma SetLeftmostBits_tl: "tl (SetLeftmostBits a b X) = tl X" using SetLeftmostBits_def by force lemma SetLeftmostBits_xor: assumes "length X = length Y" shows "SetLeftmostBits a b (xor_words X Y) = xor_words (SetLeftmostBits a b X) (SetLeftmostBits a b Y)" proof - let ?l = "length X" have "?l = length Y" using assms by blast then show ?thesis proof (induct ?l) case 0 then show ?case using SetLeftmostBits_def xor_words_def by auto next case 1: (Suc x) let ?Xh = "hd X" let ?Xt = "tl X" let ?Yh = "hd Y" let ?Yt = "tl Y" have x1: "X = ?Xh # ?Xt" by (metis 1(2) hd_Cons_tl list.size(3) nat.simps(3)) have y1: "Y = ?Yh # ?Yt" by (metis 1(2,3) hd_Cons_tl list.size(3) nat.simps(3)) let ?c = "8 - (8*a - b)" let ?sX = "SetLeftmostBits a b X" have x2: "tl ?sX = tl X" by (metis SetLeftmostBits_tl) have x3: "hd ?sX = (hd X) mod 2^?c" by (metis (no_types) SetLeftmostBits_def list.sel(1) list.simps(3) x1) let ?sY = "SetLeftmostBits a b Y" have y2: "tl ?sY = tl Y" by (metis SetLeftmostBits_tl) have y3: "hd ?sY = (hd Y) mod 2^?c" by (metis (no_types) SetLeftmostBits_def list.sel(1) list.simps(3) y1) let ?XxorY = "xor_words X Y" let ?sXxorY = "SetLeftmostBits a b (xor_words X Y)" have xy1: "tl ?XxorY = tl ?sXxorY" using SetLeftmostBits_def SetLeftmostBits_tl by presburger have xy2: "tl ?XxorY = xor_words (tl X) (tl Y)" by (metis list.sel(3) x1 xor_words_cons y1) have xy3: "tl ?sXxorY = xor_words (tl ?sX) (tl ?sY)" using xy1 xy2 x2 y2 by argo have xy4: "hd ?sXxorY = (?Xh XOR ?Yh) mod (2^?c)" by (smt (z3) SetLeftmostBits_def list.sel(1) list.simps(3) x1 xor_words_cons y1) have xy5: "(?Xh XOR ?Yh) mod (2^?c) = (?Xh mod (2^?c)) XOR (?Yh mod (2^?c))" by (metis take_bit_nat_def take_bit_xor) have xy6: "hd ?sXxorY = (hd ?sX) XOR (hd ?sY)" using xy4 xy5 x3 y3 by presburger show ?case using SetLeftmostBits_def xor_words_def xy2 xy3 xy6 zip_eq_Nil_iff by fastforce qed qed subsection \Bit Length\ text \Some of these lemmas are vestigial from the previous Octets.thy. Now they follow directly from lemmas proved above. But it's still nice to have them for the shorthand.\ lemma bit_len: "n < 2^(bit_length n)" by (metis less_numeral_extra(1) power_one_right word_length_eq2) lemma bit_len_zero: "bit_length 0 = 0" by (simp add: nat_to_words_nil_iff_zero2) lemma zero_bit_len: assumes "bit_length n = 0" shows "n = 0" by (metis assms bit_len less_one power_0) lemma bit_len_zero_eq: "(bit_length n = 0) = (n = 0)" using bit_len_zero zero_bit_len by blast lemma less_bit_len: assumes "x < bit_length n" shows "2^x \ n" by (metis (no_types) assms nat_bnd_word_len_inv2 power_one_right) lemma less_bit_len2: assumes "n < 2^x" shows "bit_length n \ x" by (meson assms less_bit_len not_less) lemma bit_len_exact1: assumes "0 < bit_length n" "bit_length n = l" shows "(n < (2^l)) \ ((2^(l-1)) \ n)" using assms bit_len less_bit_len by force lemma bit_len_exact2: assumes "0 < bit_length n" "n < 2^l" "2^(l-1) \ n" shows "bit_length n = l" by (metis assms bit_len_zero power_one_right word_len_comp word_length_eq zero_less_one) lemma bit_len_exact: assumes "0 < bit_length n" shows "(bit_length n = l) = ((n < (2^l)) \ ((2^(l-1)) \ n))" by (meson assms bit_len_exact1 bit_len_exact2) lemma bit_len_exact3: assumes "0 < n" shows "(bit_length n = l) = ((n < (2^l)) \ ((2^(l-1)) \ n))" by (metis assms bit_len_exact bit_len_zero_eq neq0_conv) lemma bit_len_lowbnd: assumes "2^l \ n" shows "l+1 \ bit_length n" by (metis Suc_eq_plus1 Suc_leI assms less_numeral_extra(1) nat_lowbnd_word_len2 power_one_right) lemma bitLen2octLen1: assumes "0 < n" "bl = bit_length n" "ol = octet_length n" shows "numBits_to_numOctets bl = ol" proof - have 2: "(2::nat)^8 = 256" by auto have bl: "2^(bl-1) \ n \ n < 2^bl" using assms(1,2) bit_len_exact3 by blast show ?thesis proof (cases "bl mod 8 = 0") case T0: True let ?x = "bl div 8" have T1: "8*?x = bl" using T0 by force have T2: "?x = numBits_to_numOctets bl" using T0 numBits_to_numOctets_def by presburger have T3: "(2::nat)^bl = (2^8)^?x" by (metis T1 power_mult) have T4: "(2::nat)^bl = 256^?x" using T3 by force have T5: "n < 256^?x" using T4 bl by presburger have T6: "(256::nat)^(?x - 1) = 2^(8*(?x - 1))" by (simp add: power_mult) have T7: "8*(?x - 1) = bl - 8" by (simp add: T1 right_diff_distrib') have T8: "2^(bl - 8) \ n" by (metis assms(1,2) diff_less less_bit_len n_div_2tow_cases zero_bit_len zero_less_numeral) have T9: "(256::nat)^(?x - 1) \ n" using T6 T7 T8 by presburger show ?thesis by (metis 2 T2 T5 T9 assms(1,3) word_length_eq zero_less_numeral) next case F0: False let ?x = "bl div 8 + 1" have F1: "?x = numBits_to_numOctets bl" using F0 numBits_to_numOctets_def by presburger let ?d = "bl div 8" let ?m = "bl mod 8" have F2: "0 < ?m \ ?m < 8" using F0 by force have F3: "bl = 8*?d + ?m" by presburger have F4: "8*?x = 8*?d + 8" by simp have F5: "8*?x > bl" using F2 F3 F4 by linarith have F6: "(256::nat)^?x = 2^(8*?x)" by (metis 2 power_mult) have F7: "(2::nat)^(8*?x) > 2^bl" using F5 by simp have F8: "(256::nat)^?x > n" using F6 F7 bl by linarith have F10: "(256::nat)^(?x-1) = 2^(8*?x - 8)" by (simp add: power_mult) have F11: "8*?x - 8 = 8*?d" using F4 by force have F12: "8*?x - 8 < bl" using F11 F3 F2 by presburger have F13: "(256::nat)^(?x-1) \ n" by (metis F10 F12 assms(2) less_bit_len) show ?thesis by (metis 2 F1 F8 F13 assms(1,3) word_length_eq zero_less_numeral) qed qed lemma bitLen2octLen2: assumes "0 < n" "bl = bit_length n" "numBits_to_numOctets bl = ol" shows "ol = octet_length n" using assms bitLen2octLen1 by blast lemma bit_len_comp: assumes "bit_length a < bit_length b" shows "a < b" using assms word_len_comp by blast lemma bit_len_head: assumes "bl = bit_length n" "ol = octet_length n" "os = nat_to_octets n" "0 < n" "hl = bit_length (hd os)" shows "bl = 8*(ol - 1) + hl" proof - have 2: "(256::nat) = 2^8" by fastforce have n1: "256^(ol-1) \ n \ n < 256^ol" by (metis 2 assms(2,4) word_length_eq zero_less_numeral) have n2: "2^(bl-1) \ n \ n < 2^bl" using assms(1,4) bit_len_exact3 by blast have ol1: "0 < ol" by (metis assms(4) neq0_conv less_one n1 power_0) let ?h = "hd os" have h1: "?h < 256" by (metis ol1 2 assms(2,3) list.size(3) nat_to_words_valid neq0_conv word_length_def words_valid_hd) have h2: "0 < ?h" using high_word_pos2 assms(3,4) by simp have h3: "?h < 2^hl \ 2^(hl-1) \ ?h" using h2 assms(5) bit_len_exact3 by blast have h4: "?h = n div 256^(ol-1)" by (smt (z3) One_nat_def Suc_leI cancel_comm_monoid_add_class.diff_cancel hd_Cons_tl le_add_diff_inverse less_mult_imp_div_less list.size(3) mod_less nat_lowbnd_word_len nat_to_words_nil_iff_zero2 neq0_conv nth_Cons_0 nth_word power_add power_one_right zero_less_numeral 2 assms(2,3) nth_word2 word_length_def ol1 h1 n1) have bo1: "ol = numBits_to_numOctets bl" using assms(1,2,4) bitLen2octLen1 by presburger have bo2: "8*(ol-1) \ bl - 1" proof (cases "bl mod 8 = 0") case True then show ?thesis using bo1 numBits_to_numOctets_def by force next case False then show ?thesis using bo1 numBits_to_numOctets_def by presburger qed have bo3: "8*(ol-1) < bl" by (metis bo2 le_less_trans less_imp_diff_less less_not_refl2 linorder_neqE_nat n2) have a11: "2^(bl - 1) div 256^(ol-1) \ n div 256^(ol-1)" using n2 div_le_mono by presburger have a12: "2^(bl-1) div 256^(ol-1) \ ?h" using a11 h4 by presburger have a13: "(2::nat)^(bl-1) div 256^(ol-1) = 2^(bl-1) div 2^(8*(ol-1))" by (metis 2 power_mult) have a14: "(2::nat)^(bl-1) div 2^(8*(ol-1)) = 2^((bl-1) - 8*(ol-1))" by (metis bo2 power_diff_power_eq zero_neq_numeral) let ?x = "bl - 8*(ol-1)" have a15: "2^(?x - 1) \ ?h" using a12 a13 a14 by force have a21: "n div 256^(ol-1) < 2^bl div 256^(ol-1)" by (smt (verit, ccfv_threshold) n2 2 div_greater_zero_iff div_le_mono le_add_diff_inverse less_mult_imp_div_less mult.commute n1 nat_less_le one_less_numeral_iff power_add power_diff power_increasing_iff power_mult semiring_norm(76) zero_less_numeral zero_less_power) have a22: "?h < 2^bl div 256^(ol-1)" using a21 h4 by presburger have a23: "(2::nat)^bl div 256^(ol-1) = 2^bl div 2^(8*(ol-1))" by (metis 2 power_mult) have a24: "(2::nat)^bl div 2^(8*(ol-1)) = 2^?x" by (metis bo3 less_imp_le_nat power_diff zero_neq_numeral) have a25: "?h < 2^?x" using a22 a23 a24 by argo have a: "?x = hl" by (metis bit_len_exact3 a15 a25 h3 h2) show ?thesis using a bo3 by fastforce qed lemma bit_len_bnd: assumes "bl = bit_length n" "ol = octet_length n" shows "bl \ 8*ol" proof (cases "ol = 0") case True then show ?thesis using assms(1,2) bit_len_zero_eq nat_to_words_nil_iff_zero2 by force next case False have "(256::nat) = 2^8" by simp then show ?thesis by (metis assms(1,2) less_bit_len2 word_length_eq2 power_mult zero_less_numeral) qed lemma bit_len_comp_head: assumes "octet_length a = octet_length b" "ha = hd (nat_to_octets a)" "hb = hd (nat_to_octets b)" "bit_length ha < bit_length hb" "0 < a" shows "a < b" proof - let ?ol = "octet_length a" let ?abl = "bit_length a" let ?bbl = "bit_length b" let ?habl = "bit_length ha" let ?hbbl = "bit_length hb" have 0: "0 < b" by (metis assms(1,5) nat_to_words_nil_iff_zero2 neq0_conv zero_less_numeral) have 1: "?bbl = 8*(?ol - 1) + ?hbbl" using 0 bit_len_head assms(1,3) by blast have 2: "?abl = 8*(?ol - 1) + ?habl" using assms(2,5) bit_len_head by blast show ?thesis using 1 2 assms(4) bit_len_comp by simp qed lemma octet_len_comp: assumes "octet_length a < octet_length b" shows "a < b" by (meson assms word_len_comp) lemma octet_len_comp_bit_len: assumes "octet_length a < octet_length b" shows "bit_length a < bit_length b" by (metis (mono_tags, opaque_lifting) assms bitLen2octLen1 bit_len_zero_eq less_nat_zero_code less_or_eq_imp_le linorder_neqE_nat not_less word_len_comp) lemma bit_len_head2: assumes "0 < length os" "hd os < 2^hl" "ol = length os" "n = octets_to_nat os" "octets_valid os" shows "bit_length n \ 8*(ol-1) + hl" proof (cases "hd os = 0") case T0: True let ?oss = "words_strip_zero os" have T1: "length ?oss < ol" by (metis T0 assms(1,3) diff_self_eq_0 drop_0 hd_conv_nth length_0_conv less_le strip_zero_drop strip_zero_head strip_zero_length) have T2: "octet_length n = length ?oss" by (metis assms(4,5) word_length_def words_to_nat_to_strip_words words_valid_def) have T3: "octet_length n \ ol-1" using T1 T2 by linarith show ?thesis by (metis T3 add.commute bit_len_bnd le_add2 le_trans mult_le_mono2) next case F0: False have F1: "nat_to_octets n = os" by (metis F0 assms(1,4,5) hd_conv_nth length_0_conv neq0_conv words_to_nat_to_words words_valid_def zero_less_numeral) have F2: "octet_length n = ol" using F1 assms(3) word_length_def by simp show ?thesis by (metis F1 F2 assms(2) bit_len_head bit_len_zero le_add2 le_add_same_cancel2 less_bit_len2 nat_add_left_cancel_le neq0_conv) qed lemma setleftmost_bit_len: assumes "Y = SetLeftmostBits a b X" "octets_valid X" "n = octets_to_nat Y" "X \ []" shows "bit_length n \ 8*((length Y) - 1) + (8 - (8*a - b))" by (metis SetLeftmostBits_def SetLeftmostBits_hd SetLeftmostBits_valid assms bit_len_head2 length_greater_0_conv list.simps(3)) lemma setleftmost_bit_len2: assumes "emLen = numBits_to_numOctets emBits" "emLen = length X" "X \ []" "octets_valid X" "Y = SetLeftmostBits emLen emBits X" "n = octets_to_nat Y" shows "bit_length n \ emBits" proof - have 0: "0 < emLen" using assms(2,3) by fast have 1: "length Y = emLen" by (simp add: SetLeftmostBits_len assms(2,5)) have 2: "bit_length n \ 8*(emLen - 1) + (8 - (8*emLen - emBits))" using 1 assms(3,4,5,6) setleftmost_bit_len by blast show ?thesis proof (cases "emBits mod 8 = 0") case T0: True have T1: "emLen = emBits div 8" using T0 assms(1) numBits_to_numOctets_def by presburger show ?thesis by (metis 2 Suc_leI T0 T1 add.commute assms(2,3) cancel_comm_monoid_add_class.diff_cancel diff_is_0_eq diff_zero dvd_mult_div_cancel le_add_diff_inverse length_greater_0_conv mod_0_imp_dvd mult_Suc_right plus_1_eq_Suc) next case F0: False let ?m = "emBits mod 8" let ?d = "emBits div 8" have F1: "emLen = ?d + 1" using F0 assms(1) numBits_to_numOctets_def by presburger have F2: "emBits = 8*?d + ?m" by simp have F3: "8*emLen - emBits = 8*?d + 8 - 8*?d - ?m" by (metis F1 F2 diff_diff_left distrib_left_numeral mult.right_neutral) have F4: "8*emLen - emBits = 8 - ?m" by (simp add: F3) have F5: "8 - (8*emLen - emBits) = ?m" by (simp add: F4) have F6: "8*(emLen - 1) + (8 - (8*emLen - emBits)) = emBits" using F1 F5 diff_add_inverse2 by presburger show ?thesis by (metis 2 F6) qed qed lemma setleftmost_bit_len3: assumes "emLen = numBits_to_numOctets emBits" "emLen = length Y" "Y \ []" "octets_valid Y" "TestLeftmostBits emLen emBits Y" "n = octets_to_nat Y" shows "bit_length n \ emBits" by (metis SetTestLeftmostBits assms setleftmost_bit_len2) lemma bit_len_shift: assumes "bit_length X = l" "0 < X" shows "bit_length (X*2^n) = l + n" by (metis assms less_numeral_extra(1) power_one_right word_len_shift) lemma bit_len_le_shift: assumes "bit_length X \ l" shows "bit_length (X*2^n) \ l + n" by (metis assms power_one_right word_len_le_shift zero_less_one) lemma bit_len_shift_add: assumes "bit_length X = l" "0 < X" "a < 2^n" shows "bit_length (X*2^n + a) = l + n" by (metis assms power_one_right word_len_shift_add zero_less_one) lemma bit_len_mult1: assumes "bit_length A = l" "bit_length B = k" "0 < A" shows "bit_length (A*B) \ l + k" using assms word_len_mult1 by blast lemma bit_len_mult2: assumes "bit_length A = l" "bit_length B = k" "0 < A" "0 < B" shows "l + k - 1 \ bit_length (A*B)" using assms word_len_mult2 by auto section \Word Conversions\ text \Different standards may be written for different word sizes. Here we define functions convenient for converting between word sizes.\ definition words_to_words :: "nat \ words \ nat \ words" where "words_to_words w1 xs w2 = nat_to_words w2 (words_to_nat w1 xs)" definition words_to_words_newLen :: "nat \ nat \ nat \ nat" where "words_to_words_newLen w1 l w2 = ( let x = w1*l in (if (w2 dvd x) then (x div w2) else ((x div w2)+1 )) )" definition words_to_words_len :: "nat \ words \ nat \ words" where "words_to_words_len w1 xs w2 = (let l = length xs; newLen = words_to_words_newLen w1 l w2 in nat_to_words_len w2 (words_to_nat w1 xs) newLen)" lemma ws2ws_newLen: assumes "w2 dvd w1" shows "words_to_words_newLen w1 l w2 = l*(w1 div w2)" using assms words_to_words_newLen_def by auto lemma ws2ws_nil [simp]: "words_to_words w1 [] w2 = []" by (simp add: words_to_words_def) lemma ws2wsnL0 [simp]: "words_to_words_newLen w1 0 w2 = 0" by (simp add: words_to_words_newLen_def) lemma ws2wsl_nil [simp]: "words_to_words_len w1 [] w2 = []" by (simp add: words_to_words_len_def) lemma ws2wsnL_same: assumes "0 < w" shows "words_to_words_newLen w l w = l" by (simp add: assms words_to_words_newLen_def) lemma ws2wsl_same: assumes "words_valid w xs" "0 < w" shows "words_to_words_len w xs w = xs" using assms words_to_nat_to_words_len2 words_to_words_len_def ws2wsnL_same by presburger lemma ws2wsl_len: assumes "words_valid w1 xs" "0 < w1" "0 < w2" shows "length (words_to_words_len w1 xs w2) = words_to_words_newLen w1 (length xs) w2" proof - let ?l = "length xs" let ?nL = "words_to_words_newLen w1 ?l w2" let ?xn = "words_to_nat w1 xs" have 1: "?xn < (2^w1)^?l" using assms(1,2) words_to_nat_len_bnd words_valid_def by blast let ?z = "w1*?l" let ?zm = "?z mod w2" let ?zd = "?z div w2" have 2: "?z = ?zd*w2 + ?zm" using div_mod_decomp by blast have 3: "?zm < w2" using assms(3) by force have 4: "?z \ ?nL * w2" proof (cases "w2 dvd ?z") case True then show ?thesis by (simp add: words_to_words_newLen_def) next case False then have "?nL = ?zd + 1" using words_to_words_newLen_def by presburger then have "?zd*w2 + w2 = ?nL * w2" by simp then have "?zd*w2 + ?zm < ?nL * w2" using 3 by linarith then show ?thesis using 2 by linarith qed have 5: "?xn < (2^w2)^?nL" by (metis 1 4 leD leI le_less_trans mult.commute nat_power_less_imp_less power_mult zero_less_numeral) have 6: "length (nat_to_words_len w2 ?xn ?nL) = ?nL" by (simp add: 5 nat_to_words_len_upbnd) show ?thesis using 6 words_to_words_len_def by auto qed lemma ws2wsl_concat: assumes "w2 dvd w1" "words_valid w1 ys" shows "words_to_words_len w1 (xs @ ys) w2 = (words_to_words_len w1 xs w2) @ (words_to_words_len w1 ys w2)" proof - let ?xl = "length xs" let ?xn = "words_to_nat w1 xs" let ?yl = "length ys" let ?yn = "words_to_nat w1 ys" let ?m = "w1 div w2" have 0: "?m*w2 = w1" using assms(1) by simp have 1: "words_to_nat w1 (xs@ys) = ?xn*((2^w1)^?yl) + ?yn" using words_to_nat_concat by fast have 10: "words_to_nat w1 (xs@ys) = ?xn*((2^w2)^(?m*?yl)) + ?yn" by (metis 0 1 mult.commute power_mult) let ?zn = "words_to_nat w1 (xs@ys)" have 2: "?yn < (2^w1)^?yl" using assms(2) strip_zero_length words_strip0_to_nat_len_bnd2 words_valid_def by blast have 20: "?yn < (2^w2)^(?m*?yl)" by (metis 0 2 mult.commute power_mult) have 3: "length (xs@ys) = ?xl + ?yl" by simp let ?newLen = "words_to_words_newLen w1 (?xl + ?yl) w2" have 30: "?newLen = ?m*?xl + ?m*?yl" using assms(1) ws2ws_newLen add_mult_distrib by force let ?newLenx = "words_to_words_newLen w1 ?xl w2" have 31: "?newLenx = ?m*?xl" by (simp add: assms(1) ws2ws_newLen) let ?newLeny = "words_to_words_newLen w1 ?yl w2" have 32: "?newLeny = ?m*?yl" by (simp add: assms(1) ws2ws_newLen) have 4: "nat_to_words_len w2 (words_to_nat w1 (xs@ys)) ?newLen = (nat_to_words_len w2 ?xn (?m*?xl)) @ (nat_to_words_len w2 ?yn (?m*?yl))" using 10 20 30 sum_to_words_len2 by fastforce show ?thesis using 3 4 31 32 words_to_words_len_def by presburger qed lemma ws2ws_append0: assumes "w2 dvd w1" "m = w1 div w2" "0 < w1" shows "words_to_words_len w1 (xs @ [0]) w2 = (words_to_words_len w1 xs w2) @ (replicate m 0)" proof - have "words_valid w1 [0]" by (simp add: words_valid_cons words_valid_nil) then show ?thesis by (metis One_nat_def Suc_eq_plus1 append.right_neutral assms(1,2) diff_zero list.size(3,4) mult.commute mult.right_neutral nat_to_words_len_def one_word_to_nat words_to_words_len_def ws2ws_newLen ws2wsl_concat zero_to_word) qed end diff --git a/thys/Native_Word/Word_Type_Copies.thy b/thys/Native_Word/Word_Type_Copies.thy --- a/thys/Native_Word/Word_Type_Copies.thy +++ b/thys/Native_Word/Word_Type_Copies.thy @@ -1,333 +1,338 @@ (* Title: Word_Type_Copies.thy Author: Florian Haftmann, TU Muenchen *) chapter \Systematic approach towards type copies of word type\ theory Word_Type_Copies imports "HOL-Library.Word" "Word_Lib.Most_significant_bit" "Word_Lib.Least_significant_bit" "Word_Lib.Generic_set_bit" "Word_Lib.Bit_Comprehension" "Code_Target_Word_Base" begin text \The lifting machinery is not localized, hence the abstract proofs are carried out using morphisms.\ locale word_type_copy = fixes of_word :: \'b::len word \ 'a\ and word_of :: \'a \ 'b word\ assumes type_definition: \type_definition word_of of_word UNIV\ begin lemma word_of_word: \word_of (of_word w) = w\ using type_definition by (simp add: type_definition_def) lemma of_word_of [code abstype]: \of_word (word_of p) = p\ \ \Use an abstract type for code generation to disable pattern matching on \<^term>\of_word\.\ using type_definition by (simp add: type_definition_def) lemma word_of_eqI: \p = q\ if \word_of p = word_of q\ proof - from that have \of_word (word_of p) = of_word (word_of q)\ by simp then show ?thesis by (simp add: of_word_of) qed lemma eq_iff_word_of: \p = q \ word_of p = word_of q\ by (auto intro: word_of_eqI) end bundle constraintless begin declaration \ let val cs = map (rpair NONE o fst o dest_Const) [\<^term>\0\, \<^term>\(+)\, \<^term>\uminus\, \<^term>\(-)\, \<^term>\1\, \<^term>\(*)\, \<^term>\(div)\, \<^term>\(mod)\, \<^term>\HOL.equal\, \<^term>\(\)\, \<^term>\(<)\, \<^term>\(dvd)\, \<^term>\of_bool\, \<^term>\numeral\, \<^term>\of_nat\, \<^term>\bit\, \<^term>\Bit_Operations.not\, \<^term>\Bit_Operations.and\, \<^term>\Bit_Operations.or\, \<^term>\Bit_Operations.xor\, \<^term>\mask\, \<^term>\push_bit\, \<^term>\drop_bit\, \<^term>\take_bit\, \<^term>\Bit_Operations.set_bit\, \<^term>\unset_bit\, \<^term>\flip_bit\, \<^term>\msb\, \<^term>\lsb\, \<^term>\size\, \<^term>\Generic_set_bit.set_bit\, \<^term>\set_bits\] in K (Context.mapping I (fold Proof_Context.add_const_constraint cs)) end \ end locale word_type_copy_ring = word_type_copy opening constraintless + constrains word_of :: \'a \ 'b::len word\ assumes word_of_0 [code]: \word_of 0 = 0\ and word_of_1 [code]: \word_of 1 = 1\ and word_of_add [code]: \word_of (p + q) = word_of p + word_of q\ and word_of_minus [code]: \word_of (- p) = - (word_of p)\ and word_of_diff [code]: \word_of (p - q) = word_of p - word_of q\ and word_of_mult [code]: \word_of (p * q) = word_of p * word_of q\ and word_of_div [code]: \word_of (p div q) = word_of p div word_of q\ and word_of_mod [code]: \word_of (p mod q) = word_of p mod word_of q\ and equal_iff_word_of [code]: \HOL.equal p q \ HOL.equal (word_of p) (word_of q)\ and less_eq_iff_word_of [code]: \p \ q \ word_of p \ word_of q\ and less_iff_word_of [code]: \p < q \ word_of p < word_of q\ begin lemma of_class_comm_ring_1: \OFCLASS('a, comm_ring_1_class)\ by standard (simp_all add: eq_iff_word_of word_of_0 word_of_1 word_of_add word_of_minus word_of_diff word_of_mult algebra_simps) lemma of_class_semiring_modulo: \OFCLASS('a, semiring_modulo_class)\ by standard (simp_all add: eq_iff_word_of word_of_0 word_of_1 word_of_add word_of_minus word_of_diff word_of_mult word_of_mod word_of_div algebra_simps mod_mult_div_eq) lemma of_class_equal: \OFCLASS('a, equal_class)\ by standard (simp add: eq_iff_word_of equal_iff_word_of equal) lemma of_class_linorder: \OFCLASS('a, linorder_class)\ by standard (auto simp add: eq_iff_word_of less_eq_iff_word_of less_iff_word_of) end locale word_type_copy_bits = word_type_copy_ring opening constraintless bit_operations_syntax + constrains word_of :: \'a::{comm_ring_1, semiring_modulo, equal, linorder} \ 'b::len word\ fixes signed_drop_bit :: \nat \ 'a \ 'a\ assumes bit_eq_word_of [code]: \bit p = bit (word_of p)\ and word_of_not [code]: \word_of (NOT p) = NOT (word_of p)\ and word_of_and [code]: \word_of (p AND q) = word_of p AND word_of q\ and word_of_or [code]: \word_of (p OR q) = word_of p OR word_of q\ and word_of_xor [code]: \word_of (p XOR q) = word_of p XOR word_of q\ and word_of_mask [code]: \word_of (mask n) = mask n\ and word_of_push_bit [code]: \word_of (push_bit n p) = push_bit n (word_of p)\ and word_of_drop_bit [code]: \word_of (drop_bit n p) = drop_bit n (word_of p)\ and word_of_signed_drop_bit [code]: \word_of (signed_drop_bit n p) = Word.signed_drop_bit n (word_of p)\ and word_of_take_bit [code]: \word_of (take_bit n p) = take_bit n (word_of p)\ and word_of_set_bit [code]: \word_of (Bit_Operations.set_bit n p) = Bit_Operations.set_bit n (word_of p)\ and word_of_unset_bit [code]: \word_of (unset_bit n p) = unset_bit n (word_of p)\ and word_of_flip_bit [code]: \word_of (flip_bit n p) = flip_bit n (word_of p)\ begin lemma word_of_bool: \word_of (of_bool n) = of_bool n\ by (simp add: word_of_0 word_of_1) lemma word_of_nat: \word_of (of_nat n) = of_nat n\ by (induction n) (simp_all add: word_of_0 word_of_1 word_of_add) lemma word_of_numeral [simp]: \word_of (numeral n) = numeral n\ proof - have \word_of (of_nat (numeral n)) = of_nat (numeral n)\ by (simp only: word_of_nat) then show ?thesis by simp qed lemma word_of_power: \word_of (p ^ n) = word_of p ^ n\ by (induction n) (simp_all add: word_of_1 word_of_mult) lemma even_iff_word_of: \2 dvd p \ 2 dvd word_of p\ (is \?P \ ?Q\) proof assume ?P then obtain q where \p = 2 * q\ .. then show ?Q by (simp add: word_of_mult) next assume ?Q then obtain w where \word_of p = 2 * w\ .. then have \of_word (word_of p) = of_word (2 * w)\ by simp then have \p = 2 * of_word w\ by (simp add: eq_iff_word_of word_of_word word_of_mult) then show ?P by simp qed lemma of_class_ring_bit_operations: \OFCLASS('a, ring_bit_operations_class)\ proof - have induct: \P p\ if stable: \(\p. p div 2 = p \ P p)\ and rec: \(\p b. P p \ (of_bool b + 2 * p) div 2 = p \ P (of_bool b + 2 * p))\ for p :: 'a and P proof - from stable have stable': \(\p. word_of p div 2 = word_of p \ P p)\ by (simp add: eq_iff_word_of word_of_div) from rec have rec': \(\p b. P p \ (of_bool b + 2 * word_of p) div 2 = word_of p \ P (of_bool b + 2 * p))\ by (simp add: eq_iff_word_of word_of_add word_of_bool word_of_mult word_of_div) define w where \w = word_of p\ then have \p = of_word w\ by (simp add: of_word_of) also have \P (of_word w)\ proof (induction w rule: bit_induct) case (stable w) show ?case by (rule stable') (simp add: word_of_word stable) next case (rec w b) have \P (of_bool b + 2 * of_word w)\ by (rule rec') (simp_all add: word_of_word rec) also have \of_bool b + 2 * of_word w = of_word (of_bool b + 2 * w)\ by (simp add: eq_iff_word_of word_of_word word_of_add word_of_1 word_of_mult word_of_0) finally show ?case . qed finally show \P p\ . qed have \class.semiring_parity_axioms (+) (0::'a) (*) 1 (mod)\ by standard (simp_all add: eq_iff_word_of word_of_0 word_of_1 even_iff_word_of word_of_mod even_iff_mod_2_eq_zero) with of_class_semiring_modulo have \OFCLASS('a, semiring_parity_class)\ by (rule semiring_parity_class.intro) + moreover have \OFCLASS('a, semiring_modulo_trivial_class)\ + apply standard + apply (simp_all only: eq_iff_word_of word_of_0 word_of_1 word_of_div) + apply simp_all + done moreover have \class.semiring_bits_axioms (+) (-) (0::'a) (*) 1 (div) (mod) bit\ apply standard apply (fact induct) apply (simp_all only: eq_iff_word_of word_of_0 word_of_1 word_of_bool word_of_numeral word_of_add word_of_diff word_of_mult word_of_div word_of_mod word_of_power bit_eq_word_of push_bit_take_bit drop_bit_take_bit even_iff_word_of fold_possible_bit flip: push_bit_eq_mult drop_bit_eq_div take_bit_eq_mod mask_eq_exp_minus_1 drop_bit_Suc) apply (simp_all add: bit_simps even_drop_bit_iff_not_bit not_less) done ultimately have \OFCLASS('a, semiring_bits_class)\ by (rule semiring_bits_class.intro) moreover have \class.semiring_bit_operations_axioms (+) (-) (0::'a) (*) (1::'a) (div) (mod) (AND) (OR) (XOR) mask Bit_Operations.set_bit unset_bit flip_bit push_bit drop_bit take_bit\ apply standard apply (simp_all add: eq_iff_word_of word_of_add word_of_push_bit word_of_power bit_eq_word_of word_of_and word_of_or word_of_xor word_of_mask word_of_diff word_of_0 word_of_1 bit_simps word_of_set_bit set_bit_eq_or word_of_unset_bit unset_bit_eq_or_xor word_of_flip_bit flip_bit_eq_xor word_of_mult word_of_drop_bit word_of_div word_of_take_bit word_of_mod and_rec [of \word_of a\ \word_of b\ for a b] or_rec [of \word_of a\ \word_of b\ for a b] xor_rec [of \word_of a\ \word_of b\ for a b] even_iff_word_of flip: mask_eq_exp_minus_1 push_bit_eq_mult drop_bit_eq_div take_bit_eq_mod) done ultimately have \OFCLASS('a, semiring_bit_operations_class)\ by (rule semiring_bit_operations_class.intro) moreover have \OFCLASS('a, ring_parity_class)\ using \OFCLASS('a, semiring_parity_class)\ by (rule ring_parity_class.intro) standard moreover have \class.ring_bit_operations_axioms (-) (1::'a) uminus NOT\ by standard (simp add: eq_iff_word_of word_of_not word_of_diff word_of_minus word_of_1 not_eq_complement) ultimately show \OFCLASS('a, ring_bit_operations_class)\ by (rule ring_bit_operations_class.intro) qed lemma [code]: \take_bit n a = a AND mask n\ for a :: 'a by (simp add: eq_iff_word_of word_of_take_bit word_of_and word_of_mask take_bit_eq_mask) lemma [code]: \mask (Suc n) = push_bit n (1 :: 'a) OR mask n\ \mask 0 = (0 :: 'a)\ by (simp_all add: eq_iff_word_of word_of_mask word_of_or word_of_push_bit word_of_0 word_of_1 mask_Suc_exp) lemma [code]: \Bit_Operations.set_bit n w = w OR push_bit n 1\ for w :: 'a by (simp add: eq_iff_word_of word_of_set_bit word_of_or word_of_push_bit word_of_1 set_bit_eq_or) lemma [code]: \unset_bit n w = w AND NOT (push_bit n 1)\ for w :: 'a by (simp add: eq_iff_word_of word_of_unset_bit word_of_and word_of_not word_of_push_bit word_of_1 unset_bit_eq_and_not) lemma [code]: \flip_bit n w = w XOR push_bit n 1\ for w :: 'a by (simp add: eq_iff_word_of word_of_flip_bit word_of_xor word_of_push_bit word_of_1 flip_bit_eq_xor) end locale word_type_copy_more = word_type_copy_bits + constrains word_of :: \'a::{ring_bit_operations, equal, linorder} \ 'b::len word\ fixes of_nat :: \nat \ 'a\ and nat_of :: \'a \ nat\ and of_int :: \int \ 'a\ and int_of :: \'a \ int\ and of_integer :: \integer \ 'a\ and integer_of :: \'a \ integer\ assumes word_of_nat_eq: \word_of (of_nat n) = word_of_nat n\ and nat_of_eq_word_of: \nat_of p = unat (word_of p)\ and word_of_int_eq: \word_of (of_int k) = word_of_int k\ and int_of_eq_word_of: \int_of p = uint (word_of p)\ and word_of_integer_eq: \word_of (of_integer l) = word_of_integer l\ and integer_of_eq_word_of: \integer_of p = unsigned (word_of p)\ begin lemma of_word_numeral [code_post]: \of_word (numeral n) = numeral n\ by (simp add: eq_iff_word_of word_of_word) lemma of_word_0 [code_post]: \of_word 0 = 0\ by (simp add: eq_iff_word_of word_of_0 word_of_word) lemma of_word_1 [code_post]: \of_word 1 = 1\ by (simp add: eq_iff_word_of word_of_1 word_of_word) text \Use pretty numerals from integer for pretty printing\ lemma numeral_eq_integer [code_unfold]: \numeral n = of_integer (numeral n)\ by (simp add: eq_iff_word_of word_of_integer_eq) lemma neg_numeral_eq_integer [code_unfold]: \- numeral n = of_integer (- numeral n)\ by (simp add: eq_iff_word_of word_of_integer_eq word_of_minus) end locale word_type_copy_misc = word_type_copy_more opening constraintless bit_operations_syntax + constrains word_of :: \'a::{ring_bit_operations, equal, linorder} \ 'b::len word\ fixes size :: nat and set_bits_aux :: \(nat \ bool) \ nat \ 'a \ 'a\ assumes size_eq_length: \size = LENGTH('b::len)\ and msb_iff_word_of [code]: \msb p \ msb (word_of p)\ and lsb_iff_word_of [code]: \lsb p \ lsb (word_of p)\ and size_eq_word_of: \Nat.size (p :: 'a) = Nat.size (word_of p)\ and word_of_generic_set_bit [code]: \word_of (Generic_set_bit.set_bit p n b) = Generic_set_bit.set_bit (word_of p) n b\ and word_of_set_bits: \word_of (set_bits P) = set_bits P\ and word_of_set_bits_aux: \word_of (set_bits_aux P n p) = Bit_Comprehension.set_bits_aux P n (word_of p)\ begin lemma size_eq [code]: \Nat.size p = size\ for p :: 'a by (simp add: size_eq_length size_eq_word_of word_size) lemma set_bits_aux_code [code]: \set_bits_aux f n w = (if n = 0 then w else let n' = n - 1 in set_bits_aux f n' (push_bit 1 w OR (if f n' then 1 else 0)))\ by (simp add: eq_iff_word_of word_of_set_bits_aux Let_def word_of_mult word_of_or word_of_0 word_of_1 set_bits_aux_rec [of f n]) lemma set_bits_code [code]: \set_bits P = set_bits_aux P size 0\ by (simp add: fun_eq_iff eq_iff_word_of word_of_set_bits word_of_set_bits_aux word_of_0 size_eq_length set_bits_conv_set_bits_aux) lemma of_class_lsb: \OFCLASS('a, lsb_class)\ by standard (simp add: fun_eq_iff lsb_iff_word_of even_iff_word_of lsb_odd) lemma of_class_set_bit: \OFCLASS('a, set_bit_class)\ by standard (simp add: eq_iff_word_of word_of_generic_set_bit bit_eq_word_of word_of_power word_of_0 bit_simps linorder_not_le) lemma of_class_bit_comprehension: \OFCLASS('a, bit_comprehension_class)\ by standard (simp add: eq_iff_word_of word_of_set_bits bit_eq_word_of set_bits_bit_eq) end end diff --git a/thys/Three_Squares/Residues_Properties.thy b/thys/Three_Squares/Residues_Properties.thy --- a/thys/Three_Squares/Residues_Properties.thy +++ b/thys/Three_Squares/Residues_Properties.thy @@ -1,381 +1,379 @@ (* Title: Three_Squares/Residues_Properties.thy Author: Anton Danilkin *) section \Properties of residues, congruences, quadratic residues and the Legendre symbol\ theory Residues_Properties imports "HOL-Number_Theory.Quadratic_Reciprocity" begin subsection \Properties of residues and congruences\ lemma mod_diff_eq_nat: fixes a b m :: nat assumes "a \ b" shows "(a - b) mod m = (m + (a mod m) - (b mod m)) mod m" proof cases assume "m = 0" thus ?thesis by auto next assume 0: "m \ 0" have "(a - b) mod m = nat (int (a - b) mod int m)" unfolding nat_mod_as_int by blast also have "... = nat ((int a - int b) mod int m)" using assms by (simp add: of_nat_diff) also have "... = nat ((int a mod int m - int b mod int m) mod int m)" using mod_diff_eq by metis also have "... = nat ((int a mod int m + (int m - int b mod int m)) mod int m)" by (metis add.left_commute add_uminus_conv_diff mod_add_self1) also have "... = nat ((int (nat (int a mod int m)) + (int m - int b mod int m)) mod int m)" by (metis nat_int of_nat_mod) also have "... = nat ((int (nat (int a mod int m)) + int (m - nat (int b mod int m))) mod int m)" by (metis 0 less_eq_nat.simps(1) mod_less_divisor nat_int nat_less_le of_nat_diff zmod_int) also have "... = nat (int (m + nat (int a mod int m) - nat (int b mod int m)) mod int m)" by (metis 0 Nat.add_diff_assoc add.commute bot_nat_0.not_eq_extremum mod_less_divisor nat_less_le nat_mod_as_int of_nat_add) also have "... = (m + (a mod m) - (b mod m)) mod m" unfolding nat_mod_as_int by blast finally show ?thesis . qed lemma prime_invertible_int: fixes a p :: int assumes "prime p" assumes "\ p dvd a" shows "\b. [a * b = 1] (mod p)" using assms coprime_commute coprime_iff_invertible_int prime_imp_coprime by blast lemma power_cong: fixes x y a m :: nat assumes "coprime a m" assumes "[x = y] (mod totient m)" shows "[a ^ x = a ^ y] (mod m)" proof - obtain k :: int where 0: "x + totient m * k = y" using assms by (metis cong_iff_lin cong_int_iff) show ?thesis proof cases assume "k \ 0" hence "x + totient m * nat k = y" using 0 by (metis int_eq_iff nat_int_add of_nat_mult) hence "[a ^ x * (a ^ totient m) ^ nat k = a ^ y] (mod m)" unfolding cong_def by (metis power_add power_mult) hence "[a ^ x * (a ^ totient m mod m) ^ nat k = a ^ y] (mod m)" unfolding cong_def by (metis mod_mult_right_eq power_mod) hence "[a ^ x * (1 mod m) ^ nat k = a ^ y] (mod m)" using euler_theorem[OF assms(1)] unfolding cong_def by argo hence "[a ^ x * 1 ^ nat k = a ^ y] (mod m)" unfolding cong_def by (metis mod_mult_right_eq power_mod) thus "[a ^ x = a ^ y] (mod m)" by auto next assume "\ k \ 0" hence "x = y + totient m * nat (- k)" using 0 by (smt (verit) int_nat_eq nat_int nat_minus_as_int of_nat_add of_nat_mult right_diff_distrib') hence "[a ^ x = a ^ y * (a ^ totient m) ^ nat (- k)] (mod m)" unfolding cong_def by (metis power_add power_mult) hence "[a ^ x = a ^ y * (a ^ totient m mod m) ^ nat (- k)] (mod m)" unfolding cong_def by (metis mod_mult_right_eq power_mod) hence "[a ^ x = a ^ y * (1 mod m) ^ nat (- k)] (mod m)" using euler_theorem[OF assms(1)] unfolding cong_def by argo hence "[a ^ x = a ^ y * 1 ^ nat (- k)] (mod m)" unfolding cong_def by (metis mod_mult_right_eq power_mod) thus "[a ^ x = a ^ y] (mod m)" by auto qed qed lemma power_cong_alt: fixes x a m :: nat assumes "coprime a m" shows "a ^ x mod m = a ^ (x mod totient m) mod m" using power_cong[OF assms] cong_def cong_mod_left by blast subsection \Properties of quadratic residues\ lemma QuadRes_cong: fixes a b p :: int assumes "[a = b] (mod p)" assumes "QuadRes p a" shows "QuadRes p b" using assms cong_trans unfolding QuadRes_def by blast lemma QuadRes_mult: fixes a b p :: int assumes "QuadRes p a" assumes "QuadRes p b" shows "QuadRes p (a * b)" using assms unfolding QuadRes_def by (metis cong_mult mult.assoc mult.commute power2_eq_square) lemma QuadRes_inv: fixes a b p :: int assumes "prime p" assumes "[a * b = 1] (mod p)" assumes "QuadRes p a" shows "QuadRes p b" proof - have 0: "\ p dvd a" using assms by (metis cong_dvd_iff dvd_mult2 not_prime_unit) obtain x where 1: "[x\<^sup>2 = a] (mod p)" using assms unfolding QuadRes_def by blast have "\ p dvd x" using 0 1 assms cong_dvd_iff pos2 prime_dvd_power_iff by blast then obtain y where "[x * y = 1] (mod p)" using assms prime_invertible_int by blast hence 2: "[(x * y)\<^sup>2 = 1] (mod p)" using cong_pow by fastforce have "[x\<^sup>2 * b = 1] (mod p)" using 1 assms cong_scalar_right cong_trans by blast hence "[y\<^sup>2 * (x\<^sup>2 * b) = y\<^sup>2 * 1] (mod p)" using cong_scalar_left by blast hence "[(x * y)\<^sup>2 * b = y\<^sup>2] (mod p)" by (simp add: algebra_simps) hence "[b = y\<^sup>2] (mod p)" using 2 by (metis cong_def cong_scalar_left mult.commute mult.right_neutral) hence "[y\<^sup>2 = b] (mod p)" by (rule cong_sym) thus ?thesis unfolding QuadRes_def by blast qed subsection \Properties of the Legendre symbol\ lemma Legendre_cong: fixes a b p :: int assumes "[a = b] (mod p)" shows "Legendre a p = Legendre b p" using assms QuadRes_cong[of a b p] QuadRes_cong[of b a p] unfolding Legendre_def cong_def by auto lemma Legendre_one: fixes p :: int assumes "p > 2" shows "Legendre 1 p = 1" using assms by (smt (verit) Legendre_def QuadRes_def cong_def cong_less_imp_eq_int one_power2) lemma Legendre_minus_one: fixes p :: int assumes "prime p" assumes "p > 2" shows "Legendre (- 1) p = 1 \ [p = 1] (mod 4)" proof - have "Legendre (- 1) p = 1 \ [Legendre (- 1) p = 1] (mod p)" using assms by (metis Legendre_def QuadRes_def cong_0_iff cong_def not_prime_unit one_power2) also have "... \ [(- 1) ^ ((nat p - 1) div 2) = 1] (mod p)" using assms euler_criterion[of "nat p" "- 1"] by (smt (verit) cong_def nat_0_le nat_one_as_int of_nat_add one_add_one prime_int_nat_transfer zless_nat_eq_int_zless) also have "... \ ((- 1 :: int) ^ ((nat p - 1) div 2)) = 1" using assms by (simp add: cong_iff_dvd_diff minus_one_power_iff zdvd_not_zless) also have "... \ even ((nat p - 1) div 2)" by (simp add: minus_one_power_iff) also have "... \ 4 dvd (nat p - 1)" using assms by (metis One_nat_def add_Suc_right div_dvd_div even_Suc even_diff_nat even_numeral even_of_nat group_cancel.rule0 nat_0_le numeral_Bit0_div_2 prime_int_nat_transfer prime_odd_int) also have "... \ [p = 1] (mod 4)" using assms unfolding cong_iff_dvd_diff int_dvd_int_iff[symmetric] by (simp add: int_ops) finally show ?thesis . qed lemma Legendre_minus_one_alt: fixes p :: int assumes "prime p" assumes "p > 2" shows "Legendre (- 1) p = (if [p = 1] (mod 4) then 1 else - 1)" using assms Legendre_minus_one[OF assms] unfolding Legendre_def cong_def by (auto simp add: zmod_minus1) lemma Legendre_two: fixes p :: int assumes "prime p" assumes "p > 2" shows "Legendre 2 p = 1 \ [p = 1] (mod 8) \ [p = 7] (mod 8)" proof - let ?n = "(p - 1) div 2 - (p - 1) div 4" have "odd p" using assms prime_odd_int by blast hence 0: "(\k. p = 8 * k + 1) \ (\k. p = 8 * k + 3) \ (\k. p = 8 * k + 5) \ (\k. p = 8 * k + 7)" by presburger have 1: "(j + 8 * k) div 4 = j div 4 + 2 * k" for k j :: int by linarith have 2: "GAUSS (nat p) 2" using assms unfolding GAUSS_def cong_def by auto have "Legendre 2 p = (- 1) ^ card (GAUSS.E (nat p) 2)" using assms GAUSS.gauss_lemma[OF 2] by auto also have "... = (- 1) ^ card ((\k. k mod p) ` (*) 2 ` {0<..(p - 1) div 2} \ {(p - 1) div 2<..})" unfolding GAUSS.E_def[OF 2] GAUSS.C_def[OF 2] GAUSS.B_def[OF 2] GAUSS.A_def[OF 2] by (simp add: algebra_simps) also have "... = (- 1) ^ card ((*) 2 ` {0<..(p - 1) div 2} \ {(p - 1) div 2<..})" by (rule arg_cong[of _ _ "\A. (- 1) ^ card (A \ {(p - 1) div 2<..})"]; force) also have "... = (- 1) ^ card {k \ (*) 2 ` {0<..(p - 1) div 2}. k > (p - 1) div 2}" by (rule arg_cong[of _ _ "\A. (- 1) ^ card A"]; blast) also have "... = (- 1) ^ card {k \ {0<..(p - 1) div 2}. 2 * k > (p - 1) div 2}" apply (rule arg_cong[of _ _ "\n. (- 1) ^ n"]) apply (rule card_bij_eq[where ?f = "\k. k div 2" and ?g = "(*) 2"]) subgoal unfolding inj_on_def by auto subgoal by auto subgoal by (simp add: inj_on_mult) subgoal by auto subgoal by (rule finite_Collect_conjI; auto) subgoal by (rule finite_Collect_conjI; auto) done also have "... = (- 1) ^ card {k \ {0<..(p - 1) div 2}. k > (p - 1) div 4}" by (rule arg_cong[of _ _ "\f. (- 1) ^ card {k \ {0<..(p - 1) div 2}. f k}"]; auto) also have "... = (- 1) ^ card {(p - 1) div 4<..(p - 1) div 2}" by (rule arg_cong[of _ _ "\A. (- 1) ^ card A"]; fastforce) also have "... = (- 1) ^ (nat ?n)" by auto finally have "Legendre 2 p = 1 \ even ?n" unfolding minus_one_power_iff by (simp add: assms even_nat_iff prime_gt_0_int zdiv_mono2) also have "... \ [p = 1] (mod 8) \ [p = 7] (mod 8)" unfolding cong_def using 0 by (auto simp add: 1) finally show ?thesis . qed lemma Legendre_two_alt: fixes p :: int assumes "prime p" assumes "p > 2" shows "Legendre 2 p = (if [p = 1] (mod 8) \ [p = 7] (mod 8) then 1 else - 1)" using assms Legendre_two[OF assms] unfolding Legendre_def cong_def by (auto simp add: zmod_minus1) lemma Legendre_mult: fixes a b p :: int assumes "prime p" shows "Legendre (a * b) p = Legendre a p * Legendre b p" proof cases assume 0: "p = 2" have 1: "QuadRes p = (\x. True)" using 0 by (metis QuadRes_def add_0 cong_iff_dvd_diff even_add odd_one power_one power_zero_numeral uminus_add_conv_diff) thus ?thesis using 0 unfolding 1 Legendre_def cong_0_iff by auto next assume "p \ 2" hence 2: "p > 2" using assms by (simp add: order_less_le prime_ge_2_int) have "[Legendre a p = a ^ ((nat p - 1) div 2)] (mod p)" "[Legendre b p = b ^ ((nat p - 1) div 2)] (mod p)" "[Legendre (a * b) p = (a * b) ^ ((nat p - 1) div 2)] (mod p)" using 2 assms euler_criterion[of "nat p" a] euler_criterion[of "nat p" b] euler_criterion[of "nat p" "a * b"] by auto hence 3: "[Legendre (a * b) p = Legendre a p * Legendre b p] (mod p)" by (smt (verit, best) cong_mult cong_sym cong_trans power_mult_distrib) have 4: "{Legendre a p, Legendre b p, Legendre (a * b) p} \ {-1, 0, 1}" unfolding Legendre_def by auto show ?thesis using 2 3 4 by (auto simp add: cong_iff_dvd_diff zdvd_not_zless) qed lemma Legendre_power: fixes a :: int fixes n :: nat fixes p :: int assumes "prime p" assumes "p > 2" shows "Legendre (a ^ n) p = (Legendre a p) ^ n" proof (induct n) case 0 thus ?case using assms Legendre_one by auto next case (Suc n) thus ?case using assms Legendre_mult by auto qed lemma Legendre_prod: fixes A :: "'a set" fixes f :: "'a \ int" fixes p :: int assumes "prime p" assumes "p > 2" shows "Legendre (prod f A) p = (\x\A. Legendre (f x) p)" proof (induct A rule: infinite_finite_induct) case (infinite A) thus ?case using assms Legendre_one by auto next case empty thus ?case using assms Legendre_one by auto next case (insert x F) thus ?case using assms Legendre_mult by auto qed lemma Legendre_equal: fixes p q :: int assumes "prime p" "prime q" assumes "p > 2" "q > 2" assumes "p \ q" assumes "[p = 1] (mod 4) \ [q = 1] (mod 4)" shows "Legendre p q = Legendre q p" proof - have 0: "even (p - 1)" "even (q - 1)" using assms prime_odd_int by auto have 1: "((p - 1) div 2) * ((q - 1) div 2) = (p - 1) * (q - 1) div 4" using 0 by fastforce have 2: "{Legendre p q, Legendre q p} \ {-1, 0, 1}" unfolding Legendre_def by auto have "Legendre p q * Legendre q p = (- 1) ^ nat (((p - 1) div 2) * ((q - 1) div 2))" using assms Quadratic_Reciprocity_int[of p q] by fastforce also have "... = (- 1) ^ nat ((p - 1) * (q - 1) div 4)" unfolding 1 by rule also have "... = 1" using 0 assms even_nat_iff unfolding minus_one_power_iff cong_iff_dvd_diff by auto finally show ?thesis using 2 by auto qed lemma Legendre_opposite: fixes p q :: int assumes "prime p" "prime q" assumes "p > 2" "q > 2" assumes "p \ q" assumes "[p = 3] (mod 4) \ [q = 3] (mod 4)" shows "Legendre p q = - Legendre q p" proof - - have 0: "even (p - 1)" "even (q - 1)" using assms prime_odd_int by auto - have 1: "((p - 1) div 2) * ((q - 1) div 2) = (p - 1) * (q - 1) div 4" - using 0 by fastforce - have "[p - 1 = 2] (mod 4) \ [q - 1 = 2] (mod 4)" - using assms - unfolding cong_iff_dvd_diff - by auto - hence "odd ((p - 1) * (q - 1) div 4)" - using assms 0 1 - by (metis bits_div_by_1 cong_dvd_iff dvd_div_iff_mult evenE even_mult_iff - even_numeral nonzero_mult_div_cancel_left numeral_One odd_one - zdiv_numeral_Bit0 zero_neq_numeral) + have \[p - 1 = 2] (mod 4)\ \[q - 1 = 2] (mod 4)\ + using \[p = 3] (mod 4) \ [q = 3] (mod 4)\ + by (simp_all add: cong_iff_dvd_diff) + from assms have \even (p - 1)\ \even (q - 1)\ + using assms by (auto dest: prime_odd_int) + then have *: "(p - 1) * (q - 1) div 4 = ((p - 1) div 2) * ((q - 1) div 2)" + by fastforce + have "odd ((p - 1) * (q - 1) div 4)" + using \even (p - 1)\ \even (q - 1)\ \[p - 1 = 2] (mod 4)\ \[q - 1 = 2] (mod 4)\ + by (auto elim!: oddE evenE dest: cong_dvd_iff simp add: *) hence 2: "odd (nat ((p - 1) * (q - 1) div 4))" using assms even_nat_iff pos_imp_zdiv_nonneg_iff by fastforce have 3: "{Legendre p q, Legendre q p} \ {-1, 0, 1}" unfolding Legendre_def by auto have "Legendre p q * Legendre q p = (- 1) ^ nat (((p - 1) div 2) * ((q - 1) div 2))" using assms Quadratic_Reciprocity_int[of p q] by fastforce - also have "... = (- 1) ^ nat ((p - 1) * (q - 1) div 4)" unfolding 1 by rule + also have "... = (- 1) ^ nat ((p - 1) * (q - 1) div 4)" unfolding * by rule also have "... = - 1" using 2 3 unfolding minus_one_power_iff by auto finally show ?thesis using 3 by auto qed end diff --git a/thys/Word_Lib/More_Word.thy b/thys/Word_Lib/More_Word.thy --- a/thys/Word_Lib/More_Word.thy +++ b/thys/Word_Lib/More_Word.thy @@ -1,2520 +1,2519 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section \Lemmas on words\ theory More_Word imports "HOL-Library.Word" More_Arithmetic More_Divides More_Bit_Ring begin context includes bit_operations_syntax begin \ \problem posed by TPHOLs referee: criterion for overflow of addition of signed integers\ lemma sofl_test: \sint x + sint y = sint (x + y) \ drop_bit (size x - 1) ((x + y XOR x) AND (x + y XOR y)) = 0\ for x y :: \'a::len word\ proof - obtain n where n: \LENGTH('a) = Suc n\ by (cases \LENGTH('a)\) simp_all have *: \sint x + sint y + 2 ^ Suc n > signed_take_bit n (sint x + sint y) \ sint x + sint y \ - (2 ^ n)\ \signed_take_bit n (sint x + sint y) > sint x + sint y - 2 ^ Suc n \ 2 ^ n > sint x + sint y\ using signed_take_bit_int_greater_eq [of \sint x + sint y\ n] signed_take_bit_int_less_eq [of n \sint x + sint y\] by (auto intro: ccontr) have \sint x + sint y = sint (x + y) \ (sint (x + y) < 0 \ sint x < 0) \ (sint (x + y) < 0 \ sint y < 0)\ using sint_less [of x] sint_greater_eq [of x] sint_less [of y] sint_greater_eq [of y] signed_take_bit_int_eq_self [of \LENGTH('a) - 1\ \sint x + sint y\] apply (auto simp add: not_less) apply (unfold sint_word_ariths) apply (subst signed_take_bit_int_eq_self) prefer 4 apply (subst signed_take_bit_int_eq_self) prefer 7 apply (subst signed_take_bit_int_eq_self) prefer 10 apply (subst signed_take_bit_int_eq_self) apply (auto simp add: signed_take_bit_int_eq_self signed_take_bit_eq_take_bit_minus take_bit_Suc_from_most n not_less intro!: *) done then show ?thesis apply (simp only: One_nat_def word_size drop_bit_eq_zero_iff_not_bit_last bit_and_iff bit_xor_iff) apply (simp add: bit_last_iff) done qed lemma unat_power_lower [simp]: "unat ((2::'a::len word) ^ n) = 2 ^ n" if "n < LENGTH('a::len)" using that by transfer simp lemma unat_p2: "n < LENGTH('a :: len) \ unat (2 ^ n :: 'a word) = 2 ^ n" by (fact unat_power_lower) lemma word_div_lt_eq_0: "x < y \ x div y = 0" for x :: "'a :: len word" by (fact div_word_less) lemma word_div_eq_1_iff: "n div m = 1 \ n \ m \ unat n < 2 * unat (m :: 'a :: len word)" apply (simp only: word_arith_nat_defs word_le_nat_alt word_of_nat_eq_iff flip: nat_div_eq_Suc_0_iff) apply (simp flip: unat_div unsigned_take_bit_eq) done lemma AND_twice [simp]: "(w AND m) AND m = w AND m" by (fact and.right_idem) lemma word_combine_masks: "w AND m = z \ w AND m' = z' \ w AND (m OR m') = (z OR z')" for w m m' z z' :: \'a::len word\ by (simp add: bit.conj_disj_distrib) lemma p2_gt_0: "(0 < (2 ^ n :: 'a :: len word)) = (n < LENGTH('a))" by (simp add : word_gt_0 not_le) lemma uint_2p_alt: \n < LENGTH('a::len) \ uint ((2::'a::len word) ^ n) = 2 ^ n\ using p2_gt_0 [of n, where ?'a = 'a] by (simp add: uint_2p) lemma p2_eq_0: \(2::'a::len word) ^ n = 0 \ LENGTH('a::len) \ n\ by (fact exp_eq_zero_iff) lemma p2len: \(2 :: 'a word) ^ LENGTH('a::len) = 0\ by (fact word_pow_0) lemma neg_mask_is_div: "w AND NOT (mask n) = (w div 2^n) * 2^n" for w :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps simp flip: push_bit_eq_mult drop_bit_eq_div) lemma neg_mask_is_div': "n < size w \ w AND NOT (mask n) = ((w div (2 ^ n)) * (2 ^ n))" for w :: \'a::len word\ by (rule neg_mask_is_div) lemma and_mask_arith: "w AND mask n = (w * 2^(size w - n)) div 2^(size w - n)" for w :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps word_size simp flip: push_bit_eq_mult drop_bit_eq_div) lemma and_mask_arith': "0 < n \ w AND mask n = (w * 2^(size w - n)) div 2^(size w - n)" for w :: \'a::len word\ by (rule and_mask_arith) lemma mask_2pm1: "mask n = 2 ^ n - (1 :: 'a::len word)" by (fact mask_eq_decr_exp) lemma add_mask_fold: "x + 2 ^ n - 1 = x + mask n" for x :: \'a::len word\ by (simp add: mask_eq_decr_exp) lemma word_and_mask_le_2pm1: "w AND mask n \ 2 ^ n - 1" for w :: \'a::len word\ by (simp add: mask_2pm1[symmetric] word_and_le1) lemma is_aligned_AND_less_0: "u AND mask n = 0 \ v < 2^n \ u AND v = 0" for u v :: \'a::len word\ apply (drule less_mask_eq) apply (simp flip: take_bit_eq_mask) apply (simp add: bit_eq_iff) apply (auto simp add: bit_simps) done lemma and_mask_eq_iff_le_mask: \w AND mask n = w \ w \ mask n\ for w :: \'a::len word\ apply (simp flip: take_bit_eq_mask) apply (cases \n \ LENGTH('a)\; transfer) apply (simp_all add: not_le min_def) apply (simp_all add: mask_eq_exp_minus_1) apply auto apply (metis take_bit_int_less_exp) apply (metis min_def nat_less_le take_bit_int_eq_self_iff take_bit_take_bit) done lemma less_eq_mask_iff_take_bit_eq_self: \w \ mask n \ take_bit n w = w\ for w :: \'a::len word\ by (simp add: and_mask_eq_iff_le_mask take_bit_eq_mask) lemma NOT_eq: "NOT (x :: 'a :: len word) = - x - 1" by (fact not_eq_complement) lemma NOT_mask: "NOT (mask n :: 'a::len word) = - (2 ^ n)" by (simp add : NOT_eq mask_2pm1) lemma le_m1_iff_lt: "(x > (0 :: 'a :: len word)) = ((y \ x - 1) = (y < x))" by uint_arith lemma gt0_iff_gem1: \0 < x \ x - 1 < x\ for x :: \'a::len word\ by (metis add.right_neutral diff_add_cancel less_irrefl measure_unat unat_arith_simps(2) word_neq_0_conv word_sub_less_iff) lemma power_2_ge_iff: \2 ^ n - (1 :: 'a::len word) < 2 ^ n \ n < LENGTH('a)\ using gt0_iff_gem1 p2_gt_0 by blast lemma le_mask_iff_lt_2n: "n < len_of TYPE ('a) = (((w :: 'a :: len word) \ mask n) = (w < 2 ^ n))" unfolding mask_2pm1 by (rule trans [OF p2_gt_0 [THEN sym] le_m1_iff_lt]) lemma mask_lt_2pn: \n < LENGTH('a) \ mask n < (2 :: 'a::len word) ^ n\ by (simp add: mask_eq_exp_minus_1 power_2_ge_iff) lemma word_unat_power: "(2 :: 'a :: len word) ^ n = of_nat (2 ^ n)" by simp lemma of_nat_mono_maybe: assumes xlt: "x < 2 ^ len_of TYPE ('a)" shows "y < x \ of_nat y < (of_nat x :: 'a :: len word)" apply (subst word_less_nat_alt) apply (subst unat_of_nat)+ apply (subst mod_less) apply (erule order_less_trans [OF _ xlt]) apply (subst mod_less [OF xlt]) apply assumption done lemma word_and_max_word: fixes a::"'a::len word" shows "x = - 1 \ a AND x = a" by simp lemma word_and_full_mask_simp: \x AND mask LENGTH('a) = x\ for x :: \'a::len word\ by (simp add: bit_eq_iff bit_simps) lemma of_int_uint: "of_int (uint x) = x" by (fact word_of_int_uint) corollary word_plus_and_or_coroll: "x AND y = 0 \ x + y = x OR y" for x y :: \'a::len word\ by (fact disjunctive_add2) corollary word_plus_and_or_coroll2: "(x AND w) + (x AND NOT w) = x" for x w :: \'a::len word\ by (fact and_plus_not_and) lemma unat_mask_eq: \unat (mask n :: 'a::len word) = mask (min LENGTH('a) n)\ by transfer (simp add: nat_mask_eq) lemma word_plus_mono_left: fixes x :: "'a :: len word" shows "\y \ z; x \ x + z\ \ y + x \ z + x" by unat_arith lemma less_Suc_unat_less_bound: "n < Suc (unat (x :: 'a :: len word)) \ n < 2 ^ LENGTH('a)" by (auto elim!: order_less_le_trans intro: Suc_leI) lemma up_ucast_inj: "\ ucast x = (ucast y::'b::len word); LENGTH('a) \ len_of TYPE ('b) \ \ x = (y::'a::len word)" by transfer (simp add: min_def split: if_splits) lemmas ucast_up_inj = up_ucast_inj lemma up_ucast_inj_eq: "LENGTH('a) \ len_of TYPE ('b) \ (ucast x = (ucast y::'b::len word)) = (x = (y::'a::len word))" by (fastforce dest: up_ucast_inj) lemma no_plus_overflow_neg: "(x :: 'a :: len word) < -y \ x \ x + y" by (metis diff_minus_eq_add less_imp_le sub_wrap_lt) lemma ucast_ucast_eq: "\ ucast x = (ucast (ucast y::'a word)::'c::len word); LENGTH('a) \ LENGTH('b); LENGTH('b) \ LENGTH('c) \ \ x = ucast y" for x :: "'a::len word" and y :: "'b::len word" apply transfer apply (cases \LENGTH('c) = LENGTH('a)\) apply (auto simp add: min_def) done lemma ucast_0_I: "x = 0 \ ucast x = 0" by simp lemma word_add_offset_less: fixes x :: "'a :: len word" assumes yv: "y < 2 ^ n" and xv: "x < 2 ^ m" and mnv: "sz < LENGTH('a :: len)" and xv': "x < 2 ^ (LENGTH('a :: len) - n)" and mn: "sz = m + n" shows "x * 2 ^ n + y < 2 ^ sz" proof (subst mn) from mnv mn have nv: "n < LENGTH('a)" and mv: "m < LENGTH('a)" by auto have uy: "unat y < 2 ^ n" by (rule order_less_le_trans [OF unat_mono [OF yv] order_eq_refl], rule unat_power_lower[OF nv]) have ux: "unat x < 2 ^ m" by (rule order_less_le_trans [OF unat_mono [OF xv] order_eq_refl], rule unat_power_lower[OF mv]) then show "x * 2 ^ n + y < 2 ^ (m + n)" using ux uy nv mnv xv' apply (subst word_less_nat_alt) apply (subst unat_word_ariths)+ apply (subst mod_less) apply simp apply (subst mult.commute) apply (rule nat_less_power_trans [OF _ order_less_imp_le [OF nv]]) apply (rule order_less_le_trans [OF unat_mono [OF xv']]) apply (cases "n = 0"; simp) apply (subst unat_power_lower[OF nv]) apply (subst mod_less) apply (erule order_less_le_trans [OF nat_add_offset_less], assumption) apply (rule mn) apply simp apply (simp add: mn mnv) apply (erule nat_add_offset_less; simp) done qed lemma word_less_power_trans: fixes n :: "'a :: len word" assumes nv: "n < 2 ^ (m - k)" and kv: "k \ m" and mv: "m < len_of TYPE ('a)" shows "2 ^ k * n < 2 ^ m" using nv kv mv apply - apply (subst word_less_nat_alt) apply (subst unat_word_ariths) apply (subst mod_less) apply simp apply (rule nat_less_power_trans) apply (erule order_less_trans [OF unat_mono]) apply simp apply simp apply simp apply (rule nat_less_power_trans) apply (subst unat_power_lower[where 'a = 'a, symmetric]) apply simp apply (erule unat_mono) apply simp done lemma word_less_power_trans2: fixes n :: "'a::len word" shows "\n < 2 ^ (m - k); k \ m; m < LENGTH('a)\ \ n * 2 ^ k < 2 ^ m" by (subst field_simps, rule word_less_power_trans) lemma Suc_unat_diff_1: fixes x :: "'a :: len word" assumes lt: "1 \ x" shows "Suc (unat (x - 1)) = unat x" proof - have "0 < unat x" by (rule order_less_le_trans [where y = 1], simp, subst unat_1 [symmetric], rule iffD1 [OF word_le_nat_alt lt]) then show ?thesis by ((subst unat_sub [OF lt])+, simp only: unat_1) qed lemma word_eq_unatI: \v = w\ if \unat v = unat w\ using that by transfer (simp add: nat_eq_iff) lemma word_div_sub: fixes x :: "'a :: len word" assumes yx: "y \ x" and y0: "0 < y" shows "(x - y) div y = x div y - 1" apply (rule word_eq_unatI) apply (subst unat_div) apply (subst unat_sub [OF yx]) apply (subst unat_sub) apply (subst word_le_nat_alt) apply (subst unat_div) apply (subst le_div_geq) apply (rule order_le_less_trans [OF _ unat_mono [OF y0]]) apply simp apply (subst word_le_nat_alt [symmetric], rule yx) apply simp apply (subst unat_div) apply (subst le_div_geq [OF _ iffD1 [OF word_le_nat_alt yx]]) apply (rule order_le_less_trans [OF _ unat_mono [OF y0]]) apply simp apply simp done lemma word_mult_less_mono1: fixes i :: "'a :: len word" assumes ij: "i < j" and knz: "0 < k" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "i * k < j * k" proof - from ij ujk knz have jk: "unat i * unat k < 2 ^ len_of TYPE ('a)" by (auto intro: order_less_subst2 simp: word_less_nat_alt elim: mult_less_mono1) then show ?thesis using ujk knz ij by (auto simp: word_less_nat_alt iffD1 [OF unat_mult_lem]) qed lemma word_mult_less_dest: fixes i :: "'a :: len word" assumes ij: "i * k < j * k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "i < j" using uik ujk ij by (auto simp: word_less_nat_alt iffD1 [OF unat_mult_lem] elim: mult_less_mono1) lemma word_mult_less_cancel: fixes k :: "'a :: len word" assumes knz: "0 < k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "(i * k < j * k) = (i < j)" by (rule iffI [OF word_mult_less_dest [OF _ uik ujk] word_mult_less_mono1 [OF _ knz ujk]]) lemma Suc_div_unat_helper: assumes szv: "sz < LENGTH('a :: len)" and usszv: "us \ sz" shows "2 ^ (sz - us) = Suc (unat (((2::'a :: len word) ^ sz - 1) div 2 ^ us))" proof - note usv = order_le_less_trans [OF usszv szv] from usszv obtain q where qv: "sz = us + q" by (auto simp: le_iff_add) have "Suc (unat (((2:: 'a word) ^ sz - 1) div 2 ^ us)) = (2 ^ us + unat ((2:: 'a word) ^ sz - 1)) div 2 ^ us" apply (subst unat_div unat_power_lower[OF usv])+ apply (subst div_add_self1, simp+) done also have "\ = ((2 ^ us - 1) + 2 ^ sz) div 2 ^ us" using szv by (simp add: unat_minus_one) also have "\ = 2 ^ q + ((2 ^ us - 1) div 2 ^ us)" apply (subst qv) apply (subst power_add) apply (subst div_mult_self2; simp) done also have "\ = 2 ^ (sz - us)" using qv by simp finally show ?thesis .. qed lemma enum_word_nth_eq: \(Enum.enum :: 'a::len word list) ! n = word_of_nat n\ if \n < 2 ^ LENGTH('a)\ for n using that by (simp add: enum_word_def) lemma length_enum_word_eq: \length (Enum.enum :: 'a::len word list) = 2 ^ LENGTH('a)\ by (simp add: enum_word_def) lemma unat_lt2p [iff]: \unat x < 2 ^ LENGTH('a)\ for x :: \'a::len word\ by transfer simp lemma of_nat_unat [simp]: "of_nat \ unat = id" by (rule ext, simp) lemma Suc_unat_minus_one [simp]: "x \ 0 \ Suc (unat (x - 1)) = unat x" by (metis Suc_diff_1 unat_gt_0 unat_minus_one) lemma word_add_le_dest: fixes i :: "'a :: len word" assumes le: "i + k \ j + k" and uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i \ j" using uik ujk le by (auto simp: word_le_nat_alt iffD1 [OF unat_add_lem] elim: add_le_mono1) lemma word_add_le_mono1: fixes i :: "'a :: len word" assumes ij: "i \ j" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i + k \ j + k" proof - from ij ujk have jk: "unat i + unat k < 2 ^ len_of TYPE ('a)" by (auto elim: order_le_less_subst2 simp: word_le_nat_alt elim: add_le_mono1) then show ?thesis using ujk ij by (auto simp: word_le_nat_alt iffD1 [OF unat_add_lem]) qed lemma word_add_le_mono2: fixes i :: "'a :: len word" shows "\i \ j; unat j + unat k < 2 ^ LENGTH('a)\ \ k + i \ k + j" by (subst field_simps, subst field_simps, erule (1) word_add_le_mono1) lemma word_add_le_iff: fixes i :: "'a :: len word" assumes uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "(i + k \ j + k) = (i \ j)" proof assume "i \ j" show "i + k \ j + k" by (rule word_add_le_mono1) fact+ next assume "i + k \ j + k" show "i \ j" by (rule word_add_le_dest) fact+ qed lemma word_add_less_mono1: fixes i :: "'a :: len word" assumes ij: "i < j" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i + k < j + k" proof - from ij ujk have jk: "unat i + unat k < 2 ^ len_of TYPE ('a)" by (auto elim: order_le_less_subst2 simp: word_less_nat_alt elim: add_less_mono1) then show ?thesis using ujk ij by (auto simp: word_less_nat_alt iffD1 [OF unat_add_lem]) qed lemma word_add_less_dest: fixes i :: "'a :: len word" assumes le: "i + k < j + k" and uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i < j" using uik ujk le by (auto simp: word_less_nat_alt iffD1 [OF unat_add_lem] elim: add_less_mono1) lemma word_add_less_iff: fixes i :: "'a :: len word" assumes uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "(i + k < j + k) = (i < j)" proof assume "i < j" show "i + k < j + k" by (rule word_add_less_mono1) fact+ next assume "i + k < j + k" show "i < j" by (rule word_add_less_dest) fact+ qed lemma word_mult_less_iff: fixes i :: "'a :: len word" assumes knz: "0 < k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "(i * k < j * k) = (i < j)" using assms by (rule word_mult_less_cancel) lemma word_le_imp_diff_le: fixes n :: "'a::len word" shows "\k \ n; n \ m\ \ n - k \ m" by (auto simp: unat_sub word_le_nat_alt) lemma word_less_imp_diff_less: fixes n :: "'a::len word" shows "\k \ n; n < m\ \ n - k < m" by (clarsimp simp: unat_sub word_less_nat_alt intro!: less_imp_diff_less) lemma word_mult_le_mono1: fixes i :: "'a :: len word" assumes ij: "i \ j" and knz: "0 < k" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "i * k \ j * k" proof - from ij ujk knz have jk: "unat i * unat k < 2 ^ len_of TYPE ('a)" by (auto elim: order_le_less_subst2 simp: word_le_nat_alt elim: mult_le_mono1) then show ?thesis using ujk knz ij by (auto simp: word_le_nat_alt iffD1 [OF unat_mult_lem]) qed lemma word_mult_le_iff: fixes i :: "'a :: len word" assumes knz: "0 < k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "(i * k \ j * k) = (i \ j)" proof assume "i \ j" show "i * k \ j * k" by (rule word_mult_le_mono1) fact+ next assume p: "i * k \ j * k" have "0 < unat k" using knz by (simp add: word_less_nat_alt) then show "i \ j" using p by (clarsimp simp: word_le_nat_alt iffD1 [OF unat_mult_lem uik] iffD1 [OF unat_mult_lem ujk]) qed lemma word_diff_less: fixes n :: "'a :: len word" shows "\0 < n; 0 < m; n \ m\ \ m - n < m" apply (subst word_less_nat_alt) apply (subst unat_sub) apply assumption apply (rule diff_less) apply (simp_all add: word_less_nat_alt) done lemma word_add_increasing: fixes x :: "'a :: len word" shows "\ p + w \ x; p \ p + w \ \ p \ x" by unat_arith lemma word_random: fixes x :: "'a :: len word" shows "\ p \ p + x'; x \ x' \ \ p \ p + x" by unat_arith lemma word_sub_mono: "\ a \ c; d \ b; a - b \ a; c - d \ c \ \ (a - b) \ (c - d :: 'a :: len word)" by unat_arith lemma power_not_zero: "n < LENGTH('a::len) \ (2 :: 'a word) ^ n \ 0" by (metis p2_gt_0 word_neq_0_conv) lemma word_gt_a_gt_0: "a < n \ (0 :: 'a::len word) < n" apply (case_tac "n = 0") apply clarsimp apply (clarsimp simp: word_neq_0_conv) done lemma word_power_less_1 [simp]: "sz < LENGTH('a::len) \ (2::'a word) ^ sz - 1 < 2 ^ sz" apply (simp add: word_less_nat_alt) apply (subst unat_minus_one) apply simp_all done lemma word_sub_1_le: "x \ 0 \ x - 1 \ (x :: 'a :: len word)" apply (subst no_ulen_sub) apply simp apply (cases "uint x = 0") apply (simp add: uint_0_iff) apply (insert uint_ge_0[where x=x]) apply arith done lemma push_bit_word_eq_nonzero: \push_bit n w \ 0\ if \w < 2 ^ m\ \m + n < LENGTH('a)\ \w \ 0\ for w :: \'a::len word\ using that apply (simp only: word_neq_0_conv word_less_nat_alt mod_0 unat_word_ariths unat_power_lower word_le_nat_alt) apply (metis add_diff_cancel_right' gr0I gr_implies_not0 less_or_eq_imp_le min_def push_bit_eq_0_iff take_bit_nat_eq_self_iff take_bit_push_bit take_bit_take_bit unsigned_push_bit_eq) done lemma unat_less_power: fixes k :: "'a::len word" assumes szv: "sz < LENGTH('a)" and kv: "k < 2 ^ sz" shows "unat k < 2 ^ sz" using szv unat_mono [OF kv] by simp lemma unat_mult_power_lem: assumes kv: "k < 2 ^ (LENGTH('a::len) - sz)" shows "unat (2 ^ sz * of_nat k :: (('a::len) word)) = 2 ^ sz * k" proof (cases \sz < LENGTH('a)\) case True with assms show ?thesis by (simp add: unat_word_ariths take_bit_eq_mod mod_simps unsigned_of_nat) (simp add: take_bit_nat_eq_self_iff nat_less_power_trans flip: take_bit_eq_mod) next case False with assms show ?thesis by simp qed lemma word_plus_mcs_4: "\v + x \ w + x; x \ v + x\ \ v \ (w::'a::len word)" by uint_arith lemma word_plus_mcs_3: "\v \ w; x \ w + x\ \ v + x \ w + (x::'a::len word)" by unat_arith lemma word_le_minus_one_leq: "x < y \ x \ y - 1" for x :: "'a :: len word" by transfer (metis le_less_trans less_irrefl take_bit_decr_eq take_bit_nonnegative zle_diff1_eq) lemma word_less_sub_le[simp]: fixes x :: "'a :: len word" assumes nv: "n < LENGTH('a)" shows "(x \ 2 ^ n - 1) = (x < 2 ^ n)" using le_less_trans word_le_minus_one_leq nv power_2_ge_iff by blast lemma unat_of_nat_len: "x < 2 ^ LENGTH('a) \ unat (of_nat x :: 'a::len word) = x" by (simp add: unsigned_of_nat take_bit_nat_eq_self_iff) lemma unat_of_nat_eq: "x < 2 ^ LENGTH('a) \ unat (of_nat x ::'a::len word) = x" by (rule unat_of_nat_len) lemma unat_eq_of_nat: "n < 2 ^ LENGTH('a) \ (unat (x :: 'a::len word) = n) = (x = of_nat n)" by transfer (auto simp add: take_bit_of_nat nat_eq_iff take_bit_nat_eq_self_iff intro: sym) lemma alignUp_div_helper: fixes a :: "'a::len word" assumes kv: "k < 2 ^ (LENGTH('a) - n)" and xk: "x = 2 ^ n * of_nat k" and le: "a \ x" and sz: "n < LENGTH('a)" and anz: "a mod 2 ^ n \ 0" shows "a div 2 ^ n < of_nat k" proof - have kn: "unat (of_nat k :: 'a word) * unat ((2::'a word) ^ n) < 2 ^ LENGTH('a)" using xk kv sz apply (subst unat_of_nat_eq) apply (erule order_less_le_trans) apply simp apply (subst unat_power_lower, simp) apply (subst mult.commute) apply (rule nat_less_power_trans) apply simp apply simp done have "unat a div 2 ^ n * 2 ^ n \ unat a" proof - have "unat a = unat a div 2 ^ n * 2 ^ n + unat a mod 2 ^ n" by (simp add: div_mult_mod_eq) also have "\ \ unat a div 2 ^ n * 2 ^ n" using sz anz by (simp add: unat_arith_simps) finally show ?thesis .. qed then have "a div 2 ^ n * 2 ^ n < a" using sz anz apply (subst word_less_nat_alt) apply (subst unat_word_ariths) apply (subst unat_div) apply simp apply (rule order_le_less_trans [OF mod_less_eq_dividend]) apply (erule order_le_neq_trans [OF div_mult_le]) done also from xk le have "\ \ of_nat k * 2 ^ n" by (simp add: field_simps) finally show ?thesis using sz kv apply - apply (erule word_mult_less_dest [OF _ _ kn]) apply (simp add: unat_div) apply (rule order_le_less_trans [OF div_mult_le]) apply (rule unat_lt2p) done qed lemma mask_out_sub_mask: "(x AND NOT (mask n)) = x - (x AND (mask n))" for x :: \'a::len word\ by (fact and_not_eq_minus_and) lemma subtract_mask: "p - (p AND mask n) = (p AND NOT (mask n))" "p - (p AND NOT (mask n)) = (p AND mask n)" for p :: \'a::len word\ by (auto simp: and_not_eq_minus_and) lemma take_bit_word_eq_self_iff: \take_bit n w = w \ n \ LENGTH('a) \ w < 2 ^ n\ for w :: \'a::len word\ using take_bit_int_eq_self_iff [of n \take_bit LENGTH('a) (uint w)\] by (transfer fixing: n) auto lemma word_power_increasing: assumes x: "2 ^ x < (2 ^ y::'a::len word)" "x < LENGTH('a)" "y < LENGTH('a)" shows "x < y" using x using assms by transfer simp lemma mask_twice: "(x AND mask n) AND mask m = x AND mask (min m n)" for x :: \'a::len word\ by (simp flip: take_bit_eq_mask) lemma plus_one_helper[elim!]: "x < n + (1 :: 'a :: len word) \ x \ n" apply (simp add: word_less_nat_alt word_le_nat_alt field_simps) apply (case_tac "1 + n = 0") apply simp_all apply (subst(asm) unatSuc, assumption) apply arith done lemma plus_one_helper2: "\ x \ n; n + 1 \ 0 \ \ x < n + (1 :: 'a :: len word)" by (simp add: word_less_nat_alt word_le_nat_alt field_simps unatSuc) lemma less_x_plus_1: "x \ - 1 \ (y < x + 1) = (y < x \ y = x)" for x :: "'a::len word" by (meson max_word_wrap plus_one_helper plus_one_helper2 word_le_less_eq) lemma word_Suc_leq: fixes k::"'a::len word" shows "k \ - 1 \ x < k + 1 \ x \ k" using less_x_plus_1 word_le_less_eq by auto lemma word_Suc_le: fixes k::"'a::len word" shows "x \ - 1 \ x + 1 \ k \ x < k" by (meson not_less word_Suc_leq) lemma word_lessThan_Suc_atMost: \{..< k + 1} = {..k}\ if \k \ - 1\ for k :: \'a::len word\ using that by (simp add: lessThan_def atMost_def word_Suc_leq) lemma word_atLeastLessThan_Suc_atLeastAtMost: \{l ..< u + 1} = {l..u}\ if \u \ - 1\ for l :: \'a::len word\ using that by (simp add: atLeastAtMost_def atLeastLessThan_def word_lessThan_Suc_atMost) lemma word_atLeastAtMost_Suc_greaterThanAtMost: \{m<..u} = {m + 1..u}\ if \m \ - 1\ for m :: \'a::len word\ using that by (simp add: greaterThanAtMost_def greaterThan_def atLeastAtMost_def atLeast_def word_Suc_le) lemma word_atLeastLessThan_Suc_atLeastAtMost_union: fixes l::"'a::len word" assumes "m \ - 1" and "l \ m" and "m \ u" shows "{l..m} \ {m+1..u} = {l..u}" proof - from ivl_disj_un_two(8)[OF assms(2) assms(3)] have "{l..u} = {l..m} \ {m<..u}" by blast with assms show ?thesis by(simp add: word_atLeastAtMost_Suc_greaterThanAtMost) qed lemma max_word_less_eq_iff [simp]: \- 1 \ w \ w = - 1\ for w :: \'a::len word\ by (fact word_order.extremum_unique) lemma word_or_zero: "(a OR b = 0) = (a = 0 \ b = 0)" for a b :: \'a::len word\ by (fact or_eq_0_iff) lemma word_2p_mult_inc: assumes x: "2 * 2 ^ n < (2::'a::len word) * 2 ^ m" assumes suc_n: "Suc n < LENGTH('a::len)" shows "2^n < (2::'a::len word)^m" by (smt suc_n le_less_trans lessI nat_less_le nat_mult_less_cancel_disj p2_gt_0 power_Suc power_Suc unat_power_lower word_less_nat_alt x) lemma power_overflow: "n \ LENGTH('a) \ 2 ^ n = (0 :: 'a::len word)" by simp lemmas extra_sle_sless_unfolds [simp] = word_sle_eq[where a=0 and b=1] word_sle_eq[where a=0 and b="numeral n"] word_sle_eq[where a=1 and b=0] word_sle_eq[where a=1 and b="numeral n"] word_sle_eq[where a="numeral n" and b=0] word_sle_eq[where a="numeral n" and b=1] word_sless_alt[where a=0 and b=1] word_sless_alt[where a=0 and b="numeral n"] word_sless_alt[where a=1 and b=0] word_sless_alt[where a=1 and b="numeral n"] word_sless_alt[where a="numeral n" and b=0] word_sless_alt[where a="numeral n" and b=1] for n lemma word_sint_1: "sint (1::'a::len word) = (if LENGTH('a) = 1 then -1 else 1)" by (fact signed_1) lemma ucast_of_nat: "is_down (ucast :: 'a :: len word \ 'b :: len word) \ ucast (of_nat n :: 'a word) = (of_nat n :: 'b word)" by transfer simp lemma scast_1': "(scast (1::'a::len word) :: 'b::len word) = (word_of_int (signed_take_bit (LENGTH('a::len) - Suc 0) (1::int)))" by transfer simp lemma scast_1: "(scast (1::'a::len word) :: 'b::len word) = (if LENGTH('a) = 1 then -1 else 1)" by (fact signed_1) lemma unat_minus_one_word: "unat (-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1" by (simp add: mask_eq_exp_minus_1 unsigned_minus_1_eq_mask) lemmas word_diff_ls'' = word_diff_ls [where xa=x and x=x for x] lemmas word_diff_ls' = word_diff_ls'' [simplified] lemmas word_l_diffs' = word_l_diffs [where xa=x and x=x for x] lemmas word_l_diffs = word_l_diffs' [simplified] lemma two_power_increasing: "\ n \ m; m < LENGTH('a) \ \ (2 :: 'a :: len word) ^ n \ 2 ^ m" by (simp add: word_le_nat_alt) lemma word_leq_le_minus_one: "\ x \ y; x \ 0 \ \ x - 1 < (y :: 'a :: len word)" apply (simp add: word_less_nat_alt word_le_nat_alt) apply (subst unat_minus_one) apply assumption apply (cases "unat x") apply (simp add: unat_eq_zero) apply arith done lemma neg_mask_combine: "NOT(mask a) AND NOT(mask b) = NOT(mask (max a b) :: 'a::len word)" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma neg_mask_twice: "x AND NOT(mask n) AND NOT(mask m) = x AND NOT(mask (max n m))" for x :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps) lemma multiple_mask_trivia: "n \ m \ (x AND NOT(mask n)) + (x AND mask n AND NOT(mask m)) = x AND NOT(mask m)" for x :: \'a::len word\ apply (rule trans[rotated], rule_tac w="mask n" in word_plus_and_or_coroll2) apply (simp add: word_bw_assocs word_bw_comms word_bw_lcs neg_mask_twice max_absorb2) done lemma word_of_nat_less: "\ n < unat x \ \ of_nat n < x" apply (simp add: word_less_nat_alt) apply (erule order_le_less_trans[rotated]) apply (simp add: unsigned_of_nat take_bit_eq_mod) done lemma unat_mask: "unat (mask n :: 'a :: len word) = 2 ^ (min n (LENGTH('a))) - 1" apply (subst min.commute) apply (simp add: mask_eq_decr_exp not_less min_def split: if_split_asm) apply (intro conjI impI) apply (simp add: unat_sub_if_size) apply (simp add: power_overflow word_size) apply (simp add: unat_sub_if_size) done lemma mask_over_length: "LENGTH('a) \ n \ mask n = (-1::'a::len word)" by (simp add: mask_eq_decr_exp) lemma Suc_2p_unat_mask: "n < LENGTH('a) \ Suc (2 ^ n * k + unat (mask n :: 'a::len word)) = 2 ^ n * (k+1)" by (simp add: unat_mask) lemma sint_of_nat_ge_zero: "x < 2 ^ (LENGTH('a) - 1) \ sint (of_nat x :: 'a :: len word) \ 0" by (simp add: bit_iff_odd signed_of_nat) lemma int_eq_sint: "x < 2 ^ (LENGTH('a) - 1) \ sint (of_nat x :: 'a :: len word) = int x" apply transfer apply (rule signed_take_bit_int_eq_self) apply simp_all apply (metis negative_zle numeral_power_eq_of_nat_cancel_iff) done lemma sint_of_nat_le: "\ b < 2 ^ (LENGTH('a) - 1); a \ b \ \ sint (of_nat a :: 'a :: len word) \ sint (of_nat b :: 'a :: len word)" apply (cases \LENGTH('a)\) apply simp_all apply transfer apply (subst signed_take_bit_eq_if_positive) apply (simp add: bit_simps) apply (metis bit_take_bit_iff nat_less_le order_less_le_trans take_bit_nat_eq_self_iff) apply (subst signed_take_bit_eq_if_positive) apply (simp add: bit_simps) apply (metis bit_take_bit_iff nat_less_le take_bit_nat_eq_self_iff) apply (simp flip: of_nat_take_bit add: take_bit_nat_eq_self) done lemma word_le_not_less: "((b::'a::len word) \ a) = (\(a < b))" by fastforce lemma less_is_non_zero_p1: fixes a :: "'a :: len word" shows "a < k \ a + 1 \ 0" apply (erule contrapos_pn) apply (drule max_word_wrap) apply (simp add: not_less) done lemma unat_add_lem': "(unat x + unat y < 2 ^ LENGTH('a)) \ (unat (x + y :: 'a :: len word) = unat x + unat y)" by (subst unat_add_lem[symmetric], assumption) lemma word_less_two_pow_divI: "\ (x :: 'a::len word) < 2 ^ (n - m); m \ n; n < LENGTH('a) \ \ x < 2 ^ n div 2 ^ m" apply (simp add: word_less_nat_alt) apply (subst unat_word_ariths) apply (subst mod_less) apply (rule order_le_less_trans [OF div_le_dividend]) apply (rule unat_lt2p) apply (simp add: power_sub) done lemma word_less_two_pow_divD: "\ (x :: 'a::len word) < 2 ^ n div 2 ^ m \ \ n \ m \ (x < 2 ^ (n - m))" apply (cases "n < LENGTH('a)") apply (cases "m < LENGTH('a)") apply (simp add: word_less_nat_alt) apply (subst(asm) unat_word_ariths) apply (subst(asm) mod_less) apply (rule order_le_less_trans [OF div_le_dividend]) apply (rule unat_lt2p) apply (clarsimp dest!: less_two_pow_divD) apply (simp add: power_overflow) apply (simp add: word_div_def) apply (simp add: power_overflow word_div_def) done lemma of_nat_less_two_pow_div_set: "\ n < LENGTH('a) \ \ {x. x < (2 ^ n div 2 ^ m :: 'a::len word)} = of_nat ` {k. k < 2 ^ n div 2 ^ m}" apply (simp add: image_def) apply (safe dest!: word_less_two_pow_divD less_two_pow_divD intro!: word_less_two_pow_divI) apply (rule_tac x="unat x" in exI) apply (simp add: power_sub[symmetric]) apply (subst unat_power_lower[symmetric, where 'a='a]) apply simp apply (erule unat_mono) apply (subst word_unat_power) apply (rule of_nat_mono_maybe) apply (rule power_strict_increasing) apply simp apply simp apply assumption done lemma ucast_less: "LENGTH('b) < LENGTH('a) \ (ucast (x :: 'b :: len word) :: ('a :: len word)) < 2 ^ LENGTH('b)" by transfer simp lemma ucast_range_less: "LENGTH('a :: len) < LENGTH('b :: len) \ range (ucast :: 'a word \ 'b word) = {x. x < 2 ^ len_of TYPE ('a)}" apply safe apply (erule ucast_less) apply (simp add: image_def) apply (rule_tac x="ucast x" in exI) apply (rule bit_word_eqI) apply (auto simp add: bit_simps) apply (metis bit_take_bit_iff take_bit_word_eq_self_iff) done lemma word_power_less_diff: "\2 ^ n * q < (2::'a::len word) ^ m; q < 2 ^ (LENGTH('a) - n)\ \ q < 2 ^ (m - n)" apply (case_tac "m \ LENGTH('a)") apply (simp add: power_overflow) apply (case_tac "n \ LENGTH('a)") apply (simp add: power_overflow) apply (cases "n = 0") apply simp apply (subst word_less_nat_alt) apply (subst unat_power_lower) apply simp apply (rule nat_power_less_diff) apply (simp add: word_less_nat_alt) apply (subst (asm) iffD1 [OF unat_mult_lem]) apply (simp add:nat_less_power_trans) apply simp done lemma word_less_sub_1: "x < (y :: 'a :: len word) \ x \ y - 1" by (fact word_le_minus_one_leq) lemma word_sub_mono2: "\ a + b \ c + d; c \ a; b \ a + b; d \ c + d \ \ b \ (d :: 'a :: len word)" by (drule(1) word_sub_mono; simp) lemma word_not_le: "(\ x \ (y :: 'a :: len word)) = (y < x)" by (fact not_le) lemma word_subset_less: "\ {x .. x + r - 1} \ {y .. y + s - 1}; x \ x + r - 1; y \ y + (s :: 'a :: len word) - 1; s \ 0 \ \ r \ s" apply (frule subsetD[where c=x]) apply simp apply (drule subsetD[where c="x + r - 1"]) apply simp apply (clarsimp simp: add_diff_eq[symmetric]) apply (drule(1) word_sub_mono2) apply (simp_all add: olen_add_eqv[symmetric]) apply (erule word_le_minus_cancel) apply (rule ccontr) apply (simp add: word_not_le) done lemma uint_power_lower: "n < LENGTH('a) \ uint (2 ^ n :: 'a :: len word) = (2 ^ n :: int)" by (rule uint_2p_alt) lemma power_le_mono: "\2 ^ n \ (2::'a::len word) ^ m; n < LENGTH('a); m < LENGTH('a)\ \ n \ m" apply (clarsimp simp add: le_less) apply safe apply (simp add: word_less_nat_alt) apply (simp only: uint_arith_simps(3)) apply (drule uint_power_lower)+ apply simp done lemma two_power_eq: "\n < LENGTH('a); m < LENGTH('a)\ \ ((2::'a::len word) ^ n = 2 ^ m) = (n = m)" apply safe apply (rule order_antisym) apply (simp add: power_le_mono[where 'a='a])+ done lemma unat_less_helper: "x < of_nat n \ unat x < n" apply (simp add: word_less_nat_alt) apply (erule order_less_le_trans) apply (simp add: take_bit_eq_mod unsigned_of_nat) done lemma nat_uint_less_helper: "nat (uint y) = z \ x < y \ nat (uint x) < z" apply (erule subst) apply (subst unat_eq_nat_uint [symmetric]) apply (subst unat_eq_nat_uint [symmetric]) by (simp add: unat_mono) lemma of_nat_0: "\of_nat n = (0::'a::len word); n < 2 ^ LENGTH('a)\ \ n = 0" by (auto simp add: word_of_nat_eq_0_iff) lemma of_nat_inj: "\x < 2 ^ LENGTH('a); y < 2 ^ LENGTH('a)\ \ (of_nat x = (of_nat y :: 'a :: len word)) = (x = y)" by (metis unat_of_nat_len) lemma div_to_mult_word_lt: "\ (x :: 'a :: len word) \ y div z \ \ x * z \ y" apply (cases "z = 0") apply simp apply (simp add: word_neq_0_conv) apply (rule order_trans) apply (erule(1) word_mult_le_mono1) apply (simp add: unat_div) apply (rule order_le_less_trans [OF div_mult_le]) apply simp apply (rule word_div_mult_le) done lemma ucast_ucast_mask: "(ucast :: 'a :: len word \ 'b :: len word) (ucast x) = x AND mask (len_of TYPE ('a))" apply (simp flip: take_bit_eq_mask) apply transfer apply (simp add: ac_simps) done lemma ucast_ucast_len: "\ x < 2 ^ LENGTH('b) \ \ ucast (ucast x::'b::len word) = (x::'a::len word)" apply (subst ucast_ucast_mask) apply (erule less_mask_eq) done lemma ucast_ucast_id: "LENGTH('a) < LENGTH('b) \ ucast (ucast (x::'a::len word)::'b::len word) = x" by (auto intro: ucast_up_ucast_id simp: is_up_def source_size_def target_size_def word_size) lemma unat_ucast: "unat (ucast x :: ('a :: len) word) = unat x mod 2 ^ (LENGTH('a))" proof - have \2 ^ LENGTH('a) = nat (2 ^ LENGTH('a))\ by simp moreover have \unat (ucast x :: 'a word) = unat x mod nat (2 ^ LENGTH('a))\ by transfer (simp flip: nat_mod_distrib take_bit_eq_mod) ultimately show ?thesis by (simp only:) qed lemma ucast_less_ucast: "LENGTH('a) \ LENGTH('b) \ (ucast x < ((ucast (y :: 'a::len word)) :: 'b::len word)) = (x < y)" apply (simp add: word_less_nat_alt unat_ucast) apply (subst mod_less) apply(rule less_le_trans[OF unat_lt2p], simp) apply (subst mod_less) apply(rule less_le_trans[OF unat_lt2p], simp) apply simp done \ \This weaker version was previously called @{text ucast_less_ucast}. We retain it to support existing proofs.\ lemmas ucast_less_ucast_weak = ucast_less_ucast[OF order.strict_implies_order] lemma unat_Suc2: fixes n :: "'a :: len word" shows "n \ -1 \ unat (n + 1) = Suc (unat n)" apply (subst add.commute, rule unatSuc) apply (subst eq_diff_eq[symmetric], simp add: minus_equation_iff) done lemma word_div_1: "(n :: 'a :: len word) div 1 = n" - by (fact bits_div_by_1) + by (fact div_by_1) lemma word_minus_one_le: "-1 \ (x :: 'a :: len word) = (x = -1)" by (fact word_order.extremum_unique) lemma up_scast_inj: "\ scast x = (scast y :: 'b :: len word); size x \ LENGTH('b) \ \ x = y" apply transfer apply (cases \LENGTH('a)\; simp) apply (metis order_refl take_bit_signed_take_bit take_bit_tightened) done lemma up_scast_inj_eq: "LENGTH('a) \ len_of TYPE ('b) \ (scast x = (scast y::'b::len word)) = (x = (y::'a::len word))" by (fastforce dest: up_scast_inj simp: word_size) lemma word_le_add: fixes x :: "'a :: len word" shows "x \ y \ \n. y = x + of_nat n" by (rule exI [where x = "unat (y - x)"]) simp lemma word_plus_mcs_4': "\x + v \ x + w; x \ x + v\ \ v \ w" for x :: "'a::len word" by (rule word_plus_mcs_4; simp add: add.commute) lemma unat_eq_1: \unat x = Suc 0 \ x = 1\ by (auto intro!: unsigned_word_eqI [where ?'a = nat]) lemma word_unat_Rep_inject1: \unat x = unat 1 \ x = 1\ by (simp add: unat_eq_1) lemma and_not_mask_twice: "(w AND NOT (mask n)) AND NOT (mask m) = w AND NOT (mask (max m n))" for w :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps) lemma word_less_cases: "x < y \ x = y - 1 \ x < y - (1 ::'a::len word)" apply (drule word_less_sub_1) apply (drule order_le_imp_less_or_eq) apply auto done lemma mask_and_mask: "mask a AND mask b = (mask (min a b) :: 'a::len word)" by (simp flip: take_bit_eq_mask ac_simps) lemma mask_eq_0_eq_x: "(x AND w = 0) = (x AND NOT w = x)" for x w :: \'a::len word\ using word_plus_and_or_coroll2[where x=x and w=w] by auto lemma mask_eq_x_eq_0: "(x AND w = x) = (x AND NOT w = 0)" for x w :: \'a::len word\ using word_plus_and_or_coroll2[where x=x and w=w] by auto lemma compl_of_1: "NOT 1 = (-2 :: 'a :: len word)" by (fact not_one_eq) lemma split_word_eq_on_mask: "(x = y) = (x AND m = y AND m \ x AND NOT m = y AND NOT m)" for x y m :: \'a::len word\ apply transfer apply (simp add: bit_eq_iff) apply (auto simp add: bit_simps ac_simps) done lemma word_FF_is_mask: "0xFF = (mask 8 :: 'a::len word)" by (simp add: mask_eq_decr_exp) lemma word_1FF_is_mask: "0x1FF = (mask 9 :: 'a::len word)" by (simp add: mask_eq_decr_exp) lemma ucast_of_nat_small: "x < 2 ^ LENGTH('a) \ ucast (of_nat x :: 'a :: len word) = (of_nat x :: 'b :: len word)" apply transfer apply (auto simp add: take_bit_of_nat min_def not_le) apply (metis linorder_not_less min_def take_bit_nat_eq_self take_bit_take_bit) done lemma word_le_make_less: fixes x :: "'a :: len word" shows "y \ -1 \ (x \ y) = (x < (y + 1))" apply safe apply (erule plus_one_helper2) apply (simp add: eq_diff_eq[symmetric]) done lemmas finite_word = finite [where 'a="'a::len word"] lemma word_to_1_set: "{0 ..< (1 :: 'a :: len word)} = {0}" by fastforce lemma word_leq_minus_one_le: fixes x :: "'a::len word" shows "\y \ 0; x \ y - 1 \ \ x < y" using le_m1_iff_lt word_neq_0_conv by blast lemma word_count_from_top: "n \ 0 \ {0 ..< n :: 'a :: len word} = {0 ..< n - 1} \ {n - 1}" apply (rule set_eqI, rule iffI) apply simp apply (drule word_le_minus_one_leq) apply (rule disjCI) apply simp apply simp apply (erule word_leq_minus_one_le) apply fastforce done lemma word_minus_one_le_leq: "\ x - 1 < y \ \ x \ (y :: 'a :: len word)" apply (cases "x = 0") apply simp apply (simp add: word_less_nat_alt word_le_nat_alt) apply (subst(asm) unat_minus_one) apply (simp add: word_less_nat_alt) apply (cases "unat x") apply (simp add: unat_eq_zero) apply arith done lemma word_must_wrap: "\ x \ n - 1; n \ x \ \ n = (0 :: 'a :: len word)" using dual_order.trans sub_wrap word_less_1 by blast lemma range_subset_card: "\ {a :: 'a :: len word .. b} \ {c .. d}; b \ a \ \ d \ c \ d - c \ b - a" using word_sub_le word_sub_mono by fastforce lemma less_1_simp: "n - 1 < m = (n \ (m :: 'a :: len word) \ n \ 0)" by unat_arith lemma word_power_mod_div: fixes x :: "'a::len word" shows "\ n < LENGTH('a); m < LENGTH('a)\ \ x mod 2 ^ n div 2 ^ m = x div 2 ^ m mod 2 ^ (n - m)" apply (simp add: word_arith_nat_div unat_mod power_mod_div) apply (subst unat_arith_simps(3)) apply (subst unat_mod) apply (subst unat_of_nat)+ apply (simp add: mod_mod_power min.commute) done lemma word_range_minus_1': fixes a :: "'a :: len word" shows "a \ 0 \ {a - 1<..b} = {a..b}" by (simp add: greaterThanAtMost_def atLeastAtMost_def greaterThan_def atLeast_def less_1_simp) lemma word_range_minus_1: fixes a :: "'a :: len word" shows "b \ 0 \ {a..b - 1} = {a.. 'b :: len word) x" by transfer simp lemma overflow_plus_one_self: "(1 + p \ p) = (p = (-1 :: 'a :: len word))" apply rule apply (rule ccontr) apply (drule plus_one_helper2) apply (rule notI) apply (drule arg_cong[where f="\x. x - 1"]) apply simp apply (simp add: field_simps) apply simp done lemma plus_1_less: "(x + 1 \ (x :: 'a :: len word)) = (x = -1)" apply (rule iffI) apply (rule ccontr) apply (cut_tac plus_one_helper2[where x=x, OF order_refl]) apply simp apply clarsimp apply (drule arg_cong[where f="\x. x - 1"]) apply simp apply simp done lemma pos_mult_pos_ge: "[|x > (0::int); n>=0 |] ==> n * x >= n*1" apply (simp only: mult_left_mono) done lemma word_plus_strict_mono_right: fixes x :: "'a :: len word" shows "\y < z; x \ x + z\ \ x + y < x + z" by unat_arith lemma word_div_mult: "0 < c \ a < b * c \ a div c < b" for a b c :: "'a::len word" by (rule classical) (use div_to_mult_word_lt [of b a c] in \auto simp add: word_less_nat_alt word_le_nat_alt unat_div\) lemma word_less_power_trans_ofnat: "\n < 2 ^ (m - k); k \ m; m < LENGTH('a)\ \ of_nat n * 2 ^ k < (2::'a::len word) ^ m" apply (subst mult.commute) apply (rule word_less_power_trans) apply (simp_all add: word_less_nat_alt unsigned_of_nat) using take_bit_nat_less_eq_self apply (rule le_less_trans) apply assumption done lemma word_1_le_power: "n < LENGTH('a) \ (1 :: 'a :: len word) \ 2 ^ n" by (rule inc_le[where i=0, simplified], erule iffD2[OF p2_gt_0]) lemma unat_1_0: "1 \ (x::'a::len word) = (0 < unat x)" by (auto simp add: word_le_nat_alt) lemma x_less_2_0_1': fixes x :: "'a::len word" shows "\LENGTH('a) \ 1; x < 2\ \ x = 0 \ x = 1" apply (cases \2 \ LENGTH('a)\; simp) apply transfer apply clarsimp apply (metis add.commute add.right_neutral even_two_times_div_two mod_div_trivial mod_pos_pos_trivial mult.commute mult_zero_left not_less not_take_bit_negative odd_two_times_div_two_succ) done lemmas word_add_le_iff2 = word_add_le_iff [folded no_olen_add_nat] lemma of_nat_power: shows "\ p < 2 ^ x; x < len_of TYPE ('a) \ \ of_nat p < (2 :: 'a :: len word) ^ x" apply (rule order_less_le_trans) apply (rule of_nat_mono_maybe) apply (erule power_strict_increasing) apply simp apply assumption apply (simp add: word_unat_power del: of_nat_power) done lemma of_nat_n_less_equal_power_2: "n < LENGTH('a::len) \ ((of_nat n)::'a word) < 2 ^ n" apply (induct n) apply clarsimp apply clarsimp apply (metis of_nat_power n_less_equal_power_2 of_nat_Suc power_Suc) done lemma eq_mask_less: fixes w :: "'a::len word" assumes eqm: "w = w AND mask n" and sz: "n < len_of TYPE ('a)" shows "w < (2::'a word) ^ n" by (subst eqm, rule and_mask_less' [OF sz]) lemma of_nat_mono_maybe': fixes Y :: "nat" assumes xlt: "x < 2 ^ len_of TYPE ('a)" assumes ylt: "y < 2 ^ len_of TYPE ('a)" shows "(y < x) = (of_nat y < (of_nat x :: 'a :: len word))" apply (subst word_less_nat_alt) apply (subst unat_of_nat)+ apply (subst mod_less) apply (rule ylt) apply (subst mod_less) apply (rule xlt) apply simp done lemma of_nat_mono_maybe_le: "\x < 2 ^ LENGTH('a); y < 2 ^ LENGTH('a)\ \ (y \ x) = ((of_nat y :: 'a :: len word) \ of_nat x)" apply (clarsimp simp: le_less) apply (rule disj_cong) apply (rule of_nat_mono_maybe', assumption+) apply auto using of_nat_inj apply blast done lemma mask_AND_NOT_mask: "(w AND NOT (mask n)) AND mask n = 0" for w :: \'a::len word\ by (rule bit_word_eqI) (simp add: bit_simps) lemma AND_NOT_mask_plus_AND_mask_eq: "(w AND NOT (mask n)) + (w AND mask n) = w" for w :: \'a::len word\ apply (subst disjunctive_add) apply (auto simp add: bit_simps) apply (rule bit_word_eqI) apply (auto simp add: bit_simps) done lemma mask_eqI: fixes x :: "'a :: len word" assumes m1: "x AND mask n = y AND mask n" and m2: "x AND NOT (mask n) = y AND NOT (mask n)" shows "x = y" proof - have *: \x = x AND mask n OR x AND NOT (mask n)\ for x :: \'a word\ by (rule bit_word_eqI) (auto simp add: bit_simps) from assms * [of x] * [of y] show ?thesis by simp qed lemma neq_0_no_wrap: fixes x :: "'a :: len word" shows "\ x \ x + y; x \ 0 \ \ x + y \ 0" by clarsimp lemma unatSuc2: fixes n :: "'a :: len word" shows "n + 1 \ 0 \ unat (n + 1) = Suc (unat n)" by (simp add: add.commute unatSuc) lemma word_of_nat_le: "n \ unat x \ of_nat n \ x" apply (simp add: word_le_nat_alt unat_of_nat) apply (erule order_trans[rotated]) apply (simp add: take_bit_eq_mod) done lemma word_unat_less_le: "a \ of_nat b \ unat a \ b" by (metis eq_iff le_cases le_unat_uoi word_of_nat_le) lemma mask_Suc_0 : "mask (Suc 0) = (1 :: 'a::len word)" by (simp add: mask_eq_decr_exp) lemma bool_mask': fixes x :: "'a :: len word" shows "2 < LENGTH('a) \ (0 < x AND 1) = (x AND 1 = 1)" by (simp add: and_one_eq mod_2_eq_odd) lemma ucast_ucast_add: fixes x :: "'a :: len word" fixes y :: "'b :: len word" shows "LENGTH('b) \ LENGTH('a) \ ucast (ucast x + y) = x + ucast y" apply transfer apply simp apply (subst (2) take_bit_add [symmetric]) apply (subst take_bit_add [symmetric]) apply simp done lemma lt1_neq0: fixes x :: "'a :: len word" shows "(1 \ x) = (x \ 0)" by unat_arith lemma word_plus_one_nonzero: fixes x :: "'a :: len word" shows "\x \ x + y; y \ 0\ \ x + 1 \ 0" apply (subst lt1_neq0 [symmetric]) apply (subst olen_add_eqv [symmetric]) apply (erule word_random) apply (simp add: lt1_neq0) done lemma word_sub_plus_one_nonzero: fixes n :: "'a :: len word" shows "\n' \ n; n' \ 0\ \ (n - n') + 1 \ 0" apply (subst lt1_neq0 [symmetric]) apply (subst olen_add_eqv [symmetric]) apply (rule word_random [where x' = n']) apply simp apply (erule word_sub_le) apply (simp add: lt1_neq0) done lemma word_le_minus_mono_right: fixes x :: "'a :: len word" shows "\ z \ y; y \ x; z \ x \ \ x - y \ x - z" apply (rule word_sub_mono) apply simp apply assumption apply (erule word_sub_le) apply (erule word_sub_le) done lemma word_0_sle_from_less: \0 \s x\ if \x < 2 ^ (LENGTH('a) - 1)\ for x :: \'a::len word\ using that apply transfer apply (cases \LENGTH('a)\) apply simp_all apply (metis bit_take_bit_iff min_def nat_less_le not_less_eq take_bit_int_eq_self_iff take_bit_take_bit) done lemma ucast_sub_ucast: fixes x :: "'a::len word" assumes "y \ x" assumes T: "LENGTH('a) \ LENGTH('b)" shows "ucast (x - y) = (ucast x - ucast y :: 'b::len word)" proof - from T have P: "unat x < 2 ^ LENGTH('b)" "unat y < 2 ^ LENGTH('b)" by (fastforce intro!: less_le_trans[OF unat_lt2p])+ then show ?thesis by (simp add: unat_arith_simps unat_ucast assms[simplified unat_arith_simps]) qed lemma word_1_0: "\a + (1::('a::len) word) \ b; a < of_nat x\ \ a < b" apply transfer apply (subst (asm) take_bit_incr_eq) apply (auto simp add: diff_less_eq) using take_bit_int_less_exp le_less_trans by blast lemma unat_of_nat_less:"\ a < b; unat b = c \ \ a < of_nat c" by fastforce lemma word_le_plus_1: "\ (y::('a::len) word) < y + n; a < n \ \ y + a \ y + a + 1" by unat_arith lemma word_le_plus:"\(a::('a::len) word) < a + b; c < b\ \ a \ a + c" by (metis order_less_imp_le word_random) lemma sint_minus1 [simp]: "(sint x = -1) = (x = -1)" apply (cases \LENGTH('a)\) apply simp_all apply transfer apply (simp only: flip: signed_take_bit_eq_iff_take_bit_eq) apply simp done lemma sint_0 [simp]: "(sint x = 0) = (x = 0)" by (fact signed_eq_0_iff) (* It is not always that case that "sint 1 = 1", because of 1-bit word sizes. * This lemma produces the different cases. *) lemma sint_1_cases: P if \\ len_of TYPE ('a::len) = 1; (a::'a word) = 0; sint a = 0 \ \ P\ \\ len_of TYPE ('a) = 1; a = 1; sint (1 :: 'a word) = -1 \ \ P\ \\ len_of TYPE ('a) > 1; sint (1 :: 'a word) = 1 \ \ P\ proof (cases \LENGTH('a) = 1\) case True then have \a = 0 \ a = 1\ by transfer auto with True that show ?thesis by auto next case False with that show ?thesis by (simp add: less_le Suc_le_eq) qed lemma sint_int_min: "sint (- (2 ^ (LENGTH('a) - Suc 0)) :: ('a::len) word) = - (2 ^ (LENGTH('a) - Suc 0))" apply (cases \LENGTH('a)\) apply simp_all apply transfer apply (simp add: signed_take_bit_int_eq_self) done lemma sint_int_max_plus_1: "sint (2 ^ (LENGTH('a) - Suc 0) :: ('a::len) word) = - (2 ^ (LENGTH('a) - Suc 0))" apply (cases \LENGTH('a)\) apply simp_all apply (subst word_of_int_2p [symmetric]) apply (subst int_word_sint) apply simp done lemma uint_range': \0 \ uint x \ uint x < 2 ^ LENGTH('a)\ for x :: \'a::len word\ by transfer simp lemma sint_of_int_eq: "\ - (2 ^ (LENGTH('a) - 1)) \ x; x < 2 ^ (LENGTH('a) - 1) \ \ sint (of_int x :: ('a::len) word) = x" by (simp add: signed_take_bit_int_eq_self signed_of_int) lemma of_int_sint: "of_int (sint a) = a" by simp lemma sint_ucast_eq_uint: "\ \ is_down (ucast :: ('a::len word \ 'b::len word)) \ \ sint ((ucast :: ('a::len word \ 'b::len word)) x) = uint x" apply transfer apply (simp add: signed_take_bit_take_bit) done lemma word_less_nowrapI': "(x :: 'a :: len word) \ z - k \ k \ z \ 0 < k \ x < x + k" by uint_arith lemma mask_plus_1: "mask n + 1 = (2 ^ n :: 'a::len word)" by (clarsimp simp: mask_eq_decr_exp) lemma unat_inj: "inj unat" by (metis eq_iff injI word_le_nat_alt) lemma unat_ucast_upcast: "is_up (ucast :: 'b word \ 'a word) \ unat (ucast x :: ('a::len) word) = unat (x :: ('b::len) word)" unfolding ucast_eq unat_eq_nat_uint apply transfer apply simp done lemma ucast_mono: "\ (x :: 'b :: len word) < y; y < 2 ^ LENGTH('a) \ \ ucast x < ((ucast y) :: 'a :: len word)" apply (simp only: flip: ucast_nat_def) apply (rule of_nat_mono_maybe) apply (rule unat_less_helper) apply simp apply (simp add: word_less_nat_alt) done lemma ucast_mono_le: "\x \ y; y < 2 ^ LENGTH('b)\ \ (ucast (x :: 'a :: len word) :: 'b :: len word) \ ucast y" apply (simp only: flip: ucast_nat_def) apply (subst of_nat_mono_maybe_le[symmetric]) apply (rule unat_less_helper) apply simp apply (rule unat_less_helper) apply (erule le_less_trans) apply (simp_all add: word_le_nat_alt) done lemma ucast_mono_le': "\ unat y < 2 ^ LENGTH('b); LENGTH('b::len) < LENGTH('a::len); x \ y \ \ ucast x \ (ucast y :: 'b word)" for x y :: \'a::len word\ by (auto simp: word_less_nat_alt intro: ucast_mono_le) lemma neg_mask_add_mask: "((x:: 'a :: len word) AND NOT (mask n)) + (2 ^ n - 1) = x OR mask n" unfolding mask_2pm1 [symmetric] apply (subst word_plus_and_or_coroll; rule bit_word_eqI) apply (auto simp add: bit_simps) done lemma le_step_down_word:"\(i::('a::len) word) \ n; i = n \ P; i \ n - 1 \ P\ \ P" by unat_arith lemma le_step_down_word_2: fixes x :: "'a::len word" shows "\x \ y; x \ y\ \ x \ y - 1" by (subst (asm) word_le_less_eq, clarsimp, simp add: word_le_minus_one_leq) lemma NOT_mask_AND_mask[simp]: "(w AND mask n) AND NOT (mask n) = 0" by (rule bit_eqI) (simp add: bit_simps) lemma and_and_not[simp]:"(a AND b) AND NOT b = 0" for a b :: \'a::len word\ apply (subst word_bw_assocs(1)) apply clarsimp done lemma ex_mask_1[simp]: "(\x. mask x = (1 :: 'a::len word))" apply (rule_tac x=1 in exI) apply (simp add:mask_eq_decr_exp) done lemma not_switch:"NOT a = x \ a = NOT x" by auto lemma test_bit_eq_iff: "bit u = bit v \ u = v" for u v :: "'a::len word" by (auto intro: bit_eqI simp add: fun_eq_iff) lemma test_bit_size: "bit w n \ n < size w" for w :: "'a::len word" by transfer simp lemma word_eq_iff: "x = y \ (\nn. n < size u \ bit u n = bit v n) \ u = v" for u :: "'a::len word" by (simp add: word_size word_eq_iff) lemma word_eqD: "u = v \ bit u x = bit v x" for u v :: "'a::len word" by simp lemma test_bit_bin': "bit w n \ n < size w \ bit (uint w) n" by transfer (simp add: bit_take_bit_iff) lemmas test_bit_bin = test_bit_bin' [unfolded word_size] lemma word_test_bit_def: \bit a = bit (uint a)\ by transfer (simp add: fun_eq_iff bit_take_bit_iff) lemmas test_bit_def' = word_test_bit_def [THEN fun_cong] lemma word_test_bit_transfer [transfer_rule]: "(rel_fun pcr_word (rel_fun (=) (=))) (\x n. n < LENGTH('a) \ bit x n) (bit :: 'a::len word \ _)" by transfer_prover lemma test_bit_wi: "bit (word_of_int x :: 'a::len word) n \ n < LENGTH('a) \ bit x n" by transfer simp lemma word_ops_nth_size: "n < size x \ bit (x OR y) n = (bit x n | bit y n) \ bit (x AND y) n = (bit x n \ bit y n) \ bit (x XOR y) n = (bit x n \ bit y n) \ bit (NOT x) n = (\ bit x n)" for x :: "'a::len word" by transfer (simp add: bit_or_iff bit_and_iff bit_xor_iff bit_not_iff) lemma word_ao_nth: "bit (x OR y) n = (bit x n | bit y n) \ bit (x AND y) n = (bit x n \ bit y n)" for x :: "'a::len word" by transfer (auto simp add: bit_or_iff bit_and_iff) lemmas lsb0 = len_gt_0 [THEN word_ops_nth_size [unfolded word_size]] lemma nth_sint: fixes w :: "'a::len word" defines "l \ LENGTH('a)" shows "bit (sint w) n = (if n < l - 1 then bit w n else bit w (l - 1))" unfolding sint_uint l_def by (auto simp: bit_signed_take_bit_iff word_test_bit_def not_less min_def) lemma test_bit_2p: "bit (word_of_int (2 ^ n)::'a::len word) m \ m = n \ m < LENGTH('a)" by transfer (auto simp add: bit_exp_iff) lemma nth_w2p: "bit ((2::'a::len word) ^ n) m \ m = n \ m < LENGTH('a::len)" by transfer (auto simp add: bit_exp_iff) lemma bang_is_le: "bit x m \ 2 ^ m \ x" for x :: "'a::len word" apply (rule xtrans(3)) apply (rule_tac [2] y = "x" in le_word_or2) apply (rule word_eqI) apply (auto simp add: word_ao_nth nth_w2p word_size) done lemmas msb0 = len_gt_0 [THEN diff_Suc_less, THEN word_ops_nth_size [unfolded word_size]] lemmas msb1 = msb0 [where i = 0] lemma test_bit_1 [iff]: "bit (1 :: 'a::len word) n \ n = 0" by transfer (auto simp add: bit_1_iff) lemma nth_0: "\ bit (0 :: 'a::len word) n" by transfer simp lemma nth_minus1: "bit (-1 :: 'a::len word) n \ n < LENGTH('a)" by transfer simp lemma nth_ucast_weak: "bit (ucast w::'a::len word) n = (bit w n \ n < LENGTH('a))" by transfer (simp add: bit_take_bit_iff ac_simps) lemma nth_ucast: "bit (ucast (w::'a::len word)::'b::len word) n = (bit w n \ n < min LENGTH('a) LENGTH('b))" by (auto simp: not_le nth_ucast_weak dest: bit_imp_le_length) lemma nth_mask: \bit (mask n :: 'a::len word) i \ i < n \ i < size (mask n :: 'a word)\ by (auto simp add: word_size Word.bit_mask_iff) lemma nth_slice: "bit (slice n w :: 'a::len word) m = (bit w (m + n) \ m < LENGTH('a))" apply (auto simp add: bit_simps less_diff_conv dest: bit_imp_le_length) using bit_imp_le_length apply fastforce done lemma test_bit_cat [OF refl]: "wc = word_cat a b \ bit wc n = (n < size wc \ (if n < size b then bit b n else bit a (n - size b)))" apply (simp add: word_size not_less; transfer) apply (auto simp add: bit_concat_bit_iff bit_take_bit_iff) done \ \keep quantifiers for use in simplification\ lemma test_bit_split': "word_split c = (a, b) \ (\n m. bit b n = (n < size b \ bit c n) \ bit a m = (m < size a \ bit c (m + size b)))" by (auto simp add: word_split_bin' bit_unsigned_iff word_size bit_drop_bit_eq ac_simps dest: bit_imp_le_length) lemma test_bit_split: "word_split c = (a, b) \ (\n::nat. bit b n \ n < size b \ bit c n) \ (\m::nat. bit a m \ m < size a \ bit c (m + size b))" by (simp add: test_bit_split') lemma test_bit_split_eq: "word_split c = (a, b) \ ((\n::nat. bit b n = (n < size b \ bit c n)) \ (\m::nat. bit a m = (m < size a \ bit c (m + size b))))" apply (rule_tac iffI) apply (rule_tac conjI) apply (erule test_bit_split [THEN conjunct1]) apply (erule test_bit_split [THEN conjunct2]) apply (case_tac "word_split c") apply (frule test_bit_split) apply (erule trans) apply (fastforce intro!: word_eqI simp add: word_size) done lemma test_bit_rcat: "sw = size (hd wl) \ rc = word_rcat wl \ bit rc n = (n < size rc \ n div sw < size wl \ bit ((rev wl) ! (n div sw)) (n mod sw))" for wl :: "'a::len word list" by (simp add: word_size word_rcat_def rev_map bit_horner_sum_uint_exp_iff bit_simps not_le) lemmas test_bit_cong = arg_cong [where f = "bit", THEN fun_cong] lemma max_test_bit: "bit (- 1::'a::len word) n \ n < LENGTH('a)" by (fact nth_minus1) lemma map_nth_0 [simp]: "map (bit (0::'a::len word)) xs = replicate (length xs) False" by (simp flip: map_replicate_const) lemma word_and_1: "n AND 1 = (if bit n 0 then 1 else 0)" for n :: "_ word" by (rule bit_word_eqI) (auto simp add: bit_and_iff bit_1_iff intro: gr0I) lemma test_bit_1': "bit (1 :: 'a :: len word) n \ 0 < LENGTH('a) \ n = 0" by simp lemma nth_w2p_same: "bit (2^n :: 'a :: len word) n = (n < LENGTH('a))" by (simp add: nth_w2p) lemma word_leI: "(\n. \n < size (u::'a::len word); bit u n \ \ bit (v::'a::len word) n) \ u <= v" apply (rule order_trans [of u \u AND v\ v]) apply (rule eq_refl) apply (rule bit_word_eqI) apply (auto simp add: bit_simps word_and_le1 word_size) done lemma bang_eq: fixes x :: "'a::len word" shows "(x = y) = (\n. bit x n = bit y n)" by (auto intro!: bit_eqI) lemma neg_mask_test_bit: "bit (NOT(mask n) :: 'a :: len word) m = (n \ m \ m < LENGTH('a))" by (auto simp add: bit_simps) lemma upper_bits_unset_is_l2p: \(\n' \ n. n' < LENGTH('a) \ \ bit p n') \ (p < 2 ^ n)\ (is \?P \ ?Q\) if \n < LENGTH('a)\ for p :: "'a :: len word" proof assume ?Q then show ?P by (meson bang_is_le le_less_trans not_le word_power_increasing) next assume ?P have \take_bit n p = p\ proof (rule bit_word_eqI) fix q assume \q < LENGTH('a)\ show \bit (take_bit n p) q \ bit p q\ proof (cases \q < n\) case True then show ?thesis by (auto simp add: bit_simps) next case False then have \n \ q\ by simp with \?P\ \q < LENGTH('a)\ have \\ bit p q\ by simp then show ?thesis by (simp add: bit_simps) qed qed with that show ?Q using take_bit_word_eq_self_iff [of n p] by auto qed lemma less_2p_is_upper_bits_unset: "p < 2 ^ n \ n < LENGTH('a) \ (\n' \ n. n' < LENGTH('a) \ \ bit p n')" for p :: "'a :: len word" by (meson le_less_trans le_mask_iff_lt_2n upper_bits_unset_is_l2p word_zero_le) lemma test_bit_over: "n \ size (x::'a::len word) \ (bit x n) = False" by transfer auto lemma le_mask_high_bits: "w \ mask n \ (\i \ {n ..< size w}. \ bit w i)" for w :: \'a::len word\ apply (auto simp add: bit_simps word_size less_eq_mask_iff_take_bit_eq_self) apply (metis bit_take_bit_iff leD) apply (metis atLeastLessThan_iff leI take_bit_word_eq_self_iff upper_bits_unset_is_l2p) done lemma test_bit_conj_lt: "(bit x m \ m < LENGTH('a)) = bit x m" for x :: "'a :: len word" using test_bit_bin by blast lemma neg_test_bit: "bit (NOT x) n = (\ bit x n \ n < LENGTH('a))" for x :: "'a::len word" by (cases "n < LENGTH('a)") (auto simp add: test_bit_over word_ops_nth_size word_size) lemma nth_bounded: "\bit (x :: 'a :: len word) n; x < 2 ^ m; m \ len_of TYPE ('a)\ \ n < m" apply (rule ccontr) apply (auto simp add: not_less) apply (meson bit_imp_le_length bit_uint_iff less_2p_is_upper_bits_unset test_bit_bin) done lemma and_neq_0_is_nth: \x AND y \ 0 \ bit x n\ if \y = 2 ^ n\ for x y :: \'a::len word\ apply (simp add: bit_eq_iff bit_simps) using that apply (simp add: bit_simps not_le) apply transfer apply auto done lemma nth_is_and_neq_0: "bit (x::'a::len word) n = (x AND 2 ^ n \ 0)" by (subst and_neq_0_is_nth; rule refl) lemma max_word_not_less [simp]: "\ - 1 < x" for x :: \'a::len word\ by (fact word_order.extremum_strict) lemma bit_twiddle_min: "(y::'a::len word) XOR (((x::'a::len word) XOR y) AND (if x < y then -1 else 0)) = min x y" by (rule bit_eqI) (auto simp add: bit_simps) lemma bit_twiddle_max: "(x::'a::len word) XOR (((x::'a::len word) XOR y) AND (if x < y then -1 else 0)) = max x y" by (rule bit_eqI) (auto simp add: bit_simps max_def) lemma swap_with_xor: "\(x::'a::len word) = a XOR b; y = b XOR x; z = x XOR y\ \ z = b \ y = a" by (auto intro: bit_word_eqI simp add: bit_simps) lemma le_mask_imp_and_mask: "(x::'a::len word) \ mask n \ x AND mask n = x" by (metis and_mask_eq_iff_le_mask) lemma or_not_mask_nop: "((x::'a::len word) OR NOT (mask n)) AND mask n = x AND mask n" by (metis word_and_not word_ao_dist2 word_bw_comms(1) word_log_esimps(3)) lemma mask_subsume: "\n \ m\ \ ((x::'a::len word) OR y AND mask n) AND NOT (mask m) = x AND NOT (mask m)" by (rule bit_word_eqI) (auto simp add: bit_simps word_size) lemma and_mask_0_iff_le_mask: fixes w :: "'a::len word" shows "(w AND NOT(mask n) = 0) = (w \ mask n)" by (simp add: mask_eq_0_eq_x le_mask_imp_and_mask and_mask_eq_iff_le_mask) lemma mask_twice2: "n \ m \ ((x::'a::len word) AND mask m) AND mask n = x AND mask n" by (metis mask_twice min_def) lemma uint_2_id: "LENGTH('a) \ 2 \ uint (2::('a::len) word) = 2" by simp lemma div_of_0_id[simp]:"(0::('a::len) word) div n = 0" by (simp add: word_div_def) lemma degenerate_word:"LENGTH('a) = 1 \ (x::('a::len) word) = 0 \ x = 1" by (metis One_nat_def less_irrefl_nat sint_1_cases) lemma div_by_0_word: "(x::'a::len word) div 0 = 0" - by (fact bits_div_by_0) + by (fact div_by_0) lemma div_less_dividend_word:"\x \ 0; n \ 1\ \ (x::('a::len) word) div n < x" apply (cases \n = 0\) apply clarsimp apply (simp add:word_neq_0_conv) apply (subst word_arith_nat_div) apply (rule word_of_nat_less) apply (rule div_less_dividend) using unat_eq_zero word_unat_Rep_inject1 apply force apply (simp add:unat_gt_0) done lemma word_less_div: fixes x :: "('a::len) word" and y :: "('a::len) word" shows "x div y = 0 \ y = 0 \ x < y" apply (case_tac "y = 0", clarsimp+) by (metis One_nat_def Suc_le_mono le0 le_div_geq not_less unat_0 unat_div unat_gt_0 word_less_nat_alt zero_less_one) lemma not_degenerate_imp_2_neq_0:"LENGTH('a) > 1 \ (2::('a::len) word) \ 0" by (metis numerals(1) power_not_zero power_zero_numeral) lemma word_overflow:"(x::('a::len) word) + 1 > x \ x + 1 = 0" apply clarsimp by (metis diff_0 eq_diff_eq less_x_plus_1) lemma word_overflow_unat:"unat ((x::('a::len) word) + 1) = unat x + 1 \ x + 1 = 0" by (metis Suc_eq_plus1 add.commute unatSuc) lemma even_word_imp_odd_next:"even (unat (x::('a::len) word)) \ x + 1 = 0 \ odd (unat (x + 1))" apply (cut_tac x=x in word_overflow_unat) apply clarsimp done lemma odd_word_imp_even_next:"odd (unat (x::('a::len) word)) \ x + 1 = 0 \ even (unat (x + 1))" apply (cut_tac x=x in word_overflow_unat) apply clarsimp done lemma overflow_imp_lsb:"(x::('a::len) word) + 1 = 0 \ bit x 0" using even_plus_one_iff [of x] by (simp add: bit_0) lemma odd_iff_lsb:"odd (unat (x::('a::len) word)) = bit x 0" by transfer (simp add: even_nat_iff bit_0) lemma of_nat_neq_iff_word: "x mod 2 ^ LENGTH('a) \ y mod 2 ^ LENGTH('a) \ (((of_nat x)::('a::len) word) \ of_nat y) = (x \ y)" apply (rule iffI) apply (case_tac "x = y") apply (subst (asm) of_nat_eq_iff[symmetric]) apply auto apply (case_tac "((of_nat x)::('a::len) word) = of_nat y") apply auto apply (metis unat_of_nat) done lemma lsb_this_or_next: "\ (bit ((x::('a::len) word) + 1) 0) \ bit x 0" by (simp add: bit_0) lemma mask_or_not_mask: "x AND mask n OR x AND NOT (mask n) = x" for x :: \'a::len word\ apply (subst word_oa_dist, simp) apply (subst word_oa_dist2, simp) done lemma word_gr0_conv_Suc: "(m::'a::len word) > 0 \ \n. m = n + 1" by (metis add.commute add_minus_cancel) lemma revcast_down_us [OF refl]: "rc = revcast \ source_size rc = target_size rc + n \ rc w = ucast (signed_drop_bit n w)" for w :: "'a::len word" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_simps ac_simps) done lemma revcast_down_ss [OF refl]: "rc = revcast \ source_size rc = target_size rc + n \ rc w = scast (signed_drop_bit n w)" for w :: "'a::len word" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_simps ac_simps) done lemma revcast_down_uu [OF refl]: "rc = revcast \ source_size rc = target_size rc + n \ rc w = ucast (drop_bit n w)" for w :: "'a::len word" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_simps ac_simps) done lemma revcast_down_su [OF refl]: "rc = revcast \ source_size rc = target_size rc + n \ rc w = scast (drop_bit n w)" for w :: "'a::len word" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_simps ac_simps) done lemma cast_down_rev [OF refl]: "uc = ucast \ source_size uc = target_size uc + n \ uc w = revcast (push_bit n w)" for w :: "'a::len word" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_simps) done lemma revcast_up [OF refl]: "rc = revcast \ source_size rc + n = target_size rc \ rc w = push_bit n (ucast w :: 'a::len word)" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_simps) apply auto apply (metis add.commute add_diff_cancel_right) apply (metis diff_add_inverse2 diff_diff_add) done lemmas rc1 = revcast_up [THEN revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]] lemmas rc2 = revcast_down_uu [THEN revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]] lemma word_ops_nth: fixes x y :: \'a::len word\ shows word_or_nth: "bit (x OR y) n = (bit x n \ bit y n)" and word_and_nth: "bit (x AND y) n = (bit x n \ bit y n)" and word_xor_nth: "bit (x XOR y) n = (bit x n \ bit y n)" by (simp_all add: bit_simps) lemma word_power_nonzero: "\ (x :: 'a::len word) < 2 ^ (LENGTH('a) - n); n < LENGTH('a); x \ 0 \ \ x * 2 ^ n \ 0" - by (metis Word.word_div_mult bits_div_0 len_gt_0 len_of_finite_2_def nat_mult_power_less_eq - p2_gt_0 unat_mono unat_power_lower word_gt_a_gt_0) + by (metis gr0I mult.commute not_less_eq p2_gt_0 power_0 word_less_1 word_power_less_diff zero_less_diff) lemma less_1_helper: "n \ m \ (n - 1 :: int) < m" by arith lemma div_power_helper: "\ x \ y; y < LENGTH('a) \ \ (2 ^ y - 1) div (2 ^ x :: 'a::len word) = 2 ^ (y - x) - 1" apply (simp flip: mask_eq_exp_minus_1 drop_bit_eq_div) apply (rule bit_word_eqI) apply (auto simp add: bit_simps not_le) done lemma max_word_mask: "(- 1 :: 'a::len word) = mask LENGTH('a)" by (fact minus_1_eq_mask) lemmas mask_len_max = max_word_mask[symmetric] lemma mask_out_first_mask_some: "\ x AND NOT (mask n) = y; n \ m \ \ x AND NOT (mask m) = y AND NOT (mask m)" for x y :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps word_size) lemma mask_lower_twice: "n \ m \ (x AND NOT (mask n)) AND NOT (mask m) = x AND NOT (mask m)" for x :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps word_size) lemma mask_lower_twice2: "(a AND NOT (mask n)) AND NOT (mask m) = a AND NOT (mask (max n m))" for a :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps) lemma ucast_and_neg_mask: "ucast (x AND NOT (mask n)) = ucast x AND NOT (mask n)" apply (rule bit_word_eqI) apply (auto simp add: bit_simps dest: bit_imp_le_length) done lemma ucast_and_mask: "ucast (x AND mask n) = ucast x AND mask n" apply (rule bit_word_eqI) apply (auto simp add: bit_simps dest: bit_imp_le_length) done lemma ucast_mask_drop: "LENGTH('a :: len) \ n \ (ucast (x AND mask n) :: 'a word) = ucast x" apply (rule bit_word_eqI) apply (auto simp add: bit_simps dest: bit_imp_le_length) done lemma mask_exceed: "n \ LENGTH('a) \ (x::'a::len word) AND NOT (mask n) = 0" by (rule bit_word_eqI) (simp add: bit_simps) lemma word_add_no_overflow:"(x::'a::len word) < - 1 \ x < x + 1" using less_x_plus_1 order_less_le by blast lemma lt_plus_1_le_word: fixes x :: "'a::len word" assumes bound:"n < unat (maxBound::'a word)" shows "x < 1 + of_nat n = (x \ of_nat n)" by (metis add.commute bound max_word_max word_Suc_leq word_not_le word_of_nat_less) lemma unat_ucast_up_simp: fixes x :: "'a::len word" assumes "LENGTH('a) \ LENGTH('b)" shows "unat (ucast x :: 'b::len word) = unat x" apply (rule bit_eqI) using assms apply (auto simp add: bit_simps dest: bit_imp_le_length) done lemma unat_ucast_less_no_overflow: "\n < 2 ^ LENGTH('a); unat f < n\ \ (f::('a::len) word) < of_nat n" by (erule (1) order_le_less_trans[OF _ of_nat_mono_maybe,rotated]) simp lemma unat_ucast_less_no_overflow_simp: "n < 2 ^ LENGTH('a) \ (unat f < n) = ((f::('a::len) word) < of_nat n)" using unat_less_helper unat_ucast_less_no_overflow by blast lemma unat_ucast_no_overflow_le: assumes no_overflow: "unat b < (2 :: nat) ^ LENGTH('a)" and upward_cast: "LENGTH('a) < LENGTH('b)" shows "(ucast (f::'a::len word) < (b :: 'b :: len word)) = (unat f < unat b)" proof - have LR: "ucast f < b \ unat f < unat b" apply (rule unat_less_helper) apply (simp add:ucast_nat_def) apply (rule_tac 'b1 = 'b in ucast_less_ucast[OF order.strict_implies_order, THEN iffD1]) apply (rule upward_cast) apply (simp add: ucast_ucast_mask less_mask_eq word_less_nat_alt unat_power_lower[OF upward_cast] no_overflow) done have RL: "unat f < unat b \ ucast f < b" proof- assume ineq: "unat f < unat b" have "ucast (f::'a::len word) < ((ucast (ucast b ::'a::len word)) :: 'b :: len word)" apply (simp add: ucast_less_ucast[OF order.strict_implies_order] upward_cast) apply (simp only: flip: ucast_nat_def) apply (rule unat_ucast_less_no_overflow[OF no_overflow ineq]) done then show ?thesis apply (rule order_less_le_trans) apply (simp add:ucast_ucast_mask word_and_le2) done qed then show ?thesis by (simp add:RL LR iffI) qed lemmas ucast_up_mono = ucast_less_ucast[THEN iffD2] lemma minus_one_word: "(-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1" by simp lemma le_2p_upper_bits: "\ (p::'a::len word) \ 2^n - 1; n < LENGTH('a) \ \ \n'\n. n' < LENGTH('a) \ \ bit p n'" by (subst upper_bits_unset_is_l2p; simp) lemma le2p_bits_unset: "p \ 2 ^ n - 1 \ \n'\n. n' < LENGTH('a) \ \ bit (p::'a::len word) n'" using upper_bits_unset_is_l2p [where p=p] by (cases "n < LENGTH('a)") auto lemma complement_nth_w2p: shows "n' < LENGTH('a) \ bit (NOT (2 ^ n :: 'a::len word)) n' = (n' \ n)" by (fastforce simp: word_ops_nth_size word_size nth_w2p) lemma word_unat_and_lt: "unat x < n \ unat y < n \ unat (x AND y) < n" by (meson le_less_trans word_and_le1 word_and_le2 word_le_nat_alt) lemma word_unat_mask_lt: "m \ size w \ unat ((w::'a::len word) AND mask m) < 2 ^ m" by (rule word_unat_and_lt) (simp add: unat_mask word_size) lemma word_sless_sint_le:"x sint x \ sint y - 1" by (metis word_sless_alt zle_diff1_eq) lemma upper_trivial: fixes x :: "'a::len word" shows "x \ 2 ^ LENGTH('a) - 1 \ x < 2 ^ LENGTH('a) - 1" by (simp add: less_le) lemma constraint_expand: fixes x :: "'a::len word" shows "x \ {y. lower \ y \ y \ upper} = (lower \ x \ x \ upper)" by (rule mem_Collect_eq) lemma card_map_elide: "card ((of_nat :: nat \ 'a::len word) ` {0.. CARD('a::len word)" proof - let ?of_nat = "of_nat :: nat \ 'a word" have "inj_on ?of_nat {i. i < CARD('a word)}" by (rule inj_onI) (simp add: card_word of_nat_inj) moreover have "{0.. {i. i < CARD('a word)}" using that by auto ultimately have "inj_on ?of_nat {0.. CARD('a::len word) \ card ((of_nat::nat \ 'a::len word) ` {0.. LENGTH('a) \ x = ucast y \ ucast x = y" for x :: "'a::len word" and y :: "'b::len word" by transfer simp lemma le_ucast_ucast_le: "x \ ucast y \ ucast x \ y" for x :: "'a::len word" and y :: "'b::len word" by (smt le_unat_uoi linorder_not_less order_less_imp_le ucast_nat_def unat_arith_simps(1)) lemma less_ucast_ucast_less: "LENGTH('b) \ LENGTH('a) \ x < ucast y \ ucast x < y" for x :: "'a::len word" and y :: "'b::len word" by (metis ucast_nat_def unat_mono unat_ucast_up_simp word_of_nat_less) lemma ucast_le_ucast: "LENGTH('a) \ LENGTH('b) \ (ucast x \ (ucast y::'b::len word)) = (x \ y)" for x :: "'a::len word" by (simp add: unat_arith_simps(1) unat_ucast_up_simp) lemmas ucast_up_mono_le = ucast_le_ucast[THEN iffD2] lemma ucast_or_distrib: fixes x :: "'a::len word" fixes y :: "'a::len word" shows "(ucast (x OR y) :: ('b::len) word) = ucast x OR ucast y" by (fact unsigned_or_eq) lemma word_exists_nth: "(w::'a::len word) \ 0 \ \i. bit w i" by (auto simp add: bit_eq_iff) lemma max_word_not_0 [simp]: "- 1 \ (0 :: 'a::len word)" by simp lemma unat_max_word_pos[simp]: "0 < unat (- 1 :: 'a::len word)" using unat_gt_0 [of \- 1 :: 'a::len word\] by simp (* Miscellaneous conditional injectivity rules. *) lemma mult_pow2_inj: assumes ws: "m + n \ LENGTH('a)" assumes le: "x \ mask m" "y \ mask m" assumes eq: "x * 2 ^ n = y * (2 ^ n::'a::len word)" shows "x = y" proof (rule bit_word_eqI) fix q assume \q < LENGTH('a)\ from eq have \push_bit n x = push_bit n y\ by (simp add: push_bit_eq_mult) moreover from le have \take_bit m x = x\ \take_bit m y = y\ by (simp_all add: less_eq_mask_iff_take_bit_eq_self) ultimately have \push_bit n (take_bit m x) = push_bit n (take_bit m y)\ by simp_all with \q < LENGTH('a)\ ws show \bit x q \ bit y q\ apply (simp add: push_bit_take_bit) unfolding bit_eq_iff apply (simp add: bit_simps not_le) apply (metis (full_types) \take_bit m x = x\ \take_bit m y = y\ add.commute add_diff_cancel_right' add_less_cancel_right bit_take_bit_iff le_add2 less_le_trans) done qed lemma word_of_nat_inj: assumes bounded: "x < 2 ^ LENGTH('a)" "y < 2 ^ LENGTH('a)" assumes of_nats: "of_nat x = (of_nat y :: 'a::len word)" shows "x = y" by (rule contrapos_pp[OF of_nats]; cases "x < y"; cases "y < x") (auto dest: bounded[THEN of_nat_mono_maybe]) lemma word_of_int_bin_cat_eq_iff: "(word_of_int (concat_bit LENGTH('b) (uint b) (uint a))::'c::len word) = word_of_int (concat_bit LENGTH('b) (uint d) (uint c)) \ b = d \ a = c" if "LENGTH('a) + LENGTH('b) \ LENGTH('c)" for a::"'a::len word" and b::"'b::len word" proof - from that show ?thesis using that concat_bit_eq_iff [of \LENGTH('b)\ \uint b\ \uint a\ \uint d\ \uint c\] apply (simp add: word_of_int_eq_iff take_bit_int_eq_self flip: word_eq_iff_unsigned) apply (simp add: concat_bit_def take_bit_int_eq_self bintr_uint take_bit_push_bit) done qed lemma word_cat_inj: "(word_cat a b::'c::len word) = word_cat c d \ a = c \ b = d" if "LENGTH('a) + LENGTH('b) \ LENGTH('c)" for a::"'a::len word" and b::"'b::len word" using word_of_int_bin_cat_eq_iff [OF that, of b a d c] by (simp add: word_cat_eq' ac_simps) lemma p2_eq_1: "2 ^ n = (1::'a::len word) \ n = 0" proof - have "2 ^ n = (1::'a word) \ n = 0" by (metis One_nat_def not_less one_less_numeral_iff p2_eq_0 p2_gt_0 power_0 power_0 power_inject_exp semiring_norm(76) unat_power_lower zero_neq_one) then show ?thesis by auto qed end lemmas word_div_less = div_word_less (* FIXME: move to Word distribution? *) lemma bin_nth_minus_Bit0[simp]: "0 < n \ bit (numeral (num.Bit0 w) :: int) n = bit (numeral w :: int) (n - 1)" by (cases n; simp) lemma bin_nth_minus_Bit1[simp]: "0 < n \ bit (numeral (num.Bit1 w) :: int) n = bit (numeral w :: int) (n - 1)" by (cases n; simp) lemma word_mod_by_0: "k mod (0::'a::len word) = k" by (simp add: word_arith_nat_mod) end