# HG changeset patch # User haftmann # Date 1748721068 -7200 # Node ID f4441890dee082a09726d4e621717182442f069f # Parent 55260840696f4c025ccda4268fcaa98eb198104b generic executable ranges diff -r 55260840696f -r f4441890dee0 NEWS --- a/NEWS Sat May 31 11:29:10 2025 +0200 +++ b/NEWS Sat May 31 21:51:08 2025 +0200 @@ -92,6 +92,9 @@ const [List.]can_select ~> Set.can_select thm can_select_def ~> Set.can_select_iff [simp] + const List.all_interval_nat List.all_interval_int + discontinued in favour of a generic List.all_range + The _def suffix for characteristic theorems is avoided to emphasize that these auxiliary operations are candidates for unfolding since they are variants of existing logical concepts. The [simp] declarations try to move the attention diff -r 55260840696f -r f4441890dee0 src/HOL/Groups_List.thy --- a/src/HOL/Groups_List.thy Sat May 31 11:29:10 2025 +0200 +++ b/src/HOL/Groups_List.thy Sat May 31 21:51:08 2025 +0200 @@ -57,6 +57,23 @@ finally show ?thesis . qed +lemma atLeastLessThan_conv_list [code_unfold]: + \set.F g {a.. + using List.atLeastLessThan_eq_range + by (simp flip: List.set_range_eq add: distinct_set_conv_list) + +lemma atLeastAtMost_conv_list [code_unfold]: + \set.F g {a..b} = list.F (map g (List.range a (b + 1)))\ + by (simp flip: List.set_range_eq add: List.atLeastAtMost_eq_range distinct_set_conv_list) + +lemma greaterThanLessThan_conv_list [code_unfold]: + \set.F g {a<.. + by (simp flip: List.set_range_eq add: List.greaterThanLessThan_eq_range distinct_set_conv_list) + +lemma greaterThanAtMost_conv_list [code_unfold]: + \set.F g {a<..b} = list.F (map g (List.range (a + 1) (b + 1)))\ + by (simp flip: List.set_range_eq add: List.greaterThanAtMost_eq_range distinct_set_conv_list) + end @@ -591,11 +608,11 @@ lemmas sum_code = sum.set_conv_list -lemma sum_set_upto_conv_sum_list_int [code_unfold]: +lemma sum_set_upto_conv_sum_list_int: "sum f (set [i..j::int]) = sum_list (map f [i..j])" by (simp add: interv_sum_list_conv_sum_set_int) -lemma sum_set_upt_conv_sum_list_nat [code_unfold]: +lemma sum_set_upt_conv_sum_list_nat: "sum f (set [m..{..n} = set [0.. + by auto + +lemma atLeast_upt: + \{.. + by auto + +lemma greaterThanLessThan_upt: + \{n<.. + by auto + +lemma atLeastLessThan_upt: + \{i.. + by auto + +lemma greaterThanAtMost_upt: + "{n<..m} = set [Suc n.. set xs \ \ys zs. xs = ys @ x # zs" proof (induct xs) @@ -6585,6 +6608,10 @@ "sorted_list_of_set {..Suc k} = sorted_list_of_set {..k} @ [Suc k]" using lessThan_Suc_atMost sorted_list_of_set_lessThan_Suc by fastforce +lemma sorted_lift_of_set_eq_upto [simp]: + \sorted_list_of_set {k..l} = [k..l]\ + by (rule sorted_distinct_set_unique) simp_all + lemma sorted_list_of_set_nonempty: assumes "finite A" "A \ {}" shows "sorted_list_of_set A = Min A # sorted_list_of_set (A - {Min A})" @@ -7943,116 +7970,133 @@ by (simp_all add: add: lexordp_def) -text \Intervals over \<^typ>\nat\.\ - -lemma atMost_upto [code_unfold]: - \{..n} = set [0.. - by auto - -lemma atLeast_upt [code_unfold]: - \{.. - by auto - -lemma greaterThanLessThan_upt [code_unfold]: - \{n<.. - by auto - -lemma atLeastLessThan_upt [code_unfold]: - \{i.. - by auto - -lemma greaterThanAtMost_upt [code_unfold]: - "{n<..m} = set [Suc n..Optimized code for \\n:{a.. and similiarly for \\\.\ - -definition all_interval_nat :: \(nat \ bool) \ nat \ nat \ bool\ - where \all_interval_nat P i j \ (\n \ {i.. - -lemma [code]: - \all_interval_nat P i j \ i \ j \ P i \ all_interval_nat P (Suc i) j\ -proof - - have *: \\n. P i \ \n\{Suc i.. i \ n \ n < j \ P n\ - using le_less_Suc_eq by fastforce +text \Executable ranges\ + +class discrete_linordered_semidom_with_range = discrete_linordered_semidom + + assumes finite_atLeastLessThan: \finite {a.. +begin + +context +begin + +qualified definition range :: \'a \ 'a \ 'a list\ \ \only for code generation\ + where range_eq: \range a b = sorted_list_of_set {a.. + +qualified lemma set_range_eq [simp]: + \set (range a b) = {a.. + using finite_atLeastLessThan by (simp add: range_eq) + +qualified lemma distinct_range [simp]: + \distinct (range a b)\ + by (simp add: range_eq) + +qualified lemma range_code [code]: + \range a b = (if a < b then a # range (a + 1) b else [])\ +proof (cases \a < b\) + case True + then have \{a.. + by (auto simp flip: less_iff_succ_less_eq less_eq_iff_succ_less) + then show ?thesis + by (auto simp add: range_eq finite_atLeastLessThan intro!: insort_is_Cons) +next + case False then show ?thesis - by (auto simp add: all_interval_nat_def) -qed - -lemma list_all_iff_all_interval_nat [code_unfold]: - \list_all P [i.. all_interval_nat P i j\ - by (simp add: list_all_iff all_interval_nat_def) - -lemma list_ex_iff_not_all_inverval_nat [code_unfold]: - \list_ex P [i.. \ (all_interval_nat (Not \ P) i j)\ - by (simp add: list_ex_iff all_interval_nat_def) - -hide_const (open) all_interval_nat + by (simp add: range_eq) +qed + +qualified lemma atLeastLessThan_eq_range [code]: + \{a.. + by simp + +qualified lemma atLeastAtMost_eq_range [code]: + \{a..b} = set (range a (b + 1))\ + by (auto simp flip: less_eq_iff_succ_less) + +qualified lemma greaterThanLessThan_eq_range [code]: + \{a<.. + by (auto simp add: less_eq_iff_succ_less) + +qualified lemma greaterThanAtMost_eq_range [code]: + \{a<..b} = set (range (a + 1) (b + 1))\ + by (auto simp add: less_eq_iff_succ_less) + +qualified definition all_range :: \('a \ bool) \ 'a \ 'a \ bool\ \ \only for code generation\ + where all_range_iff [code_abbrev, simp]: \all_range P a b \ (\n\{a.. + +qualified lemma all_range_code [code]: + \all_range P a b \ (a < b \ P a \ all_range P (a + 1) b)\ + by (auto simp add: Ball_def simp flip: less_iff_succ_less_eq) + (auto simp add: le_less) + +lemma forall_atLeastAtMost_iff [code_unfold]: + \(\n\{a..b}. P n) \ all_range P a (b + 1)\ + by (simp add: atLeastAtMost_eq_range all_range_iff) + +lemma forall_greaterThanLessThan_iff [code_unfold]: + \(\n\{a<.. all_range P (a + 1) b\ + by (simp add: greaterThanLessThan_eq_range) + +lemma forall_greaterThanAtMost_iff [code_unfold]: + \(\n\{a<..b}. P n) \ all_range P (a + 1) (b + 1)\ + by (simp add: greaterThanAtMost_eq_range) + +lemma exists_atLeastLessThan [code_unfold]: + \(\n\{a.. \ all_range (Not \ P) a b\ + by simp + +lemma exists_atLeastAtMost_iff [code_unfold]: + \(\n\{a..b}. P n) \ \ all_range (Not \ P) a (b + 1)\ + by (simp add: atLeastAtMost_eq_range) + +lemma exists_greaterThanLessThan_iff [code_unfold]: + \(\n\{a<.. \ all_range (Not \ P) (a + 1) b\ + by (simp add: greaterThanLessThan_eq_range) + +lemma exists_greaterThanAtMost_iff [code_unfold]: + \(\n\{a<..b}. P n) \ \ all_range (Not \ P) (a + 1) (b + 1)\ + by (simp add: greaterThanAtMost_eq_range) + +end + +end + +hide_const (open) range \ \TODO: should not be necessary\ + +instance nat :: discrete_linordered_semidom_with_range + by standard auto + +instance int :: discrete_linordered_semidom_with_range + by standard auto lemma all_nat_less_eq [code_unfold]: - "(\m (\m \ {0..m (\m \ {0..m\n. P m) \ (\m \ {0..n}. P m)" for n :: nat by auto lemma ex_nat_less_eq [code_unfold]: - "(\m (\m \ {0..m\n::nat. P m) \ (\m \ {0..n}. P m)" + "(\m (\m \ {0..m\n::nat. P m) \ (\m \ {0..n}. P m)" - by auto - - -text \Intervals over \<^typ>\int\.\ - -lemma greaterThanLessThan_upto [code_unfold]: - \{i<.. - by auto - -lemma atLeastLessThan_upto [code_unfold]: - \{i.. - by auto - -lemma greaterThanAtMost_upto [code_unfold]: - \{i<..j::int} = set [i+1..j]\ - by auto - -lemma atLeastAtMost_upto [code_unfold]: - \{i..j} = set [i..j]\ + "(\m\n. P m) \ (\m \ {0..n}. P m)" for n :: nat by auto - -text \Optimized code for \\i\{a..b::int}\ and similiarly for \\\.\ - -definition all_interval_int :: \(int \ bool) \ int \ int \ bool\ - where \all_interval_int P i j \ (\k \ {i..j}. P k)\ - -lemma [code]: - \all_interval_int P i j \ i > j \ P i \ all_interval_int P (i + 1) j\ -proof - - have *: \\k. P i \ \k\{i+1..j}. P k \ i \ k \ k \ j \ P k\ - by (smt (verit, best) atLeastAtMost_iff) - show ?thesis - by (auto simp add: all_interval_int_def intro: *) -qed - -lemma list_all_iff_all_interval_int [code_unfold]: - \list_all P [i..j] \ all_interval_int P i j\ - by (simp add: list_all_iff all_interval_int_def) - -lemma list_ex_iff_not_all_inverval_int [code_unfold]: - \list_ex P [i..j] \ \ (all_interval_int (Not \ P) i j)\ - by (simp add: list_ex_iff all_interval_int_def) - -hide_const (open) all_interval_int +context +begin + +qualified lemma range_eq_upt [simp]: + \List.range m n = [m.. + by (simp add: List.range_eq) + +qualified lemma range_eq_upto [simp]: + \List.range i k = [i..k - 1]\ + using atLeastLessThanPlusOne_atLeastAtMost_int [of i \k - 1\] + by (simp add: List.range_eq) + +end text \Bounded \LEAST\ operator.\