# HG changeset patch # User nipkow # Date 1217588270 -7200 # Node ID 087db5aacac369ea7fb408fd02490ea2b20f47c1 # Parent 27b4d7c01f8b50cd5edda141b0e8278e5e713f45 made setsum executable on int. diff -r 27b4d7c01f8b -r 087db5aacac3 src/HOL/List.thy --- a/src/HOL/List.thy Thu Jul 31 09:49:21 2008 +0200 +++ b/src/HOL/List.thy Fri Aug 01 12:57:50 2008 +0200 @@ -2926,9 +2926,29 @@ apply (metis successor_incr leD less_imp_le order_trans) done +lemma sorted_list_of_set_rec2: "i \ j \ + sorted_list_of_set {i..successor j} = + sorted_list_of_set {i..j} @ [successor j]" +apply(simp add:sorted_list_of_set_def upto_def) +apply (rule the1_equality[OF finite_sorted_distinct_unique]) + apply (simp add:finite_intvl) +apply(rule the1I2[OF finite_sorted_distinct_unique]) + apply (simp add:finite_intvl) +apply (simp add: sorted_append Ball_def expand_set_eq) +apply(rule conjI) +apply (metis eq_iff leD linear not_leE ord_discrete order_trans successor_incr) +apply (metis leD linear order_trans successor_incr) +done + lemma upto_rec[code]: "[i..j] = (if i \ j then i # [successor i..j] else [])" by(simp add: upto_def sorted_list_of_set_rec) +lemma upto_empty[simp]: "j < i \ [i..j] = []" +by(simp add: upto_rec) + +lemma upto_rec2: "i \ j \ [i..successor j] = [i..j] @ [successor j]" +by(simp add: upto_def sorted_list_of_set_rec2) + end text{* The integers are an instance of the above class: *} @@ -2937,10 +2957,10 @@ begin definition - successor_int_def: "successor = (%i\int. i+1)" +successor_int_def: "successor = (%i\int. i+1)" instance - by intro_classes (simp_all add: successor_int_def) +by intro_classes (simp_all add: successor_int_def) end @@ -2948,6 +2968,9 @@ hide (open) const successor +lemma upto_rec2_int: "(i::int) \ j \ [i..j+1] = [i..j] @ [j+1]" +by(rule upto_rec2[where 'a = int, simplified successor_int_def]) + subsubsection {* @{text lists}: the list-forming operator over sets *} @@ -2956,7 +2979,7 @@ for A :: "'a set" where Nil [intro!]: "[]: lists A" - | Cons [intro!,noatp]: "[| a: A;l: lists A|] ==> a#l : lists A" + | Cons [intro!,noatp]: "[| a: A; l: lists A|] ==> a#l : lists A" inductive_cases listsE [elim!,noatp]: "x#l : lists A" inductive_cases listspE [elim!,noatp]: "listsp A (x # l)" @@ -3619,11 +3642,11 @@ "{n..m\n\nat. P m) \ (\m \ {0..n}. P m)" by auto +lemma setsum_set_distinct_conv_listsum: + "distinct xs \ setsum f (set xs) = listsum (map f xs)" +by (induct xs) simp_all + lemma setsum_set_upt_conv_listsum [code unfold]: - "setsum f (set [k..