# HG changeset patch # User nipkow # Date 1528276357 -7200 # Node ID 98cf1c823c486adce22d3f835a610e0f2085a638 # Parent 5c579bb9adb11cf7ad098f09bf26cfa71aacbead Keep filter input syntax diff -r 5c579bb9adb1 -r 98cf1c823c48 src/Doc/Tutorial/Inductive/AB.thy --- a/src/Doc/Tutorial/Inductive/AB.thy Sun Jun 03 19:06:56 2018 +0200 +++ b/src/Doc/Tutorial/Inductive/AB.thy Wed Jun 06 11:12:37 2018 +0200 @@ -64,13 +64,15 @@ \ lemma correctness: - "(w \ S \ size(filter (\x. x=a) w) = size(filter (\x. x=b) w)) \ - (w \ A \ size(filter (\x. x=a) w) = size(filter (\x. x=b) w) + 1) \ - (w \ B \ size(filter (\x. x=b) w) = size(filter (\x. x=a) w) + 1)" + "(w \ S \ size[x\w. x=a] = size[x\w. x=b]) \ + (w \ A \ size[x\w. x=a] = size[x\w. x=b] + 1) \ + (w \ B \ size[x\w. x=b] = size[x\w. x=a] + 1)" txt\\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\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: \ @@ -110,8 +112,8 @@ \ lemma step1: "\i < size w. - \(int(size(filter P (take (i+1) w)))-int(size(filter (\x. \P x) (take (i+1) w)))) - - (int(size(filter P (take i w)))-int(size(filter (\x. \P x) (take i w))))\ \ 1" + \(int(size[x\take (i+1) w. P x])-int(size[x\take (i+1) w. \P x])) + - (int(size[x\take i w. P x])-int(size[x\take i w. \P x]))\ \ 1" txt\\noindent The lemma is a bit hard to read because of the coercion function @@ -135,8 +137,8 @@ \ lemma part1: - "size(filter P w) = size(filter (\x. \P x) w)+2 \ - \i\size w. size(filter P (take i w)) = size(filter (\x. \P x) (take i w))+1" + "size[x\w. P x] = size[x\w. \P x]+2 \ + \i\size w. size[x\take i w. P x] = size[x\take i w. \P x]+1" txt\\noindent This is proved by @{text force} with the help of the intermediate value theorem, @@ -155,10 +157,10 @@ lemma part2: - "\size(filter P (take i w @ drop i w)) = - size(filter (\x. \P x) (take i w @ drop i w))+2; - size(filter P (take i w)) = size(filter (\x. \P x) (take i w))+1\ - \ size(filter P (drop i w)) = size(filter (\x. \P x) (drop i w))+1" + "\size[x\take i w @ drop i w. P x] = + size[x\take i w @ drop i w. \P x]+2; + size[x\take i w. P x] = size[x\take i w. \P x]+1\ + \ size[x\drop i w. P x] = size[x\drop i w. \P x]+1" by(simp del: append_take_drop_id) text\\noindent @@ -185,9 +187,9 @@ \ theorem completeness: - "(size(filter (\x. x=a) w) = size(filter (\x. x=b) w) \ w \ S) \ - (size(filter (\x. x=a) w) = size(filter (\x. x=b) w) + 1 \ w \ A) \ - (size(filter (\x. x=b) w) = size(filter (\x. x=a) w) + 1 \ w \ B)" + "(size[x\w. x=a] = size[x\w. x=b] \ w \ S) \ + (size[x\w. x=a] = size[x\w. x=b] + 1 \ w \ A) \ + (size[x\w. x=b] = size[x\w. x=a] + 1 \ w \ B)" txt\\noindent The proof is by induction on @{term w}. Structural induction would fail here @@ -232,9 +234,9 @@ apply(clarify) txt\\noindent This yields an index @{prop"i \ length v"} such that -@{prop[display]"length (filter (\x. x=a) (take i v)) = length (filter (\x. x=b) (take i v)) + 1"} +@{prop[display]"length [x\take i v . x = a] = length [x\take i v . x = b] + 1"} With the help of @{thm[source]part2} it follows that -@{prop[display]"length (filter (\x. x=a) (drop i v)) = length (filter (\x. x=b) (drop i v)) + 1"} +@{prop[display]"length [x\drop i v . x = a] = length [x\drop i v . x = b] + 1"} \ apply(drule part2[of "\x. x=a", simplified]) diff -r 5c579bb9adb1 -r 98cf1c823c48 src/HOL/Library/AList.thy --- a/src/HOL/Library/AList.thy Sun Jun 03 19:06:56 2018 +0200 +++ b/src/HOL/Library/AList.thy Wed Jun 06 11:12:37 2018 +0200 @@ -45,7 +45,7 @@ using assms by (simp add: update_keys) lemma update_filter: - "a \ k \ update k v (filter (\q. fst q \ a) ps) = filter (\q. fst q \ a) (update k v ps)" + "a \ k \ update k v [q\ps. fst q \ a] = [q\update k v ps. fst q \ a]" by (induct ps) auto lemma update_triv: "map_of al k = Some v \ update k v al = al" @@ -361,12 +361,12 @@ proof - have "y \ x \ x \ y" for y by auto - then show "filter (\p. fst p \ D \ x \ fst p) al = filter (\p. fst p \ D \ fst p \ x) al" + then show "[p \ al. fst p \ D \ x \ fst p] = [p \ al. fst p \ D \ fst p \ x]" by simp assume "x \ D" then have "y \ D \ y \ D \ x \ y" for y by auto - then show "filter (\p. fst p \ D \ x \ fst p) al = filter (\p. fst p \ D) al" + then show "[p \ al . fst p \ D \ x \ fst p] = [p \ al . fst p \ D]" by simp qed @@ -492,7 +492,7 @@ lemma distinct_map_ran: "distinct (map fst al) \ 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 (\p. fst p \ a) ps) = filter (\p. fst p \ a) (map_ran f ps)" +lemma map_ran_filter: "map_ran f [p\ps. fst p \ a] = [p\map_ran f ps. fst p \ 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)" diff -r 5c579bb9adb1 -r 98cf1c823c48 src/HOL/Library/Finite_Map.thy --- a/src/HOL/Library/Finite_Map.thy Sun Jun 03 19:06:56 2018 +0200 +++ b/src/HOL/Library/Finite_Map.thy Wed Jun 06 11:12:37 2018 +0200 @@ -84,10 +84,10 @@ definition map_filter :: "('a \ bool) \ ('a \ 'b) \ ('a \ 'b)" where "map_filter P m = (\x. if P x then m x else None)" -lemma map_filter_map_of[simp]: "map_filter P (map_of m) = map_of (filter (\(k, _). P k) m)" +lemma map_filter_map_of[simp]: "map_filter P (map_of m) = map_of [(k, _) \ m. P k]" proof fix x - show "map_filter P (map_of m) x = map_of (filter (\(k, _). P k) m) x" + show "map_filter P (map_of m) x = map_of [(k, _) \ m. P k] x" by (induct m) (auto simp: map_filter_def) qed diff -r 5c579bb9adb1 -r 98cf1c823c48 src/HOL/Library/Multiset.thy --- a/src/HOL/Library/Multiset.thy Sun Jun 03 19:06:56 2018 +0200 +++ b/src/HOL/Library/Multiset.thy Wed Jun 06 11:12:37 2018 +0200 @@ -1888,7 +1888,7 @@ apply simp done -lemma mset_compl_union [simp]: "mset (filter P xs) + mset (filter (\x. \P x) xs) = mset xs" +lemma mset_compl_union [simp]: "mset [x\xs. P x] + mset [x\xs. \P x] = mset xs" by (induct xs) auto lemma nth_mem_mset: "i < length ls \ (ls ! i) \# mset ls" @@ -2582,14 +2582,14 @@ show "mset ys = mset xs" by simp from \sorted (map f ys)\ show "sorted (map f ys)" . - show "filter (\x. f k = f x) ys = filter (\x. f k = f x) xs" if "k \ set ys" for k + show "[x\ys . f k = f x] = [x\xs . f k = f x]" if "k \ 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 \inj_on f (set xs)\ have inj: "inj_on f (insert k (set ys))" by (simp add: set_equal) - from inj have "filter (\x. f k = f x) ys = filter (HOL.eq k) ys" + from inj have "[x\ys . f k = f x] = filter (HOL.eq k) ys" by (auto intro!: inj_on_filter_key_eq) also have "\ = 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 "\ = filter (HOL.eq k) xs" by (simp add: replicate_count_mset_eq_filter_eq) - also have "\ = filter (\x. f k = f x) xs" + also have "\ = [x\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 (\x. f x < f (xs ! (length xs div 2))) xs) - @ filter (\x. f x = f (xs ! (length xs div 2))) xs - @ sort_key f (filter (\x. f x > f (xs ! (length xs div 2))) xs)" (is "sort_key f ?lhs = ?rhs") + "sort_key f xs = sort_key f [x\xs. f x < f (xs ! (length xs div 2))] + @ [x\xs. f x = f (xs ! (length xs div 2))] + @ sort_key f [x\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 \ set ?rhs" let ?pivot = "f (xs ! (length xs div 2))" have *: "\x. f l = f x \ f x = f l" by auto - have "filter (\x. f x = f l) (sort_key f xs) = filter (\x. f x = f l) xs" + have "[x \ sort_key f xs . f x = f l] = [x \ xs. f x = f l]" unfolding filter_sort by (rule properties_for_sort_key) (auto intro: sorted_map_same) - with * have **: "filter (\x. f l = f x) (sort_key f xs) = filter (\x. f l = f x) xs" by simp + with * have **: "[x \ sort_key f xs . f l = f x] = [x \ xs. f l = f x]" by simp have "\x P. P (f x) ?pivot \ f l = f x \ P (f l) ?pivot \ f l = f x" by auto - then have "\P. filter (\x. P (f x) ?pivot \ f l = f x) (sort_key f xs) = - filter (\x. P (f l) ?pivot \ f l = f x) (sort_key f xs)" by simp + then have "\P. [x \ sort_key f xs . P (f x) ?pivot \ f l = f x] = + [x \ sort_key f xs. P (f l) ?pivot \ f l = f x]" by simp note *** = this [of "(<)"] this [of "(>)"] this [of "(=)"] - show "filter (\x. f l = f x) ?rhs = filter (\x. f l = f x) ?lhs" + show "[x \ ?rhs. f l = f x] = [x \ ?lhs. f l = f x]" proof (cases "f l" ?pivot rule: linorder_cases) case less then have "f l \ ?pivot" and "\ f l > ?pivot" by auto @@ -2648,15 +2648,15 @@ qed lemma sort_by_quicksort: - "sort xs = sort (filter (\x. x < xs ! (length xs div 2)) xs) - @ filter (\x. x = xs ! (length xs div 2)) xs - @ sort (filter (\x. x > xs ! (length xs div 2)) xs)" (is "sort ?lhs = ?rhs") + "sort xs = sort [x\xs. x < xs ! (length xs div 2)] + @ [x\xs. x = xs ! (length xs div 2)] + @ sort [x\xs. x > xs ! (length xs div 2)]" (is "sort ?lhs = ?rhs") using sort_key_by_quicksort [of "\x. x", symmetric] by simp text \A stable parametrized quicksort\ definition part :: "('b \ 'a) \ 'a \ 'b list \ 'b list \ 'b list \ 'b list" where - "part f pivot xs = (filter (\x. f x < pivot) xs, filter (\x. f x = pivot) xs, filter (\x. f x > pivot) xs)" + "part f pivot xs = ([x \ xs. f x < pivot], [x \ xs. f x = pivot], [x \ xs. pivot < f x])" lemma part_code [code]: "part f pivot [] = ([], [], [])" diff -r 5c579bb9adb1 -r 98cf1c823c48 src/HOL/MicroJava/DFA/Kildall.thy --- a/src/HOL/MicroJava/DFA/Kildall.thy Sun Jun 03 19:06:56 2018 +0200 +++ b/src/HOL/MicroJava/DFA/Kildall.thy Wed Jun 06 11:12:37 2018 +0200 @@ -38,7 +38,7 @@ lemma (in Semilat) nth_merges: "\ss. \p < length ss; ss \ list n A; \(p,t)\set ps. p t\A \ \ - (merges f ps ss)!p = map snd (filter (\(p',t'). p'=p) ps) ++_f ss!p" + (merges f ps ss)!p = map snd [(p',t') \ ps. p'=p] ++_f ss!p" (is "\ss. \_; _; ?steptype ps\ \ ?P ss ps") proof (induct ps) show "\ss. ?P ss []" by simp diff -r 5c579bb9adb1 -r 98cf1c823c48 src/HOL/MicroJava/DFA/LBVComplete.thy --- a/src/HOL/MicroJava/DFA/LBVComplete.thy Sun Jun 03 19:06:56 2018 +0200 +++ b/src/HOL/MicroJava/DFA/LBVComplete.thy Wed Jun 06 11:12:37 2018 +0200 @@ -80,13 +80,13 @@ assume merge: "?s1 \ T" from x ss1 have "?s1 = (if \(pc', s')\set ss1. pc' \ pc + 1 \ s' <=_r c!pc' - then (map snd (filter (\(p', t'). p'=pc+1) ss1)) ++_f x + then (map snd [(p', t') \ ss1 . p'=pc+1]) ++_f x else \)" by (rule merge_def) with merge obtain app: "\(pc',s')\set ss1. pc' \ pc+1 \ s' <=_r c!pc'" (is "?app ss1") and - sum: "(map snd (filter (\(p',t'). p' = pc+1) ss1) ++_f x) = ?s1" + sum: "(map snd [(p',t') \ 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 \(pc', s')\set ss2. pc' \ pc + 1 \ s' <=_r c!pc' - then map snd (filter (\(p', t'). p' = pc + 1) ss2) ++_f x + then map snd [(p', t') \ ss2 . p' = pc + 1] ++_f x else \)" by (rule merge_def) ultimately have ?thesis by simp @@ -194,7 +194,7 @@ ultimately have "merge c pc ?step (c!Suc pc) = (if \(pc',s')\set ?step. pc'\pc+1 \ s' <=_r c!pc' - then map snd (filter (\(p',t').p'=pc+1) ?step) ++_f c!Suc pc + then map snd [(p',t') \ ?step.p'=pc+1] ++_f c!Suc pc else \)" unfolding mrg_def by (rule lbv.merge_def [OF lbvc.axioms(1), OF lbvc_axioms]) moreover { fix pc' s' assume s': "(pc', s') \ set ?step" and suc_pc: "pc' \ pc+1" @@ -207,7 +207,7 @@ } hence "\(pc',s')\set ?step. pc'\pc+1 \ s' <=_r c!pc'" by auto moreover from pc have "Suc pc = length \ \ Suc pc < length \" by auto - hence "map snd (filter (\(p',t').p'=pc+1) ?step) ++_f c!Suc pc \ \" + hence "map snd [(p',t') \ ?step.p'=pc+1] ++_f c!Suc pc \ \" (is "?map ++_f _ \ _") proof (rule disjE) assume pc': "Suc pc = length \" @@ -215,7 +215,7 @@ moreover from pc' bounded pc have "\(p',t')\set ?step. p'\pc+1" by clarify (drule boundedD, auto) - hence "filter (\(p',t').p'=pc+1) ?step = []" by (blast intro: filter_False) + hence "[(p',t') \ ?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) \ \" by (simp add: wti) ultimately have "merge c pc ?step (c!Suc pc) = - map snd (filter (\(p',t').p'=pc+1) ?step) ++_f c!Suc pc" by (rule merge_not_top_s) + map snd [(p',t')\ ?step.p'=pc+1] ++_f c!Suc pc" by (rule merge_not_top_s) hence "?wti = \" (is "_ = (?map ++_f _)" is "_ = ?sum") by (simp add: wti) also { from suc_pc phi have "\!Suc pc \ A" by simp diff -r 5c579bb9adb1 -r 98cf1c823c48 src/HOL/MicroJava/DFA/LBVCorrect.thy --- a/src/HOL/MicroJava/DFA/LBVCorrect.thy Sun Jun 03 19:06:56 2018 +0200 +++ b/src/HOL/MicroJava/DFA/LBVCorrect.thy Wed Jun 06 11:12:37 2018 +0200 @@ -88,7 +88,7 @@ also from s2 merge have "\ \ \" (is "?merge \ _") by simp with cert_in_A step_in_A - have "?merge = (map snd (filter (\(p',t').p'=pc+1) (?step pc)) ++_f (c!(pc+1)))" + have "?merge = (map snd [(p',t') \ ?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 diff -r 5c579bb9adb1 -r 98cf1c823c48 src/HOL/MicroJava/DFA/LBVSpec.thy --- a/src/HOL/MicroJava/DFA/LBVSpec.thy Sun Jun 03 19:06:56 2018 +0200 +++ b/src/HOL/MicroJava/DFA/LBVSpec.thy Wed Jun 06 11:12:37 2018 +0200 @@ -113,7 +113,7 @@ lemma (in Semilat) pp_ub1': assumes S: "snd`set S \ A" assumes y: "y \ A" and ab: "(a, b) \ set S" - shows "b <=_r map snd (filter (\(p', t'). p' = a) S) ++_f y" + shows "b <=_r map snd [(p', t') \ S . p' = a] ++_f y" proof - from S have "\(x,y) \ set S. y \ A" by auto with semilat y ab show ?thesis by - (rule ub1') @@ -172,7 +172,7 @@ "\x. x \ A \ snd`set ss \ A \ merge c pc ss x = (if \(pc',s') \ set ss. pc'\pc+1 \ s' <=_r c!pc' then - map snd (filter (\(p',t'). p'=pc+1) ss) ++_f x + map snd [(p',t') \ ss. p'=pc+1] ++_f x else \)" (is "\x. _ \ _ \ ?merge ss x = ?if ss x" is "\x. _ \ _ \ ?P ss x") proof (induct ss) @@ -202,15 +202,15 @@ hence "\(pc', s')\set ls. pc'\pc+1 \ s' <=_r c!pc'" by auto moreover from True have - "map snd (filter (\(p',t'). p'=pc+1) ls) ++_f ?if' = - (map snd (filter (\(p',t'). p'=pc+1) (l#ls)) ++_f x)" + "map snd [(p',t')\ls . p'=pc+1] ++_f ?if' = + (map snd [(p',t')\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 (\(p', t'). p' = Suc pc) ls)) \ A" by auto + from ls have "set (map snd [(p', t')\ls . p' = Suc pc]) \ 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 \ A" and ss: "snd`set ss \ A" assumes m: "merge c pc ss x \ \" - shows "merge c pc ss x = (map snd (filter (\(p',t'). p'=pc+1) ss) ++_f x)" + shows "merge c pc ss x = (map snd [(p',t') \ ss. p'=pc+1] ++_f x)" proof - from ss m have "\(pc',s') \ set ss. (pc' \ pc+1 \ s' <=_r c!pc')" by (rule merge_not_top) @@ -307,8 +307,8 @@ assumes s0: "snd`set ss \ A" and x: "x \ A" shows "merge c pc ss x \ A" proof - - from s0 have "set (map snd (filter (\(p', t'). p'=pc+1) ss)) \ A" by auto - with x have "(map snd (filter (\(p', t'). p'=pc+1) ss) ++_f x) \ A" + from s0 have "set (map snd [(p', t')\ss . p'=pc+1]) \ A" by auto + with x have "(map snd [(p', t')\ss . p'=pc+1] ++_f x) \ A" by (auto intro!: plusplus_closed semilat) with s0 x show ?thesis by (simp add: merge_def T_A) qed diff -r 5c579bb9adb1 -r 98cf1c823c48 src/HOL/MicroJava/DFA/SemilatAlg.thy --- a/src/HOL/MicroJava/DFA/SemilatAlg.thy Sun Jun 03 19:06:56 2018 +0200 +++ b/src/HOL/MicroJava/DFA/SemilatAlg.thy Wed Jun 06 11:12:37 2018 +0200 @@ -155,7 +155,7 @@ lemma ub1': assumes "semilat (A, r, f)" shows "\\(p,s) \ set S. s \ A; y \ A; (a,b) \ set S\ - \ b <=_r map snd (filter (\(p', t'). p' = a) S) ++_f y" + \ b <=_r map snd [(p', t')\S. p' = a] ++_f y" proof - interpret Semilat A r f using assms by (rule Semilat.intro) @@ -175,7 +175,7 @@ lemma plusplus_empty: "\s'. (q, s') \ set S \ s' +_f ss ! q = ss ! q \ - (map snd (filter (\(p', t'). p' = q) S) ++_f ss ! q) = ss ! q" + (map snd [(p', t') \ S. p' = q] ++_f ss ! q) = ss ! q" by (induct S) auto end diff -r 5c579bb9adb1 -r 98cf1c823c48 src/HOL/Nitpick_Examples/Manual_Nits.thy --- a/src/HOL/Nitpick_Examples/Manual_Nits.thy Sun Jun 03 19:06:56 2018 +0200 +++ b/src/HOL/Nitpick_Examples/Manual_Nits.thy Wed Jun 06 11:12:37 2018 +0200 @@ -317,7 +317,7 @@ | "\v \ B\<^sub>1; v \ B\<^sub>1\ \ a # v @ w \ B\<^sub>1" theorem S\<^sub>1_sound: -"w \ S\<^sub>1 \ length (filter (\x. x = a) w) = length (filter (\x. x = b) w)" +"w \ S\<^sub>1 \ length [x \ w. x = a] = length [x \ w. x = b]" nitpick [expect = genuine] oops @@ -330,7 +330,7 @@ | "\v \ B\<^sub>2; v \ B\<^sub>2\ \ a # v @ w \ B\<^sub>2" theorem S\<^sub>2_sound: -"w \ S\<^sub>2 \ length (filter (\x. x = a) w) = length (filter (\x. x = b) w)" +"w \ S\<^sub>2 \ length [x \ w. x = a] = length [x \ w. x = b]" nitpick [expect = genuine] oops @@ -343,12 +343,12 @@ | "\v \ B\<^sub>3; w \ B\<^sub>3\ \ a # v @ w \ B\<^sub>3" theorem S\<^sub>3_sound: -"w \ S\<^sub>3 \ length (filter (\x. x = a) w) = length (filter (\x. x = b) w)" +"w \ S\<^sub>3 \ length [x \ w. x = a] = length [x \ w. x = b]" nitpick [card = 1-5, expect = none] sorry theorem S\<^sub>3_complete: -"length (filter (\x. x = a) w) = length (filter (\x. x = b) w) \ w \ S\<^sub>3" +"length [x \ w. x = a] = length [x \ w. x = b] \ w \ S\<^sub>3" nitpick [expect = genuine] oops @@ -362,19 +362,19 @@ | "\v \ B\<^sub>4; w \ B\<^sub>4\ \ a # v @ w \ B\<^sub>4" theorem S\<^sub>4_sound: -"w \ S\<^sub>4 \ length (filter (\x. x = a) w) = length (filter (\x. x = b) w)" +"w \ S\<^sub>4 \ length [x \ w. x = a] = length [x \ w. x = b]" nitpick [card = 1-5, expect = none] sorry theorem S\<^sub>4_complete: -"length (filter (\x. x = a) w) = length (filter (\x. x = b) w) \ w \ S\<^sub>4" +"length [x \ w. x = a] = length [x \ w. x = b] \ w \ S\<^sub>4" nitpick [card = 1-5, expect = none] sorry theorem S\<^sub>4_A\<^sub>4_B\<^sub>4_sound_and_complete: -"w \ S\<^sub>4 \ length (filter (\x. x = a) w) = length (filter (\x. x = b) w)" -"w \ A\<^sub>4 \ length (filter (\x. x = a) w) = length (filter (\x. x = b) w) + 1" -"w \ B\<^sub>4 \ length (filter (\x. x = b) w) = length (filter (\x. x = a) w) + 1" +"w \ S\<^sub>4 \ length [x \ w. x = a] = length [x \ w. x = b]" +"w \ A\<^sub>4 \ length [x \ w. x = a] = length [x \ w. x = b] + 1" +"w \ B\<^sub>4 \ length [x \ w. x = b] = length [x \ w. x = a] + 1" nitpick [card = 1-5, expect = none] sorry diff -r 5c579bb9adb1 -r 98cf1c823c48 src/HOL/Nominal/Examples/W.thy --- a/src/HOL/Nominal/Examples/W.thy Sun Jun 03 19:06:56 2018 +0200 +++ b/src/HOL/Nominal/Examples/W.thy Wed Jun 06 11:12:37 2018 +0200 @@ -9,7 +9,7 @@ abbreviation "difference_list" :: "'a list \ 'a list \ 'a list" ("_ - _" [60,60] 60) where - "xs - ys \ filter (\x. x\set ys) xs" + "xs - ys \ [x \ xs. x\set ys]" lemma difference_eqvt_tvar[eqvt]: fixes pi::"tvar prm" diff -r 5c579bb9adb1 -r 98cf1c823c48 src/HOL/Predicate_Compile_Examples/Predicate_Compile_Quickcheck_Examples.thy --- a/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Quickcheck_Examples.thy Sun Jun 03 19:06:56 2018 +0200 +++ b/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Quickcheck_Examples.thy Wed Jun 06 11:12:37 2018 +0200 @@ -71,7 +71,7 @@ oops theorem S\<^sub>1_sound: -"S\<^sub>1p w \ length (filter (\x. x=a) w) = length (filter (\x. x=b) w)" +"S\<^sub>1p w \ length [x \ w. x = a] = length [x \ w. x = b]" quickcheck[tester=smart_exhaustive, size=15] oops @@ -113,7 +113,7 @@ oops *) theorem S\<^sub>2_sound: -"S\<^sub>2p w \ length (filter (\x. x=a) w) = length (filter (\x. x=b) w)" +"S\<^sub>2p w \ length [x \ w. x = a] = length [x \ 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 \ length (filter (\x. x=a) w) = length (filter (\x. x=b) w)" +"S\<^sub>3p w \ length [x \ w. x = a] = length [x \ w. x = b]" quickcheck[tester=smart_exhaustive, size=10, iterations=10] oops -lemma "\ (length w > 2) \ \ (length (filter (\x. x=a) w) = length (filter (\x. x=b) w))" +lemma "\ (length w > 2) \ \ (length [x \ w. x = a] = length [x \ w. x = b])" quickcheck[size=10, tester = smart_exhaustive] oops theorem S\<^sub>3_complete: -"length (filter (\x. x=a) w) = length (filter (\x. x=b) w) \ S\<^sub>3p w" +"length [x \ w. x = a] = length [x \ w. b = x] \ S\<^sub>3p w" (*quickcheck[generator=SML]*) quickcheck[tester=smart_exhaustive, size=10, iterations=100] oops @@ -158,12 +158,12 @@ | "\v \ B\<^sub>4; w \ B\<^sub>4\ \ a # v @ w \ B\<^sub>4" theorem S\<^sub>4_sound: -"S\<^sub>4p w \ length (filter (\x. x=a) w) = length (filter (\x. x=b) w)" +"S\<^sub>4p w \ length [x \ w. x = a] = length [x \ w. x = b]" quickcheck[tester = smart_exhaustive, size=5, iterations=1] oops theorem S\<^sub>4_complete: -"length (filter (\x. x=a) w) = length (filter (\x. x=b) w) \ S\<^sub>4p w" +"length [x \ w. x = a] = length [x \ w. x = b] \ S\<^sub>4p w" quickcheck[tester = smart_exhaustive, size=5, iterations=1] oops @@ -301,7 +301,7 @@ subsection "Compressed matrix" -definition "sparsify xs = filter (\i. snd i \ 0) (zip [0.. zip [0.. 0]" (* lemma sparsify_length: "(i, x) \ set (sparsify xs) \ i < length xs" by (auto simp: sparsify_def set_zip) diff -r 5c579bb9adb1 -r 98cf1c823c48 src/HOL/Probability/Probability_Mass_Function.thy --- a/src/HOL/Probability/Probability_Mass_Function.thy Sun Jun 03 19:06:56 2018 +0200 +++ b/src/HOL/Probability/Probability_Mass_Function.thy Wed Jun 06 11:12:37 2018 +0200 @@ -2028,7 +2028,7 @@ private lemma pmf_of_list_aux: assumes "\x. x \ set (map snd xs) \ x \ 0" assumes "sum_list (map snd xs) = 1" - shows "(\\<^sup>+ x. ennreal (sum_list (map snd (filter (\z. fst z = x) xs))) \count_space UNIV) = 1" + shows "(\\<^sup>+ x. ennreal (sum_list (map snd [z\xs . fst z = x])) \count_space UNIV) = 1" proof - have "(\\<^sup>+ x. ennreal (sum_list (map snd (filter (\z. fst z = x) xs))) \count_space UNIV) = (\\<^sup>+ x. ennreal (sum_list (map (\(x',p). indicator {x'} x * p) xs)) \count_space UNIV)" @@ -2083,7 +2083,7 @@ show "x \ set (map fst xs)" proof (rule ccontr) assume "x \ set (map fst xs)" - hence "filter (\z. fst z = x) xs = []" by (auto simp: filter_empty_conv) + hence "[z\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 "\ = (\x\set_pmf (pmf_of_list xs) \ A. ennreal (sum_list (map snd (filter (\z. fst z = x) xs))))" + have "\ = (\x\set_pmf (pmf_of_list xs) \ A. ennreal (sum_list (map snd [z\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 "\ = ennreal (\x\set_pmf (pmf_of_list xs) \ A. sum_list (map snd (filter (\z. fst z = x) xs)))" + have "\ = ennreal (\x\set_pmf (pmf_of_list xs) \ A. sum_list (map snd [z\xs . fst z = x]))" by (subst sum_ennreal) (auto simp: pmf_of_list_wf_def intro!: sum_list_nonneg) also have "\ = ennreal (\x\set_pmf (pmf_of_list xs) \ A. indicator A x * pmf (pmf_of_list xs) x)" (is "_ = ennreal ?S") @@ -2184,7 +2184,7 @@ { fix x assume A: "x \ fst ` set xs" and B: "x \ set_pmf (pmf_of_list xs)" then obtain y where y: "(x, y) \ set xs" by auto - from B have "sum_list (map snd (filter (\z. fst z = x) xs)) = 0" + from B have "sum_list (map snd [z\xs. fst z = x]) = 0" by (simp add: pmf_pmf_of_list[OF assms(1)] set_pmf_eq) moreover from y have "y \ snd ` {xa \ set xs. fst xa = x}" by force ultimately have "y = 0" using assms(1) @@ -2199,11 +2199,11 @@ defines "xs' \ filter (\z. snd z \ 0) xs" shows "pmf_of_list_wf xs'" "pmf_of_list xs' = pmf_of_list xs" proof - - have "map snd (filter (\z. snd z \ 0) xs) = filter (\x. x \ 0) (map snd xs)" + have "map snd [z\xs . snd z \ 0] = filter (\x. x \ 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 (\z. fst z = i) xs')) = sum_list (map snd (filter (\z. fst z = i) xs))" for i + have "sum_list (map snd [z\xs' . fst z = i]) = sum_list (map snd [z\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) diff -r 5c579bb9adb1 -r 98cf1c823c48 src/HOL/Quickcheck_Examples/Quickcheck_Examples.thy --- a/src/HOL/Quickcheck_Examples/Quickcheck_Examples.thy Sun Jun 03 19:06:56 2018 +0200 +++ b/src/HOL/Quickcheck_Examples/Quickcheck_Examples.thy Wed Jun 06 11:12:37 2018 +0200 @@ -186,7 +186,7 @@ oops lemma - "i < length (List.transpose (List.transpose xs)) ==> xs ! i = map (%xs. xs ! i) (filter (\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 (\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 (\ys. length ys \ {.. xs ! i = map (%ys. ys ! i) [ys<-List.transpose xs. length ys \ {.. 'a list \ 'a list" where - [simp]: "inter_list xs ys = filter (\x. x \ set xs \ x \ set ys) xs" + [simp]: "inter_list xs ys = [x \ xs. x \ set xs \ x \ set ys]" definition diff_list :: "'a list \ 'a list \ 'a list" where - [simp]: "diff_list xs ys = filter (\x. x \ set ys) xs" + [simp]: "diff_list xs ys = [x \ xs. x \ set ys]" definition rsp_fold :: "('a \ 'b \ 'b) \ bool" diff -r 5c579bb9adb1 -r 98cf1c823c48 src/HOL/Random.thy --- a/src/HOL/Random.thy Sun Jun 03 19:06:56 2018 +0200 +++ b/src/HOL/Random.thy Wed Jun 06 11:12:37 2018 +0200 @@ -116,7 +116,7 @@ lemma select_weight_drop_zero: "select_weight (filter (\(k, _). k > 0) xs) = select_weight xs" proof - - have "sum_list (map fst (filter (\(k, _). 0 < k) xs)) = sum_list (map fst xs)" + have "sum_list (map fst [(k, _)\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 diff -r 5c579bb9adb1 -r 98cf1c823c48 src/HOL/ex/Quicksort.thy --- a/src/HOL/ex/Quicksort.thy Sun Jun 03 19:06:56 2018 +0200 +++ b/src/HOL/ex/Quicksort.thy Wed Jun 06 11:12:37 2018 +0200 @@ -13,11 +13,11 @@ fun quicksort :: "'a list \ 'a list" where "quicksort [] = []" -| "quicksort (x#xs) = quicksort (filter (\y. \ x\y) xs) @ [x] @ quicksort (filter (\y. x\y) xs)" +| "quicksort (x#xs) = quicksort [y\xs. \ x\y] @ [x] @ quicksort [y\xs. x\y]" lemma [code]: "quicksort [] = []" - "quicksort (x#xs) = quicksort (filter (\y. \ x\y) xs) @ [x] @ quicksort (filter (\y. x\y) xs)" + "quicksort (x#xs) = quicksort [y\xs. yxs. x\y]" by (simp_all add: not_le) lemma quicksort_permutes [simp]: diff -r 5c579bb9adb1 -r 98cf1c823c48 src/HOL/ex/Radix_Sort.thy --- a/src/HOL/ex/Radix_Sort.thy Sun Jun 03 19:06:56 2018 +0200 +++ b/src/HOL/ex/Radix_Sort.thy Wed Jun 06 11:12:37 2018 +0200 @@ -44,7 +44,7 @@ lemma sorted_from_Suc2: "\ cols xss n; i < n; sorted_col i xss; - \x. sorted_from (i+1) (filter (\ys. ys!i = x) xss) \ + \x. sorted_from (i+1) [ys \ xss. ys!i = x] \ \ 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 "\x. sorted_from (i+1) (filter (\ys. ys ! i = x) (xs2 # xss))" + show "\x. sorted_from (i+1) [ys\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 (\ys. ys ! i = x) (sort_col i xss))" + fix x show "sorted_from (i+1) [ys \ sort_col i xss . ys ! i = x]" proof - from assms(3) have "sorted_from (i+1) (filter (\ys. ys!i = x) xss)"