# HG changeset patch # User wenzelm # Date 1608762111 -3600 # Node ID eac16c76273ee27fc764a8d707c1a3198c1fc85e # Parent 4fc3dc37f4069963db1bd7fd142190537f14d6f9# Parent 055f44891643d94c10d6d1f691f7762df7ab0ffb merged diff -r 055f44891643 -r eac16c76273e src/HOL/Data_Structures/Trie_Fun.thy --- a/src/HOL/Data_Structures/Trie_Fun.thy Wed Dec 23 23:20:52 2020 +0100 +++ b/src/HOL/Data_Structures/Trie_Fun.thy Wed Dec 23 23:21:51 2020 +0100 @@ -20,7 +20,7 @@ fun insert :: "'a list \ 'a trie \ 'a trie" where "insert [] (Nd b m) = Nd True m" | "insert (x#xs) (Nd b m) = - Nd b (m(x := Some(insert xs (case m x of None \ empty | Some t \ t))))" + (let s = (case m x of None \ empty | Some t \ t) in Nd b (m(x := Some(insert xs s))))" fun delete :: "'a list \ 'a trie \ 'a trie" where "delete [] (Nd b m) = Nd False m" | @@ -29,62 +29,39 @@ None \ m | Some t \ m(x := Some(delete xs t)))" -text \The actual definition of \set\ is a bit cryptic but canonical, to enable -primrec to prove termination:\ - -primrec set :: "'a trie \ 'a list set" where -"set (Nd b m) = (if b then {[]} else {}) \ - (\a. case (map_option set o m) a of None \ {} | Some t \ (#) a ` t)" +text \Use (a tuned version of) @{const isin} as an abstraction function:\ -text \This is the more human-readable version:\ +lemma isin_case: "isin (Nd b m) xs = + (case xs of + [] \ b | + x # ys \ (case m x of None \ False | Some t \ isin t ys))" +by(cases xs)auto -lemma set_Nd: - "set (Nd b m) = - (if b then {[]} else {}) \ - (\a. case m a of None \ {} | Some t \ (#) a ` set t)" -by (auto simp: split: option.splits) +definition set :: "'a trie \ 'a list set" where +[simp]: "set t = {xs. isin t xs}" lemma isin_set: "isin t xs = (xs \ set t)" -apply(induction t xs rule: isin.induct) -apply (auto split: option.split) -done +by simp lemma set_insert: "set (insert xs t) = set t \ {xs}" -proof(induction xs t rule: insert.induct) - case 1 thus ?case by simp -next - case 2 - thus ?case - apply(simp) - apply(subst set_eq_iff) - apply(auto split!: if_splits option.splits) - apply fastforce - by (metis imageI option.sel) -qed +by (induction xs t rule: insert.induct) + (auto simp: isin_case split!: if_splits option.splits list.splits) lemma set_delete: "set (delete xs t) = set t - {xs}" -proof(induction xs t rule: delete.induct) - case 1 thus ?case by (force split: option.splits) -next - case 2 - show ?case - apply (auto simp add: image_iff 2 split!: if_splits option.splits) - apply (metis DiffI empty_iff insert_iff option.inject) - apply (metis DiffI empty_iff insert_iff option.inject) - done -qed +by (induction xs t rule: delete.induct) + (auto simp: isin_case split!: if_splits option.splits list.splits) interpretation S: Set where empty = empty and isin = isin and insert = insert and delete = delete and set = set and invar = "\_. True" proof (standard, goal_cases) - case 1 show ?case by (simp) + case 1 show ?case by (simp add: isin_case split: list.split) next - case 2 thus ?case by(simp add: isin_set) + case 2 show ?case by(rule isin_set) next - case 3 thus ?case by(simp add: set_insert) + case 3 show ?case by(rule set_insert) next - case 4 thus ?case by(simp add: set_delete) + case 4 show ?case by(rule set_delete) qed (rule TrueI)+ end diff -r 055f44891643 -r eac16c76273e src/HOL/Data_Structures/Trie_Map.thy --- a/src/HOL/Data_Structures/Trie_Map.thy Wed Dec 23 23:20:52 2020 +0100 +++ b/src/HOL/Data_Structures/Trie_Map.thy Wed Dec 23 23:21:51 2020 +0100 @@ -91,13 +91,13 @@ where empty = empty and isin = isin and insert = insert and delete = delete and set = "set o abs" and invar = invar proof (standard, goal_cases) - case 1 show ?case by (simp) + case 1 show ?case by (simp add: isin_case split: list.split) next - case 2 thus ?case by (simp add: isin_set isin_abs) + case 2 thus ?case by (simp add: isin_abs) next - case 3 thus ?case by (simp add: set_insert abs_insert) + case 3 thus ?case by (simp add: set_insert abs_insert del: set_def) next - case 4 thus ?case by (simp add: set_delete abs_delete) + case 4 thus ?case by (simp add: set_delete abs_delete del: set_def) next case 5 thus ?case by (simp add: M.map_specs RBT_Set.empty_def[symmetric]) next diff -r 055f44891643 -r eac16c76273e src/HOL/Series.thy --- a/src/HOL/Series.thy Wed Dec 23 23:20:52 2020 +0100 +++ b/src/HOL/Series.thy Wed Dec 23 23:21:51 2020 +0100 @@ -614,7 +614,7 @@ lemma suminf_geometric: "norm c < 1 \ suminf (\n. c^n) = 1 / (1 - c)" by (rule sums_unique[symmetric]) (rule geometric_sums) -lemma summable_geometric_iff: "summable (\n. c ^ n) \ norm c < 1" +lemma summable_geometric_iff [simp]: "summable (\n. c ^ n) \ norm c < 1" proof assume "summable (\n. c ^ n :: 'a :: real_normed_field)" then have "(\n. norm c ^ n) \ 0" diff -r 055f44891643 -r eac16c76273e src/HOL/Transcendental.thy --- a/src/HOL/Transcendental.thy Wed Dec 23 23:20:52 2020 +0100 +++ b/src/HOL/Transcendental.thy Wed Dec 23 23:21:51 2020 +0100 @@ -1886,10 +1886,8 @@ next fix x :: real assume "x \ {- 1<..<1}" - then have "norm (-x) < 1" by auto - show "summable (\n. (- 1) ^ n * (1 / real (n + 1)) * real (Suc n) * x^n)" - unfolding One_nat_def - by (auto simp: power_mult_distrib[symmetric] summable_geometric[OF \norm (-x) < 1\]) + then show "summable (\n. (- 1) ^ n * (1 / real (n + 1)) * real (Suc n) * x^n)" + by (simp add: abs_if flip: power_mult_distrib) qed then have "DERIV (\x. suminf (?f x)) (x - 1) :> suminf (?f' x)" unfolding One_nat_def by auto