# HG changeset patch # User haftmann # Date 1288006498 -7200 # Node ID 1d8ad2ff3e011bde4d52d7ae65743899912d847b # Parent e7a80c6752c905a55d74ab3f39d79fa302257a13 dropped (almost) redundant distinct.induct rule; distinct_simps again named distinct.simps diff -r e7a80c6752c9 -r 1d8ad2ff3e01 src/HOL/Library/Dlist.thy --- a/src/HOL/Library/Dlist.thy Mon Oct 25 13:34:57 2010 +0200 +++ b/src/HOL/Library/Dlist.thy Mon Oct 25 13:34:58 2010 +0200 @@ -143,11 +143,11 @@ proof (induct xs) case Nil from empty show ?case by (simp add: empty_def) next - case (insert x xs) + case (Cons x xs) then have "\ member (Dlist xs) x" and "P (Dlist xs)" by (simp_all add: member_def List.member_def) with insrt have "P (insert x (Dlist xs))" . - with insert show ?case by (simp add: insert_def distinct_remdups_id) + with Cons show ?case by (simp add: insert_def distinct_remdups_id) qed with dxs show "P dxs" by simp qed diff -r e7a80c6752c9 -r 1d8ad2ff3e01 src/HOL/Library/Permutation.thy --- a/src/HOL/Library/Permutation.thy Mon Oct 25 13:34:57 2010 +0200 +++ b/src/HOL/Library/Permutation.thy Mon Oct 25 13:34:58 2010 +0200 @@ -168,7 +168,7 @@ apply simp apply (subgoal_tac "a#list <~~> a#ysa@zs") apply (metis Cons_eq_appendI perm_append_Cons trans) - apply (metis Cons Cons_eq_appendI distinct_simps(2) + apply (metis Cons Cons_eq_appendI distinct.simps(2) distinct_remdups distinct_remdups_id perm_append_swap perm_distinct_iff) apply (subgoal_tac "set (a#list) = set (ysa@a#zs) & distinct (a#list) & distinct (ysa@a#zs)") apply (fastsimp simp add: insert_ident) diff -r e7a80c6752c9 -r 1d8ad2ff3e01 src/HOL/List.thy --- a/src/HOL/List.thy Mon Oct 25 13:34:57 2010 +0200 +++ b/src/HOL/List.thy Mon Oct 25 13:34:58 2010 +0200 @@ -5,7 +5,7 @@ header {* The datatype of finite lists *} theory List -imports Plain Quotient Presburger Code_Numeral Recdef +imports Plain Presburger Recdef Code_Numeral Quotient uses ("Tools/list_code.ML") begin @@ -174,15 +174,10 @@ "removeAll x [] = []" | "removeAll x (y # xs) = (if x = y then removeAll x xs else y # removeAll x xs)" -inductive +primrec distinct :: "'a list \ bool" where - Nil: "distinct []" - | insert: "x \ set xs \ distinct xs \ distinct (x # xs)" - -lemma distinct_simps [simp, code]: - "distinct [] \ True" - "distinct (x # xs) \ x \ set xs \ distinct xs" - by (auto intro: distinct.intros elim: distinct.cases) + "distinct [] \ True" + | "distinct (x # xs) \ x \ set xs \ distinct xs" primrec remdups :: "'a list \ 'a list" where @@ -786,9 +781,8 @@ by (induct xs) auto lemma map_cong [fundef_cong, recdef_cong]: -"xs = ys ==> (!!x. x : set ys ==> f x = g x) ==> map f xs = map g ys" --- {* a congruence rule for @{text map} *} -by simp + "xs = ys \ (\x. x \ set ys \ f x = g x) \ map f xs = map g ys" + by simp lemma map_is_Nil_conv [iff]: "(map f xs = []) = (xs = [])" by (cases xs) auto @@ -990,14 +984,14 @@ "(x : set xs) = (\ys zs. xs = ys @ x # zs \ x \ set ys)" by (auto dest!: split_list_first) -lemma split_list_last: "x : set xs \ \ys zs. xs = ys @ x # zs \ x \ set zs" -proof (induct xs rule:rev_induct) +lemma split_list_last: "x \ set xs \ \ys zs. xs = ys @ x # zs \ x \ set zs" +proof (induct xs rule: rev_induct) case Nil thus ?case by simp next case (snoc a xs) show ?case proof cases - assume "x = a" thus ?case using snoc by simp (metis ex_in_conv set_empty2) + assume "x = a" thus ?case using snoc by (metis List.set.simps(1) emptyE) next assume "x \ a" thus ?case using snoc by fastsimp qed @@ -1030,8 +1024,7 @@ show ?case proof cases assume "P x" - thus ?thesis by simp - (metis Un_upper1 contra_subsetD in_set_conv_decomp_first self_append_conv2 set_append) + thus ?thesis by simp (metis Un_upper1 contra_subsetD in_set_conv_decomp_first self_append_conv2 set_append) next assume "\ P x" hence "\x\set xs. P x" using Cons(2) by simp