diff -r 206d779ba96d -r ee6cd48cf840 src/HOL/SetInterval.thy --- a/src/HOL/SetInterval.thy Mon Feb 21 18:04:28 2005 +0100 +++ b/src/HOL/SetInterval.thy Mon Feb 21 19:23:46 2005 +0100 @@ -534,6 +534,24 @@ lemmas ivl_disj_int = ivl_disj_int_singleton ivl_disj_int_one ivl_disj_int_two +subsubsection {* Some Differences *} + +lemma ivl_diff[simp]: + "i \ n \ {i.. {m.. i | m \ i & j \ (n::'a::linorder))" +apply(auto simp:linorder_not_le) +apply(rule ccontr) +apply(insert linorder_le_less_linear[of i n]) +apply(clarsimp simp:linorder_not_le) +apply(fastsimp) +done + subsection {* Summation indexed over intervals *} @@ -585,6 +603,17 @@ not provide all lemmas available for @{term"{m..a = c; b = d; !!x. \ c \ x; x < d \ \ f x = g x \ \ + setsum f {a..i < Suc n. b i) = b n + (\i < n. b i)" by (simp add:lessThan_Suc)