# HG changeset patch # User nipkow # Date 1521801422 -3600 # Node ID 30486b96274d51ee823201dc383d20b1488ab69c # Parent 0b70405b39698f7ff3179869c7da76a8495925b3 eliminated "elems" diff -r 0b70405b3969 -r 30486b96274d src/HOL/Data_Structures/Brother12_Set.thy --- a/src/HOL/Data_Structures/Brother12_Set.thy Thu Mar 22 17:18:33 2018 +0100 +++ b/src/HOL/Data_Structures/Brother12_Set.thy Fri Mar 23 11:37:02 2018 +0100 @@ -150,13 +150,9 @@ subsubsection "Proofs for isin" -lemma - "t \ T h \ sorted(inorder t) \ isin t x = (x \ elems(inorder t))" -by(induction h arbitrary: t) (fastforce simp: elems_simps1 split: if_splits)+ - -lemma isin_set: "t \ T h \ - sorted(inorder t) \ isin t x = (x \ elems(inorder t))" -by(induction h arbitrary: t) (auto simp: elems_simps2 split: if_splits) +lemma isin_set: + "t \ T h \ sorted(inorder t) \ isin t x = (x \ set(inorder t))" +by(induction h arbitrary: t) (fastforce simp: isin_simps split: if_splits)+ subsubsection "Proofs for insertion" diff -r 0b70405b3969 -r 30486b96274d src/HOL/Data_Structures/Isin2.thy --- a/src/HOL/Data_Structures/Isin2.thy Thu Mar 22 17:18:33 2018 +0100 +++ b/src/HOL/Data_Structures/Isin2.thy Fri Mar 23 11:37:02 2018 +0100 @@ -17,10 +17,7 @@ EQ \ True | GT \ isin r x)" -lemma "sorted(inorder t) \ isin t x = (x \ elems(inorder t))" -by (induction t) (auto simp: elems_simps1) - -lemma isin_set: "sorted(inorder t) \ isin t x = (x \ elems(inorder t))" -by (induction t) (auto simp: elems_simps2) +lemma isin_set: "sorted(inorder t) \ isin t x = (x \ set(inorder t))" +by (induction t) (auto simp: isin_simps) end diff -r 0b70405b3969 -r 30486b96274d src/HOL/Data_Structures/List_Ins_Del.thy --- a/src/HOL/Data_Structures/List_Ins_Del.thy Thu Mar 22 17:18:33 2018 +0100 +++ b/src/HOL/Data_Structures/List_Ins_Del.thy Fri Mar 23 11:37:02 2018 +0100 @@ -8,38 +8,29 @@ subsection \Elements in a list\ -fun elems :: "'a list \ 'a set" where -"elems [] = {}" | -"elems (x#xs) = Set.insert x (elems xs)" - -lemma elems_app: "elems (xs @ ys) = (elems xs \ elems ys)" -by (induction xs) auto - -lemma elems_eq_set: "elems xs = set xs" -by (induction xs) auto - lemma sorted_Cons_iff: - "sorted(x # xs) = ((\y \ elems xs. x < y) \ sorted xs)" -by(simp add: elems_eq_set sorted_wrt_Cons) + "sorted(x # xs) = ((\y \ set xs. x < y) \ sorted xs)" +by(simp add: sorted_wrt_Cons) lemma sorted_snoc_iff: - "sorted(xs @ [x]) = (sorted xs \ (\y \ elems xs. y < x))" -by(simp add: elems_eq_set sorted_wrt_append) - + "sorted(xs @ [x]) = (sorted xs \ (\y \ set xs. y < x))" +by(simp add: sorted_wrt_append) +(* text\The above two rules introduce quantifiers. It turns out that in practice this is not a problem because of the simplicity of -the "isin" functions that implement @{const elems}. Nevertheless +the "isin" functions that implement @{const set}. Nevertheless it is possible to avoid the quantifiers with the help of some rewrite rules:\ -lemma sorted_ConsD: "sorted (y # xs) \ x \ y \ x \ elems xs" +lemma sorted_ConsD: "sorted (y # xs) \ x \ y \ x \ set xs" by (auto simp: sorted_Cons_iff) -lemma sorted_snocD: "sorted (xs @ [y]) \ y \ x \ x \ elems xs" +lemma sorted_snocD: "sorted (xs @ [y]) \ y \ x \ x \ set xs" by (auto simp: sorted_snoc_iff) -lemmas elems_simps = sorted_lems elems_app -lemmas elems_simps1 = elems_simps sorted_Cons_iff sorted_snoc_iff -lemmas elems_simps2 = elems_simps sorted_ConsD sorted_snocD +lemmas isin_simps2 = sorted_lems sorted_ConsD sorted_snocD +*) + +lemmas isin_simps = sorted_lems sorted_Cons_iff sorted_snoc_iff subsection \Inserting into an ordered list without duplicates:\ @@ -49,7 +40,7 @@ "ins_list x (a#xs) = (if x < a then x#a#xs else if x=a then a#xs else a # ins_list x xs)" -lemma set_ins_list: "elems (ins_list x xs) = insert x (elems xs)" +lemma set_ins_list: "set (ins_list x xs) = insert x (set xs)" by(induction xs) auto lemma distinct_if_sorted: "sorted xs \ distinct xs" @@ -93,16 +84,12 @@ "del_list x [] = []" | "del_list x (a#xs) = (if x=a then xs else a # del_list x xs)" -lemma del_list_idem: "x \ elems xs \ del_list x xs = xs" +lemma del_list_idem: "x \ set xs \ del_list x xs = xs" by (induct xs) simp_all -lemma elems_del_list_eq: - "distinct xs \ elems (del_list x xs) = elems xs - {x}" -apply(induct xs) - apply simp -apply (simp add: elems_eq_set) -apply blast -done +lemma set_del_list_eq: + "distinct xs \ set (del_list x xs) = set xs - {x}" +by(induct xs) auto lemma sorted_del_list: "sorted xs \ sorted(del_list x xs)" apply(induction xs rule: induct_list012) @@ -112,7 +99,7 @@ lemma del_list_sorted: "sorted (xs @ a # ys) \ del_list x (xs @ a # ys) = (if x < a then del_list x xs @ a # ys else xs @ del_list x (a # ys))" by(induction xs) - (fastforce simp: sorted_lems sorted_Cons_iff elems_eq_set intro!: del_list_idem)+ + (fastforce simp: sorted_lems sorted_Cons_iff intro!: del_list_idem)+ text\In principle, @{thm del_list_sorted} suffices, but the following corollaries speed up proofs.\ diff -r 0b70405b3969 -r 30486b96274d src/HOL/Data_Structures/Set_by_Ordered.thy --- a/src/HOL/Data_Structures/Set_by_Ordered.thy Thu Mar 22 17:18:33 2018 +0100 +++ b/src/HOL/Data_Structures/Set_by_Ordered.thy Fri Mar 23 11:37:02 2018 +0100 @@ -30,7 +30,7 @@ fixes inv :: "'t \ bool" assumes empty: "inorder empty = []" assumes isin: "inv t \ sorted(inorder t) \ - isin t x = (x \ elems (inorder t))" + isin t x = (x \ set (inorder t))" assumes insert: "inv t \ sorted(inorder t) \ inorder(insert x t) = ins_list x (inorder t)" assumes delete: "inv t \ sorted(inorder t) \ @@ -41,7 +41,7 @@ begin sublocale Set - empty insert delete isin "elems o inorder" "\t. inv t \ sorted(inorder t)" + empty insert delete isin "set o inorder" "\t. inv t \ sorted(inorder t)" proof(standard, goal_cases) case 1 show ?case by (auto simp: empty) next @@ -50,7 +50,7 @@ case 3 thus ?case by(simp add: insert set_ins_list) next case (4 s x) thus ?case - using delete[OF 4, of x] by (auto simp: distinct_if_sorted elems_del_list_eq) + using delete[OF 4, of x] by (auto simp: distinct_if_sorted set_del_list_eq) next case 5 thus ?case by(simp add: empty inv_empty) next diff -r 0b70405b3969 -r 30486b96274d src/HOL/Data_Structures/Tree234_Set.thy --- a/src/HOL/Data_Structures/Tree234_Set.thy Thu Mar 22 17:18:33 2018 +0100 +++ b/src/HOL/Data_Structures/Tree234_Set.thy Fri Mar 23 11:37:02 2018 +0100 @@ -202,11 +202,8 @@ subsubsection \Functional correctness of isin:\ -lemma "sorted(inorder t) \ isin t x = (x \ elems (inorder t))" -by (induction t) (auto simp: elems_simps1 ball_Un) - -lemma isin_set: "sorted(inorder t) \ isin t x = (x \ elems (inorder t))" -by (induction t) (auto simp: elems_simps2 split!: if_splits) +lemma isin_set: "sorted(inorder t) \ isin t x = (x \ set (inorder t))" +by (induction t) (auto simp: isin_simps ball_Un) subsubsection \Functional correctness of insert:\ diff -r 0b70405b3969 -r 30486b96274d src/HOL/Data_Structures/Tree23_Set.thy --- a/src/HOL/Data_Structures/Tree23_Set.thy Thu Mar 22 17:18:33 2018 +0100 +++ b/src/HOL/Data_Structures/Tree23_Set.thy Fri Mar 23 11:37:02 2018 +0100 @@ -142,11 +142,8 @@ subsubsection "Proofs for isin" -lemma "sorted(inorder t) \ isin t x = (x \ elems (inorder t))" -by (induction t) (auto simp: elems_simps1 ball_Un) - -lemma isin_set: "sorted(inorder t) \ isin t x = (x \ elems (inorder t))" -by (induction t) (auto simp: elems_simps2) +lemma isin_set: "sorted(inorder t) \ isin t x = (x \ set (inorder t))" +by (induction t) (auto simp: isin_simps ball_Un) subsubsection "Proofs for insert" diff -r 0b70405b3969 -r 30486b96274d src/HOL/Data_Structures/Tree_Set.thy --- a/src/HOL/Data_Structures/Tree_Set.thy Thu Mar 22 17:18:33 2018 +0100 +++ b/src/HOL/Data_Structures/Tree_Set.thy Fri Mar 23 11:37:02 2018 +0100 @@ -42,12 +42,8 @@ subsection "Functional Correctness Proofs" -lemma "sorted(inorder t) \ isin t x = (x \ elems (inorder t))" -by (induction t) (auto simp: elems_simps1) - -lemma isin_set: "sorted(inorder t) \ isin t x = (x \ elems (inorder t))" -by (induction t) (auto simp: elems_simps2) - +lemma isin_set: "sorted(inorder t) \ isin t x = (x \ set (inorder t))" +by (induction t) (auto simp: isin_simps) lemma inorder_insert: "sorted(inorder t) \ inorder(insert x t) = ins_list x (inorder t)"