--- a/src/Doc/Tutorial/Inductive/AB.thy Wed Jun 06 14:22:54 2018 +0200
+++ b/src/Doc/Tutorial/Inductive/AB.thy Wed Jun 06 14:23:13 2018 +0200
@@ -64,13 +64,15 @@
\<close>
lemma correctness:
- "(w \<in> S \<longrightarrow> size(filter (\<lambda>x. x=a) w) = size(filter (\<lambda>x. x=b) w)) \<and>
- (w \<in> A \<longrightarrow> size(filter (\<lambda>x. x=a) w) = size(filter (\<lambda>x. x=b) w) + 1) \<and>
- (w \<in> B \<longrightarrow> size(filter (\<lambda>x. x=b) w) = size(filter (\<lambda>x. x=a) w) + 1)"
+ "(w \<in> S \<longrightarrow> size[x\<leftarrow>w. x=a] = size[x\<leftarrow>w. x=b]) \<and>
+ (w \<in> A \<longrightarrow> size[x\<leftarrow>w. x=a] = size[x\<leftarrow>w. x=b] + 1) \<and>
+ (w \<in> B \<longrightarrow> size[x\<leftarrow>w. x=b] = size[x\<leftarrow>w. x=a] + 1)"
txt\<open>\noindent
These propositions are expressed with the help of the predefined @{term
-filter} function on lists. Remember that on lists @{text size} and @{text length} are synonymous.
+filter} function on lists, which has the convenient syntax @{text"[x\<leftarrow>xs. P
+x]"}, the list of all elements @{term x} in @{term xs} such that @{prop"P x"}
+holds. Remember that on lists @{text size} and @{text length} are synonymous.
The proof itself is by rule induction and afterwards automatic:
\<close>
@@ -110,8 +112,8 @@
\<close>
lemma step1: "\<forall>i < size w.
- \<bar>(int(size(filter P (take (i+1) w)))-int(size(filter (\<lambda>x. \<not>P x) (take (i+1) w))))
- - (int(size(filter P (take i w)))-int(size(filter (\<lambda>x. \<not>P x) (take i w))))\<bar> \<le> 1"
+ \<bar>(int(size[x\<leftarrow>take (i+1) w. P x])-int(size[x\<leftarrow>take (i+1) w. \<not>P x]))
+ - (int(size[x\<leftarrow>take i w. P x])-int(size[x\<leftarrow>take i w. \<not>P x]))\<bar> \<le> 1"
txt\<open>\noindent
The lemma is a bit hard to read because of the coercion function
@@ -135,8 +137,8 @@
\<close>
lemma part1:
- "size(filter P w) = size(filter (\<lambda>x. \<not>P x) w)+2 \<Longrightarrow>
- \<exists>i\<le>size w. size(filter P (take i w)) = size(filter (\<lambda>x. \<not>P x) (take i w))+1"
+ "size[x\<leftarrow>w. P x] = size[x\<leftarrow>w. \<not>P x]+2 \<Longrightarrow>
+ \<exists>i\<le>size w. size[x\<leftarrow>take i w. P x] = size[x\<leftarrow>take i w. \<not>P x]+1"
txt\<open>\noindent
This is proved by @{text force} with the help of the intermediate value theorem,
@@ -155,10 +157,10 @@
lemma part2:
- "\<lbrakk>size(filter P (take i w @ drop i w)) =
- size(filter (\<lambda>x. \<not>P x) (take i w @ drop i w))+2;
- size(filter P (take i w)) = size(filter (\<lambda>x. \<not>P x) (take i w))+1\<rbrakk>
- \<Longrightarrow> size(filter P (drop i w)) = size(filter (\<lambda>x. \<not>P x) (drop i w))+1"
+ "\<lbrakk>size[x\<leftarrow>take i w @ drop i w. P x] =
+ size[x\<leftarrow>take i w @ drop i w. \<not>P x]+2;
+ size[x\<leftarrow>take i w. P x] = size[x\<leftarrow>take i w. \<not>P x]+1\<rbrakk>
+ \<Longrightarrow> size[x\<leftarrow>drop i w. P x] = size[x\<leftarrow>drop i w. \<not>P x]+1"
by(simp del: append_take_drop_id)
text\<open>\noindent
@@ -185,9 +187,9 @@
\<close>
theorem completeness:
- "(size(filter (\<lambda>x. x=a) w) = size(filter (\<lambda>x. x=b) w) \<longrightarrow> w \<in> S) \<and>
- (size(filter (\<lambda>x. x=a) w) = size(filter (\<lambda>x. x=b) w) + 1 \<longrightarrow> w \<in> A) \<and>
- (size(filter (\<lambda>x. x=b) w) = size(filter (\<lambda>x. x=a) w) + 1 \<longrightarrow> w \<in> B)"
+ "(size[x\<leftarrow>w. x=a] = size[x\<leftarrow>w. x=b] \<longrightarrow> w \<in> S) \<and>
+ (size[x\<leftarrow>w. x=a] = size[x\<leftarrow>w. x=b] + 1 \<longrightarrow> w \<in> A) \<and>
+ (size[x\<leftarrow>w. x=b] = size[x\<leftarrow>w. x=a] + 1 \<longrightarrow> w \<in> B)"
txt\<open>\noindent
The proof is by induction on @{term w}. Structural induction would fail here
@@ -232,9 +234,9 @@
apply(clarify)
txt\<open>\noindent
This yields an index @{prop"i \<le> length v"} such that
-@{prop[display]"length (filter (\<lambda>x. x=a) (take i v)) = length (filter (\<lambda>x. x=b) (take i v)) + 1"}
+@{prop[display]"length [x\<leftarrow>take i v . x = a] = length [x\<leftarrow>take i v . x = b] + 1"}
With the help of @{thm[source]part2} it follows that
-@{prop[display]"length (filter (\<lambda>x. x=a) (drop i v)) = length (filter (\<lambda>x. x=b) (drop i v)) + 1"}
+@{prop[display]"length [x\<leftarrow>drop i v . x = a] = length [x\<leftarrow>drop i v . x = b] + 1"}
\<close>
apply(drule part2[of "\<lambda>x. x=a", simplified])
--- a/src/HOL/Library/AList.thy Wed Jun 06 14:22:54 2018 +0200
+++ b/src/HOL/Library/AList.thy Wed Jun 06 14:23:13 2018 +0200
@@ -45,7 +45,7 @@
using assms by (simp add: update_keys)
lemma update_filter:
- "a \<noteq> k \<Longrightarrow> update k v (filter (\<lambda>q. fst q \<noteq> a) ps) = filter (\<lambda>q. fst q \<noteq> a) (update k v ps)"
+ "a \<noteq> k \<Longrightarrow> update k v [q\<leftarrow>ps. fst q \<noteq> a] = [q\<leftarrow>update k v ps. fst q \<noteq> a]"
by (induct ps) auto
lemma update_triv: "map_of al k = Some v \<Longrightarrow> update k v al = al"
@@ -361,12 +361,12 @@
proof -
have "y \<noteq> x \<longleftrightarrow> x \<noteq> y" for y
by auto
- then show "filter (\<lambda>p. fst p \<in> D \<and> x \<noteq> fst p) al = filter (\<lambda>p. fst p \<in> D \<and> fst p \<noteq> x) al"
+ then show "[p \<leftarrow> al. fst p \<in> D \<and> x \<noteq> fst p] = [p \<leftarrow> al. fst p \<in> D \<and> fst p \<noteq> x]"
by simp
assume "x \<notin> D"
then have "y \<in> D \<longleftrightarrow> y \<in> D \<and> x \<noteq> y" for y
by auto
- then show "filter (\<lambda>p. fst p \<in> D \<and> x \<noteq> fst p) al = filter (\<lambda>p. fst p \<in> D) al"
+ then show "[p \<leftarrow> al . fst p \<in> D \<and> x \<noteq> fst p] = [p \<leftarrow> al . fst p \<in> D]"
by simp
qed
@@ -492,7 +492,7 @@
lemma distinct_map_ran: "distinct (map fst al) \<Longrightarrow> distinct (map fst (map_ran f al))"
by (simp add: map_ran_def split_def comp_def)
-lemma map_ran_filter: "map_ran f (filter (\<lambda>p. fst p \<noteq> a) ps) = filter (\<lambda>p. fst p \<noteq> a) (map_ran f ps)"
+lemma map_ran_filter: "map_ran f [p\<leftarrow>ps. fst p \<noteq> a] = [p\<leftarrow>map_ran f ps. fst p \<noteq> a]"
by (simp add: map_ran_def filter_map split_def comp_def)
lemma clearjunk_map_ran: "clearjunk (map_ran f al) = map_ran f (clearjunk al)"
--- a/src/HOL/Library/Finite_Map.thy Wed Jun 06 14:22:54 2018 +0200
+++ b/src/HOL/Library/Finite_Map.thy Wed Jun 06 14:23:13 2018 +0200
@@ -84,10 +84,10 @@
definition map_filter :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<rightharpoonup> 'b) \<Rightarrow> ('a \<rightharpoonup> 'b)" where
"map_filter P m = (\<lambda>x. if P x then m x else None)"
-lemma map_filter_map_of[simp]: "map_filter P (map_of m) = map_of (filter (\<lambda>(k, _). P k) m)"
+lemma map_filter_map_of[simp]: "map_filter P (map_of m) = map_of [(k, _) \<leftarrow> m. P k]"
proof
fix x
- show "map_filter P (map_of m) x = map_of (filter (\<lambda>(k, _). P k) m) x"
+ show "map_filter P (map_of m) x = map_of [(k, _) \<leftarrow> m. P k] x"
by (induct m) (auto simp: map_filter_def)
qed
--- a/src/HOL/Library/Multiset.thy Wed Jun 06 14:22:54 2018 +0200
+++ b/src/HOL/Library/Multiset.thy Wed Jun 06 14:23:13 2018 +0200
@@ -1888,7 +1888,7 @@
apply simp
done
-lemma mset_compl_union [simp]: "mset (filter P xs) + mset (filter (\<lambda>x. \<not>P x) xs) = mset xs"
+lemma mset_compl_union [simp]: "mset [x\<leftarrow>xs. P x] + mset [x\<leftarrow>xs. \<not>P x] = mset xs"
by (induct xs) auto
lemma nth_mem_mset: "i < length ls \<Longrightarrow> (ls ! i) \<in># mset ls"
@@ -2582,14 +2582,14 @@
show "mset ys = mset xs" by simp
from \<open>sorted (map f ys)\<close>
show "sorted (map f ys)" .
- show "filter (\<lambda>x. f k = f x) ys = filter (\<lambda>x. f k = f x) xs" if "k \<in> set ys" for k
+ show "[x\<leftarrow>ys . f k = f x] = [x\<leftarrow>xs . f k = f x]" if "k \<in> set ys" for k
proof -
from mset_equal
have set_equal: "set xs = set ys" by (rule mset_eq_setD)
with that have "insert k (set ys) = set ys" by auto
with \<open>inj_on f (set xs)\<close> have inj: "inj_on f (insert k (set ys))"
by (simp add: set_equal)
- from inj have "filter (\<lambda>x. f k = f x) ys = filter (HOL.eq k) ys"
+ from inj have "[x\<leftarrow>ys . f k = f x] = filter (HOL.eq k) ys"
by (auto intro!: inj_on_filter_key_eq)
also have "\<dots> = replicate (count (mset ys) k) k"
by (simp add: replicate_count_mset_eq_filter_eq)
@@ -2597,7 +2597,7 @@
using mset_equal by simp
also have "\<dots> = filter (HOL.eq k) xs"
by (simp add: replicate_count_mset_eq_filter_eq)
- also have "\<dots> = filter (\<lambda>x. f k = f x) xs"
+ also have "\<dots> = [x\<leftarrow>xs . f k = f x]"
using inj by (auto intro!: inj_on_filter_key_eq [symmetric] simp add: set_equal)
finally show ?thesis .
qed
@@ -2610,9 +2610,9 @@
by (rule sort_key_inj_key_eq) (simp_all add: assms)
lemma sort_key_by_quicksort:
- "sort_key f xs = sort_key f (filter (\<lambda>x. f x < f (xs ! (length xs div 2))) xs)
- @ filter (\<lambda>x. f x = f (xs ! (length xs div 2))) xs
- @ sort_key f (filter (\<lambda>x. f x > f (xs ! (length xs div 2))) xs)" (is "sort_key f ?lhs = ?rhs")
+ "sort_key f xs = sort_key f [x\<leftarrow>xs. f x < f (xs ! (length xs div 2))]
+ @ [x\<leftarrow>xs. f x = f (xs ! (length xs div 2))]
+ @ sort_key f [x\<leftarrow>xs. f x > f (xs ! (length xs div 2))]" (is "sort_key f ?lhs = ?rhs")
proof (rule properties_for_sort_key)
show "mset ?rhs = mset ?lhs"
by (rule multiset_eqI) (auto simp add: mset_filter)
@@ -2623,14 +2623,14 @@
assume "l \<in> set ?rhs"
let ?pivot = "f (xs ! (length xs div 2))"
have *: "\<And>x. f l = f x \<longleftrightarrow> f x = f l" by auto
- have "filter (\<lambda>x. f x = f l) (sort_key f xs) = filter (\<lambda>x. f x = f l) xs"
+ have "[x \<leftarrow> sort_key f xs . f x = f l] = [x \<leftarrow> xs. f x = f l]"
unfolding filter_sort by (rule properties_for_sort_key) (auto intro: sorted_map_same)
- with * have **: "filter (\<lambda>x. f l = f x) (sort_key f xs) = filter (\<lambda>x. f l = f x) xs" by simp
+ with * have **: "[x \<leftarrow> sort_key f xs . f l = f x] = [x \<leftarrow> xs. f l = f x]" by simp
have "\<And>x P. P (f x) ?pivot \<and> f l = f x \<longleftrightarrow> P (f l) ?pivot \<and> f l = f x" by auto
- then have "\<And>P. filter (\<lambda>x. P (f x) ?pivot \<and> f l = f x) (sort_key f xs) =
- filter (\<lambda>x. P (f l) ?pivot \<and> f l = f x) (sort_key f xs)" by simp
+ then have "\<And>P. [x \<leftarrow> sort_key f xs . P (f x) ?pivot \<and> f l = f x] =
+ [x \<leftarrow> sort_key f xs. P (f l) ?pivot \<and> f l = f x]" by simp
note *** = this [of "(<)"] this [of "(>)"] this [of "(=)"]
- show "filter (\<lambda>x. f l = f x) ?rhs = filter (\<lambda>x. f l = f x) ?lhs"
+ show "[x \<leftarrow> ?rhs. f l = f x] = [x \<leftarrow> ?lhs. f l = f x]"
proof (cases "f l" ?pivot rule: linorder_cases)
case less
then have "f l \<noteq> ?pivot" and "\<not> f l > ?pivot" by auto
@@ -2648,15 +2648,15 @@
qed
lemma sort_by_quicksort:
- "sort xs = sort (filter (\<lambda>x. x < xs ! (length xs div 2)) xs)
- @ filter (\<lambda>x. x = xs ! (length xs div 2)) xs
- @ sort (filter (\<lambda>x. x > xs ! (length xs div 2)) xs)" (is "sort ?lhs = ?rhs")
+ "sort xs = sort [x\<leftarrow>xs. x < xs ! (length xs div 2)]
+ @ [x\<leftarrow>xs. x = xs ! (length xs div 2)]
+ @ sort [x\<leftarrow>xs. x > xs ! (length xs div 2)]" (is "sort ?lhs = ?rhs")
using sort_key_by_quicksort [of "\<lambda>x. x", symmetric] by simp
text \<open>A stable parametrized quicksort\<close>
definition part :: "('b \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'b list \<Rightarrow> 'b list \<times> 'b list \<times> 'b list" where
- "part f pivot xs = (filter (\<lambda>x. f x < pivot) xs, filter (\<lambda>x. f x = pivot) xs, filter (\<lambda>x. f x > pivot) xs)"
+ "part f pivot xs = ([x \<leftarrow> xs. f x < pivot], [x \<leftarrow> xs. f x = pivot], [x \<leftarrow> xs. pivot < f x])"
lemma part_code [code]:
"part f pivot [] = ([], [], [])"
--- a/src/HOL/MicroJava/DFA/Kildall.thy Wed Jun 06 14:22:54 2018 +0200
+++ b/src/HOL/MicroJava/DFA/Kildall.thy Wed Jun 06 14:23:13 2018 +0200
@@ -38,7 +38,7 @@
lemma (in Semilat) nth_merges:
"\<And>ss. \<lbrakk>p < length ss; ss \<in> list n A; \<forall>(p,t)\<in>set ps. p<n \<and> t\<in>A \<rbrakk> \<Longrightarrow>
- (merges f ps ss)!p = map snd (filter (\<lambda>(p',t'). p'=p) ps) ++_f ss!p"
+ (merges f ps ss)!p = map snd [(p',t') \<leftarrow> ps. p'=p] ++_f ss!p"
(is "\<And>ss. \<lbrakk>_; _; ?steptype ps\<rbrakk> \<Longrightarrow> ?P ss ps")
proof (induct ps)
show "\<And>ss. ?P ss []" by simp
--- a/src/HOL/MicroJava/DFA/LBVComplete.thy Wed Jun 06 14:22:54 2018 +0200
+++ b/src/HOL/MicroJava/DFA/LBVComplete.thy Wed Jun 06 14:23:13 2018 +0200
@@ -80,13 +80,13 @@
assume merge: "?s1 \<noteq> T"
from x ss1 have "?s1 =
(if \<forall>(pc', s')\<in>set ss1. pc' \<noteq> pc + 1 \<longrightarrow> s' <=_r c!pc'
- then (map snd (filter (\<lambda>(p', t'). p'=pc+1) ss1)) ++_f x
+ then (map snd [(p', t') \<leftarrow> ss1 . p'=pc+1]) ++_f x
else \<top>)"
by (rule merge_def)
with merge obtain
app: "\<forall>(pc',s')\<in>set ss1. pc' \<noteq> pc+1 \<longrightarrow> s' <=_r c!pc'"
(is "?app ss1") and
- sum: "(map snd (filter (\<lambda>(p',t'). p' = pc+1) ss1) ++_f x) = ?s1"
+ sum: "(map snd [(p',t') \<leftarrow> ss1 . p' = pc+1] ++_f x) = ?s1"
(is "?map ss1 ++_f x = _" is "?sum ss1 = _")
by (simp split: if_split_asm)
from app less
@@ -115,7 +115,7 @@
from x ss2 have
"?s2 =
(if \<forall>(pc', s')\<in>set ss2. pc' \<noteq> pc + 1 \<longrightarrow> s' <=_r c!pc'
- then map snd (filter (\<lambda>(p', t'). p' = pc + 1) ss2) ++_f x
+ then map snd [(p', t') \<leftarrow> ss2 . p' = pc + 1] ++_f x
else \<top>)"
by (rule merge_def)
ultimately have ?thesis by simp
@@ -194,7 +194,7 @@
ultimately
have "merge c pc ?step (c!Suc pc) =
(if \<forall>(pc',s')\<in>set ?step. pc'\<noteq>pc+1 \<longrightarrow> s' <=_r c!pc'
- then map snd (filter (\<lambda>(p',t').p'=pc+1) ?step) ++_f c!Suc pc
+ then map snd [(p',t') \<leftarrow> ?step.p'=pc+1] ++_f c!Suc pc
else \<top>)" unfolding mrg_def by (rule lbv.merge_def [OF lbvc.axioms(1), OF lbvc_axioms])
moreover {
fix pc' s' assume s': "(pc', s') \<in> set ?step" and suc_pc: "pc' \<noteq> pc+1"
@@ -207,7 +207,7 @@
} hence "\<forall>(pc',s')\<in>set ?step. pc'\<noteq>pc+1 \<longrightarrow> s' <=_r c!pc'" by auto
moreover
from pc have "Suc pc = length \<phi> \<or> Suc pc < length \<phi>" by auto
- hence "map snd (filter (\<lambda>(p',t').p'=pc+1) ?step) ++_f c!Suc pc \<noteq> \<top>"
+ hence "map snd [(p',t') \<leftarrow> ?step.p'=pc+1] ++_f c!Suc pc \<noteq> \<top>"
(is "?map ++_f _ \<noteq> _")
proof (rule disjE)
assume pc': "Suc pc = length \<phi>"
@@ -215,7 +215,7 @@
moreover
from pc' bounded pc
have "\<forall>(p',t')\<in>set ?step. p'\<noteq>pc+1" by clarify (drule boundedD, auto)
- hence "filter (\<lambda>(p',t').p'=pc+1) ?step = []" by (blast intro: filter_False)
+ hence "[(p',t') \<leftarrow> ?step.p'=pc+1] = []" by (blast intro: filter_False)
hence "?map = []" by simp
ultimately show ?thesis by (simp add: B_neq_T)
next
@@ -262,7 +262,7 @@
hence "merge c pc ?step (c!Suc pc) \<noteq> \<top>" by (simp add: wti)
ultimately
have "merge c pc ?step (c!Suc pc) =
- map snd (filter (\<lambda>(p',t').p'=pc+1) ?step) ++_f c!Suc pc" by (rule merge_not_top_s)
+ map snd [(p',t')\<leftarrow> ?step.p'=pc+1] ++_f c!Suc pc" by (rule merge_not_top_s)
hence "?wti = \<dots>" (is "_ = (?map ++_f _)" is "_ = ?sum") by (simp add: wti)
also {
from suc_pc phi have "\<phi>!Suc pc \<in> A" by simp
--- a/src/HOL/MicroJava/DFA/LBVCorrect.thy Wed Jun 06 14:22:54 2018 +0200
+++ b/src/HOL/MicroJava/DFA/LBVCorrect.thy Wed Jun 06 14:23:13 2018 +0200
@@ -88,7 +88,7 @@
also
from s2 merge have "\<dots> \<noteq> \<top>" (is "?merge \<noteq> _") by simp
with cert_in_A step_in_A
- have "?merge = (map snd (filter (\<lambda>(p',t').p'=pc+1) (?step pc)) ++_f (c!(pc+1)))"
+ have "?merge = (map snd [(p',t') \<leftarrow> ?step pc. p'=pc+1] ++_f (c!(pc+1)))"
by (rule merge_not_top_s)
finally
have "s' <=_r ?s2" using step_in_A cert_in_A True step
--- a/src/HOL/MicroJava/DFA/LBVSpec.thy Wed Jun 06 14:22:54 2018 +0200
+++ b/src/HOL/MicroJava/DFA/LBVSpec.thy Wed Jun 06 14:23:13 2018 +0200
@@ -113,7 +113,7 @@
lemma (in Semilat) pp_ub1':
assumes S: "snd`set S \<subseteq> A"
assumes y: "y \<in> A" and ab: "(a, b) \<in> set S"
- shows "b <=_r map snd (filter (\<lambda>(p', t'). p' = a) S) ++_f y"
+ shows "b <=_r map snd [(p', t') \<leftarrow> S . p' = a] ++_f y"
proof -
from S have "\<forall>(x,y) \<in> set S. y \<in> A" by auto
with semilat y ab show ?thesis by - (rule ub1')
@@ -172,7 +172,7 @@
"\<And>x. x \<in> A \<Longrightarrow> snd`set ss \<subseteq> A \<Longrightarrow>
merge c pc ss x =
(if \<forall>(pc',s') \<in> set ss. pc'\<noteq>pc+1 \<longrightarrow> s' <=_r c!pc' then
- map snd (filter (\<lambda>(p',t'). p'=pc+1) ss) ++_f x
+ map snd [(p',t') \<leftarrow> ss. p'=pc+1] ++_f x
else \<top>)"
(is "\<And>x. _ \<Longrightarrow> _ \<Longrightarrow> ?merge ss x = ?if ss x" is "\<And>x. _ \<Longrightarrow> _ \<Longrightarrow> ?P ss x")
proof (induct ss)
@@ -202,15 +202,15 @@
hence "\<forall>(pc', s')\<in>set ls. pc'\<noteq>pc+1 \<longrightarrow> s' <=_r c!pc'" by auto
moreover
from True have
- "map snd (filter (\<lambda>(p',t'). p'=pc+1) ls) ++_f ?if' =
- (map snd (filter (\<lambda>(p',t'). p'=pc+1) (l#ls)) ++_f x)"
+ "map snd [(p',t')\<leftarrow>ls . p'=pc+1] ++_f ?if' =
+ (map snd [(p',t')\<leftarrow>l#ls . p'=pc+1] ++_f x)"
by simp
ultimately
show ?thesis using True by simp
next
case False
moreover
- from ls have "set (map snd (filter (\<lambda>(p', t'). p' = Suc pc) ls)) \<subseteq> A" by auto
+ from ls have "set (map snd [(p', t')\<leftarrow>ls . p' = Suc pc]) \<subseteq> A" by auto
ultimately show ?thesis by auto
qed
finally show "?P (l#ls) x" .
@@ -219,7 +219,7 @@
lemma (in lbv) merge_not_top_s:
assumes x: "x \<in> A" and ss: "snd`set ss \<subseteq> A"
assumes m: "merge c pc ss x \<noteq> \<top>"
- shows "merge c pc ss x = (map snd (filter (\<lambda>(p',t'). p'=pc+1) ss) ++_f x)"
+ shows "merge c pc ss x = (map snd [(p',t') \<leftarrow> ss. p'=pc+1] ++_f x)"
proof -
from ss m have "\<forall>(pc',s') \<in> set ss. (pc' \<noteq> pc+1 \<longrightarrow> s' <=_r c!pc')"
by (rule merge_not_top)
@@ -307,8 +307,8 @@
assumes s0: "snd`set ss \<subseteq> A" and x: "x \<in> A"
shows "merge c pc ss x \<in> A"
proof -
- from s0 have "set (map snd (filter (\<lambda>(p', t'). p'=pc+1) ss)) \<subseteq> A" by auto
- with x have "(map snd (filter (\<lambda>(p', t'). p'=pc+1) ss) ++_f x) \<in> A"
+ from s0 have "set (map snd [(p', t')\<leftarrow>ss . p'=pc+1]) \<subseteq> A" by auto
+ with x have "(map snd [(p', t')\<leftarrow>ss . p'=pc+1] ++_f x) \<in> A"
by (auto intro!: plusplus_closed semilat)
with s0 x show ?thesis by (simp add: merge_def T_A)
qed
--- a/src/HOL/MicroJava/DFA/SemilatAlg.thy Wed Jun 06 14:22:54 2018 +0200
+++ b/src/HOL/MicroJava/DFA/SemilatAlg.thy Wed Jun 06 14:23:13 2018 +0200
@@ -155,7 +155,7 @@
lemma ub1':
assumes "semilat (A, r, f)"
shows "\<lbrakk>\<forall>(p,s) \<in> set S. s \<in> A; y \<in> A; (a,b) \<in> set S\<rbrakk>
- \<Longrightarrow> b <=_r map snd (filter (\<lambda>(p', t'). p' = a) S) ++_f y"
+ \<Longrightarrow> b <=_r map snd [(p', t')\<leftarrow>S. p' = a] ++_f y"
proof -
interpret Semilat A r f using assms by (rule Semilat.intro)
@@ -175,7 +175,7 @@
lemma plusplus_empty:
"\<forall>s'. (q, s') \<in> set S \<longrightarrow> s' +_f ss ! q = ss ! q \<Longrightarrow>
- (map snd (filter (\<lambda>(p', t'). p' = q) S) ++_f ss ! q) = ss ! q"
+ (map snd [(p', t') \<leftarrow> S. p' = q] ++_f ss ! q) = ss ! q"
by (induct S) auto
end
--- a/src/HOL/Nitpick_Examples/Manual_Nits.thy Wed Jun 06 14:22:54 2018 +0200
+++ b/src/HOL/Nitpick_Examples/Manual_Nits.thy Wed Jun 06 14:23:13 2018 +0200
@@ -317,7 +317,7 @@
| "\<lbrakk>v \<in> B\<^sub>1; v \<in> B\<^sub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>1"
theorem S\<^sub>1_sound:
-"w \<in> S\<^sub>1 \<longrightarrow> length (filter (\<lambda>x. x = a) w) = length (filter (\<lambda>x. x = b) w)"
+"w \<in> S\<^sub>1 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
nitpick [expect = genuine]
oops
@@ -330,7 +330,7 @@
| "\<lbrakk>v \<in> B\<^sub>2; v \<in> B\<^sub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>2"
theorem S\<^sub>2_sound:
-"w \<in> S\<^sub>2 \<longrightarrow> length (filter (\<lambda>x. x = a) w) = length (filter (\<lambda>x. x = b) w)"
+"w \<in> S\<^sub>2 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
nitpick [expect = genuine]
oops
@@ -343,12 +343,12 @@
| "\<lbrakk>v \<in> B\<^sub>3; w \<in> B\<^sub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>3"
theorem S\<^sub>3_sound:
-"w \<in> S\<^sub>3 \<longrightarrow> length (filter (\<lambda>x. x = a) w) = length (filter (\<lambda>x. x = b) w)"
+"w \<in> S\<^sub>3 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
nitpick [card = 1-5, expect = none]
sorry
theorem S\<^sub>3_complete:
-"length (filter (\<lambda>x. x = a) w) = length (filter (\<lambda>x. x = b) w) \<longrightarrow> w \<in> S\<^sub>3"
+"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^sub>3"
nitpick [expect = genuine]
oops
@@ -362,19 +362,19 @@
| "\<lbrakk>v \<in> B\<^sub>4; w \<in> B\<^sub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>4"
theorem S\<^sub>4_sound:
-"w \<in> S\<^sub>4 \<longrightarrow> length (filter (\<lambda>x. x = a) w) = length (filter (\<lambda>x. x = b) w)"
+"w \<in> S\<^sub>4 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
nitpick [card = 1-5, expect = none]
sorry
theorem S\<^sub>4_complete:
-"length (filter (\<lambda>x. x = a) w) = length (filter (\<lambda>x. x = b) w) \<longrightarrow> w \<in> S\<^sub>4"
+"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^sub>4"
nitpick [card = 1-5, expect = none]
sorry
theorem S\<^sub>4_A\<^sub>4_B\<^sub>4_sound_and_complete:
-"w \<in> S\<^sub>4 \<longleftrightarrow> length (filter (\<lambda>x. x = a) w) = length (filter (\<lambda>x. x = b) w)"
-"w \<in> A\<^sub>4 \<longleftrightarrow> length (filter (\<lambda>x. x = a) w) = length (filter (\<lambda>x. x = b) w) + 1"
-"w \<in> B\<^sub>4 \<longleftrightarrow> length (filter (\<lambda>x. x = b) w) = length (filter (\<lambda>x. x = a) w) + 1"
+"w \<in> S\<^sub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+"w \<in> A\<^sub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] + 1"
+"w \<in> B\<^sub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = b] = length [x \<leftarrow> w. x = a] + 1"
nitpick [card = 1-5, expect = none]
sorry
--- a/src/HOL/Nominal/Examples/W.thy Wed Jun 06 14:22:54 2018 +0200
+++ b/src/HOL/Nominal/Examples/W.thy Wed Jun 06 14:23:13 2018 +0200
@@ -9,7 +9,7 @@
abbreviation
"difference_list" :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" ("_ - _" [60,60] 60)
where
- "xs - ys \<equiv> filter (\<lambda>x. x\<notin>set ys) xs"
+ "xs - ys \<equiv> [x \<leftarrow> xs. x\<notin>set ys]"
lemma difference_eqvt_tvar[eqvt]:
fixes pi::"tvar prm"
--- a/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Quickcheck_Examples.thy Wed Jun 06 14:22:54 2018 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Quickcheck_Examples.thy Wed Jun 06 14:23:13 2018 +0200
@@ -71,7 +71,7 @@
oops
theorem S\<^sub>1_sound:
-"S\<^sub>1p w \<Longrightarrow> length (filter (\<lambda>x. x=a) w) = length (filter (\<lambda>x. x=b) w)"
+"S\<^sub>1p w \<Longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
quickcheck[tester=smart_exhaustive, size=15]
oops
@@ -113,7 +113,7 @@
oops
*)
theorem S\<^sub>2_sound:
-"S\<^sub>2p w \<longrightarrow> length (filter (\<lambda>x. x=a) w) = length (filter (\<lambda>x. x=b) w)"
+"S\<^sub>2p w \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
quickcheck[tester=smart_exhaustive, size=5, iterations=10]
oops
@@ -133,16 +133,16 @@
lemma S\<^sub>3_sound:
-"S\<^sub>3p w \<longrightarrow> length (filter (\<lambda>x. x=a) w) = length (filter (\<lambda>x. x=b) w)"
+"S\<^sub>3p w \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
quickcheck[tester=smart_exhaustive, size=10, iterations=10]
oops
-lemma "\<not> (length w > 2) \<or> \<not> (length (filter (\<lambda>x. x=a) w) = length (filter (\<lambda>x. x=b) w))"
+lemma "\<not> (length w > 2) \<or> \<not> (length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b])"
quickcheck[size=10, tester = smart_exhaustive]
oops
theorem S\<^sub>3_complete:
-"length (filter (\<lambda>x. x=a) w) = length (filter (\<lambda>x. x=b) w) \<longrightarrow> S\<^sub>3p w"
+"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. b = x] \<longrightarrow> S\<^sub>3p w"
(*quickcheck[generator=SML]*)
quickcheck[tester=smart_exhaustive, size=10, iterations=100]
oops
@@ -158,12 +158,12 @@
| "\<lbrakk>v \<in> B\<^sub>4; w \<in> B\<^sub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>4"
theorem S\<^sub>4_sound:
-"S\<^sub>4p w \<longrightarrow> length (filter (\<lambda>x. x=a) w) = length (filter (\<lambda>x. x=b) w)"
+"S\<^sub>4p w \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
quickcheck[tester = smart_exhaustive, size=5, iterations=1]
oops
theorem S\<^sub>4_complete:
-"length (filter (\<lambda>x. x=a) w) = length (filter (\<lambda>x. x=b) w) \<longrightarrow> S\<^sub>4p w"
+"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> S\<^sub>4p w"
quickcheck[tester = smart_exhaustive, size=5, iterations=1]
oops
@@ -301,7 +301,7 @@
subsection "Compressed matrix"
-definition "sparsify xs = filter (\<lambda>i. snd i \<noteq> 0) (zip [0..<length xs] xs)"
+definition "sparsify xs = [i \<leftarrow> zip [0..<length xs] xs. snd i \<noteq> 0]"
(*
lemma sparsify_length: "(i, x) \<in> set (sparsify xs) \<Longrightarrow> i < length xs"
by (auto simp: sparsify_def set_zip)
--- a/src/HOL/Probability/Probability_Mass_Function.thy Wed Jun 06 14:22:54 2018 +0200
+++ b/src/HOL/Probability/Probability_Mass_Function.thy Wed Jun 06 14:23:13 2018 +0200
@@ -2028,7 +2028,7 @@
private lemma pmf_of_list_aux:
assumes "\<And>x. x \<in> set (map snd xs) \<Longrightarrow> x \<ge> 0"
assumes "sum_list (map snd xs) = 1"
- shows "(\<integral>\<^sup>+ x. ennreal (sum_list (map snd (filter (\<lambda>z. fst z = x) xs))) \<partial>count_space UNIV) = 1"
+ shows "(\<integral>\<^sup>+ x. ennreal (sum_list (map snd [z\<leftarrow>xs . fst z = x])) \<partial>count_space UNIV) = 1"
proof -
have "(\<integral>\<^sup>+ x. ennreal (sum_list (map snd (filter (\<lambda>z. fst z = x) xs))) \<partial>count_space UNIV) =
(\<integral>\<^sup>+ x. ennreal (sum_list (map (\<lambda>(x',p). indicator {x'} x * p) xs)) \<partial>count_space UNIV)"
@@ -2083,7 +2083,7 @@
show "x \<in> set (map fst xs)"
proof (rule ccontr)
assume "x \<notin> set (map fst xs)"
- hence "filter (\<lambda>z. fst z = x) xs = []" by (auto simp: filter_empty_conv)
+ hence "[z\<leftarrow>xs . fst z = x] = []" by (auto simp: filter_empty_conv)
with A assms show False by (simp add: pmf_pmf_of_list set_pmf_eq)
qed
qed
@@ -2122,10 +2122,10 @@
have "emeasure (pmf_of_list xs) A = nn_integral (measure_pmf (pmf_of_list xs)) (indicator A)"
by simp
also from assms
- have "\<dots> = (\<Sum>x\<in>set_pmf (pmf_of_list xs) \<inter> A. ennreal (sum_list (map snd (filter (\<lambda>z. fst z = x) xs))))"
+ have "\<dots> = (\<Sum>x\<in>set_pmf (pmf_of_list xs) \<inter> A. ennreal (sum_list (map snd [z\<leftarrow>xs . fst z = x])))"
by (subst nn_integral_measure_pmf_finite) (simp_all add: finite_set_pmf_of_list pmf_pmf_of_list Int_def)
also from assms
- have "\<dots> = ennreal (\<Sum>x\<in>set_pmf (pmf_of_list xs) \<inter> A. sum_list (map snd (filter (\<lambda>z. fst z = x) xs)))"
+ have "\<dots> = ennreal (\<Sum>x\<in>set_pmf (pmf_of_list xs) \<inter> A. sum_list (map snd [z\<leftarrow>xs . fst z = x]))"
by (subst sum_ennreal) (auto simp: pmf_of_list_wf_def intro!: sum_list_nonneg)
also have "\<dots> = ennreal (\<Sum>x\<in>set_pmf (pmf_of_list xs) \<inter> A.
indicator A x * pmf (pmf_of_list xs) x)" (is "_ = ennreal ?S")
@@ -2184,7 +2184,7 @@
{
fix x assume A: "x \<in> fst ` set xs" and B: "x \<notin> set_pmf (pmf_of_list xs)"
then obtain y where y: "(x, y) \<in> set xs" by auto
- from B have "sum_list (map snd (filter (\<lambda>z. fst z = x) xs)) = 0"
+ from B have "sum_list (map snd [z\<leftarrow>xs. fst z = x]) = 0"
by (simp add: pmf_pmf_of_list[OF assms(1)] set_pmf_eq)
moreover from y have "y \<in> snd ` {xa \<in> set xs. fst xa = x}" by force
ultimately have "y = 0" using assms(1)
@@ -2199,11 +2199,11 @@
defines "xs' \<equiv> filter (\<lambda>z. snd z \<noteq> 0) xs"
shows "pmf_of_list_wf xs'" "pmf_of_list xs' = pmf_of_list xs"
proof -
- have "map snd (filter (\<lambda>z. snd z \<noteq> 0) xs) = filter (\<lambda>x. x \<noteq> 0) (map snd xs)"
+ have "map snd [z\<leftarrow>xs . snd z \<noteq> 0] = filter (\<lambda>x. x \<noteq> 0) (map snd xs)"
by (induction xs) simp_all
with assms(1) show wf: "pmf_of_list_wf xs'"
by (auto simp: pmf_of_list_wf_def xs'_def sum_list_filter_nonzero)
- have "sum_list (map snd (filter (\<lambda>z. fst z = i) xs')) = sum_list (map snd (filter (\<lambda>z. fst z = i) xs))" for i
+ have "sum_list (map snd [z\<leftarrow>xs' . fst z = i]) = sum_list (map snd [z\<leftarrow>xs . fst z = i])" for i
unfolding xs'_def by (induction xs) simp_all
with assms(1) wf show "pmf_of_list xs' = pmf_of_list xs"
by (intro pmf_eqI) (simp_all add: pmf_pmf_of_list)
--- a/src/HOL/Quickcheck_Examples/Quickcheck_Examples.thy Wed Jun 06 14:22:54 2018 +0200
+++ b/src/HOL/Quickcheck_Examples/Quickcheck_Examples.thy Wed Jun 06 14:23:13 2018 +0200
@@ -186,7 +186,7 @@
oops
lemma
- "i < length (List.transpose (List.transpose xs)) ==> xs ! i = map (%xs. xs ! i) (filter (\<lambda>ys. i < length ys) xs)"
+ "i < length (List.transpose (List.transpose xs)) ==> xs ! i = map (%xs. xs ! i) [ys<-xs. i < length ys]"
quickcheck[random, iterations = 10000]
quickcheck[exhaustive, expect = counterexample]
oops
@@ -228,13 +228,13 @@
oops
lemma
- "[| sorted (rev (map length xs)); i < length xs |] ==> xs ! i = map (%ys. ys ! i) (filter (\<lambda>ys. i < length ys) (remdups xs))"
+ "[| sorted (rev (map length xs)); i < length xs |] ==> xs ! i = map (%ys. ys ! i) [ys<-remdups xs. i < length ys]"
quickcheck[random]
quickcheck[exhaustive, expect = counterexample]
oops
lemma
- "[| sorted (rev (map length xs)); i < length xs |] ==> xs ! i = map (%ys. ys ! i) (filter (\<lambda>ys. length ys \<in> {..<i}) (List.transpose xs))"
+ "[| sorted (rev (map length xs)); i < length xs |] ==> xs ! i = map (%ys. ys ! i) [ys<-List.transpose xs. length ys \<in> {..<i}]"
quickcheck[random]
quickcheck[exhaustive, expect = counterexample]
oops
--- a/src/HOL/Quotient_Examples/Quotient_FSet.thy Wed Jun 06 14:22:54 2018 +0200
+++ b/src/HOL/Quotient_Examples/Quotient_FSet.thy Wed Jun 06 14:23:13 2018 +0200
@@ -69,12 +69,12 @@
definition
inter_list :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
where
- [simp]: "inter_list xs ys = filter (\<lambda>x. x \<in> set xs \<and> x \<in> set ys) xs"
+ [simp]: "inter_list xs ys = [x \<leftarrow> xs. x \<in> set xs \<and> x \<in> set ys]"
definition
diff_list :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
where
- [simp]: "diff_list xs ys = filter (\<lambda>x. x \<notin> set ys) xs"
+ [simp]: "diff_list xs ys = [x \<leftarrow> xs. x \<notin> set ys]"
definition
rsp_fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> bool"
--- a/src/HOL/Random.thy Wed Jun 06 14:22:54 2018 +0200
+++ b/src/HOL/Random.thy Wed Jun 06 14:23:13 2018 +0200
@@ -116,7 +116,7 @@
lemma select_weight_drop_zero:
"select_weight (filter (\<lambda>(k, _). k > 0) xs) = select_weight xs"
proof -
- have "sum_list (map fst (filter (\<lambda>(k, _). 0 < k) xs)) = sum_list (map fst xs)"
+ have "sum_list (map fst [(k, _)\<leftarrow>xs . 0 < k]) = sum_list (map fst xs)"
by (induct xs) (auto simp add: less_natural_def natural_eq_iff)
then show ?thesis by (simp only: select_weight_def pick_drop_zero)
qed
--- a/src/HOL/ex/Quicksort.thy Wed Jun 06 14:22:54 2018 +0200
+++ b/src/HOL/ex/Quicksort.thy Wed Jun 06 14:23:13 2018 +0200
@@ -13,11 +13,11 @@
fun quicksort :: "'a list \<Rightarrow> 'a list" where
"quicksort [] = []"
-| "quicksort (x#xs) = quicksort (filter (\<lambda>y. \<not> x\<le>y) xs) @ [x] @ quicksort (filter (\<lambda>y. x\<le>y) xs)"
+| "quicksort (x#xs) = quicksort [y\<leftarrow>xs. \<not> x\<le>y] @ [x] @ quicksort [y\<leftarrow>xs. x\<le>y]"
lemma [code]:
"quicksort [] = []"
- "quicksort (x#xs) = quicksort (filter (\<lambda>y. \<not> x\<le>y) xs) @ [x] @ quicksort (filter (\<lambda>y. x\<le>y) xs)"
+ "quicksort (x#xs) = quicksort [y\<leftarrow>xs. y<x] @ [x] @ quicksort [y\<leftarrow>xs. x\<le>y]"
by (simp_all add: not_le)
lemma quicksort_permutes [simp]:
--- a/src/HOL/ex/Radix_Sort.thy Wed Jun 06 14:22:54 2018 +0200
+++ b/src/HOL/ex/Radix_Sort.thy Wed Jun 06 14:23:13 2018 +0200
@@ -44,7 +44,7 @@
lemma sorted_from_Suc2:
"\<lbrakk> cols xss n; i < n;
sorted_col i xss;
- \<And>x. sorted_from (i+1) (filter (\<lambda>ys. ys!i = x) xss) \<rbrakk>
+ \<And>x. sorted_from (i+1) [ys \<leftarrow> xss. ys!i = x] \<rbrakk>
\<Longrightarrow> sorted_from i xss"
proof(induction xss rule: induct_list012)
case 1 show ?case by simp
@@ -68,7 +68,7 @@
proof(rule "3.IH"[OF _ "3.prems"(2)])
show "cols (xs2 # xss) n" using "3.prems"(1) by(simp add: cols_def)
show "sorted_col i (xs2 # xss)" using "3.prems"(3) by simp
- show "\<And>x. sorted_from (i+1) (filter (\<lambda>ys. ys ! i = x) (xs2 # xss))"
+ show "\<And>x. sorted_from (i+1) [ys\<leftarrow>xs2 # xss . ys ! i = x]"
using "3.prems"(4)
sorted_antimono_suffix[OF map_mono_suffix[OF filter_mono_suffix[OF suffix_ConsI[OF suffix_order.order.refl]]]]
by fastforce
@@ -81,7 +81,7 @@
shows "sorted_from i (sort_col i xss)"
proof (rule sorted_from_Suc2[OF cols_sort1[OF assms(1)] assms(2)])
show "sorted_col i (sort_col i xss)" by(simp add: sorted)
- fix x show "sorted_from (i+1) (filter (\<lambda>ys. ys ! i = x) (sort_col i xss))"
+ fix x show "sorted_from (i+1) [ys \<leftarrow> sort_col i xss . ys ! i = x]"
proof -
from assms(3)
have "sorted_from (i+1) (filter (\<lambda>ys. ys!i = x) xss)"