# HG changeset patch # User paulson # Date 1568910248 -3600 # Node ID e4825ec2046893c15240bedf6eb1292de0cdcb59 # Parent 91587befabfd945eb3554d26179e9718564cbf8e Tidying and one more theorem diff -r 91587befabfd -r e4825ec20468 src/HOL/Analysis/Derivative.thy --- a/src/HOL/Analysis/Derivative.thy Thu Sep 19 14:49:19 2019 +0100 +++ b/src/HOL/Analysis/Derivative.thy Thu Sep 19 17:24:08 2019 +0100 @@ -239,6 +239,9 @@ lemma frechet_derivative_id [simp]: "frechet_derivative id (at a) = id" using differentiable_def frechet_derivative_works has_derivative_id has_derivative_unique by blast +lemma frechet_derivative_ident [simp]: "frechet_derivative (\x. x) (at a) = (\x. x)" + by (metis eq_id_iff frechet_derivative_id) + subsection \Differentiability implies continuity\ diff -r 91587befabfd -r e4825ec20468 src/HOL/Library/Boolean_Algebra.thy --- a/src/HOL/Library/Boolean_Algebra.thy Thu Sep 19 14:49:19 2019 +0100 +++ b/src/HOL/Library/Boolean_Algebra.thy Thu Sep 19 17:24:08 2019 +0100 @@ -102,17 +102,7 @@ subsection \Conjunction\ lemma conj_zero_right [simp]: "x \<^bold>\ \ = \" -proof - - from conj_cancel_right have "x \<^bold>\ \ = x \<^bold>\ (x \<^bold>\ \ x)" - by simp - also from conj_assoc have "\ = (x \<^bold>\ x) \<^bold>\ \ x" - by (simp only: ac_simps) - also from conj_absorb have "\ = x \<^bold>\ \ x" - by simp - also have "\ = \" - by simp - finally show ?thesis . -qed + using conj.left_idem conj_cancel_right by fastforce lemma compl_one [simp]: "\ \ = \" by (rule compl_unique [OF conj_zero_right disj_zero_right]) @@ -241,7 +231,7 @@ lemmas xor_ac = xor.assoc xor.commute xor.left_commute lemma xor_def2: "x \ y = (x \<^bold>\ y) \<^bold>\ (\ x \<^bold>\ \ y)" - by (simp only: xor_def conj_disj_distribs disj_ac conj_ac conj_cancel_right disj_zero_left) + using conj.commute conj_disj_distrib2 disj.commute xor_def by auto lemma xor_zero_right [simp]: "x \ \ = x" by (simp only: xor_def compl_zero conj_one_right conj_zero_right disj_zero_right) @@ -262,20 +252,10 @@ by (simp only: xor_assoc [symmetric] xor_self xor_zero_left) lemma xor_compl_left [simp]: "\ x \ y = \ (x \ y)" - apply (simp only: xor_def de_Morgan_disj de_Morgan_conj double_compl) - apply (simp only: conj_disj_distribs) - apply (simp only: conj_cancel_right conj_cancel_left) - apply (simp only: disj_zero_left disj_zero_right) - apply (simp only: disj_ac conj_ac) - done + by (metis xor_assoc xor_one_left) lemma xor_compl_right [simp]: "x \ \ y = \ (x \ y)" - apply (simp only: xor_def de_Morgan_disj de_Morgan_conj double_compl) - apply (simp only: conj_disj_distribs) - apply (simp only: conj_cancel_right conj_cancel_left) - apply (simp only: disj_zero_left disj_zero_right) - apply (simp only: disj_ac conj_ac) - done + using xor_commute xor_compl_left by auto lemma xor_cancel_right: "x \ \ x = \" by (simp only: xor_compl_right xor_self compl_zero) @@ -295,12 +275,7 @@ qed lemma conj_xor_distrib2: "(y \ z) \<^bold>\ x = (y \<^bold>\ x) \ (z \<^bold>\ x)" -proof - - have "x \<^bold>\ (y \ z) = (x \<^bold>\ y) \ (x \<^bold>\ z)" - by (rule conj_xor_distrib) - then show "(y \ z) \<^bold>\ x = (y \<^bold>\ x) \ (z \<^bold>\ x)" - by (simp only: conj_commute) -qed + by (simp add: conj.commute conj_xor_distrib) lemmas conj_xor_distribs = conj_xor_distrib conj_xor_distrib2