# HG changeset patch # User nipkow # Date 1447758079 -3600 # Node ID f6b9f528c89c78214301057a14912cfc78e6bf7a # Parent cb595e12451d8bd5b66620e3d9978e07235a3414 converted lookup to cmp diff -r cb595e12451d -r f6b9f528c89c src/HOL/Data_Structures/AList_Upd_Del.thy --- a/src/HOL/Data_Structures/AList_Upd_Del.thy Tue Nov 17 11:44:10 2015 +0100 +++ b/src/HOL/Data_Structures/AList_Upd_Del.thy Tue Nov 17 12:01:19 2015 +0100 @@ -80,23 +80,6 @@ lemmas upd_list_simps = sorted_lems upd_list_sorted1 upd_list_sorted2 -(* -lemma set_ins_list[simp]: "set (ins_list x xs) = insert x (set xs)" -by(induction xs) auto - -lemma distinct_if_sorted: "sorted xs \ distinct xs" -apply(induction xs rule: sorted.induct) -apply auto -by (metis in_set_conv_decomp_first less_imp_not_less sorted_mid_iff2) - -lemma set_del_list_eq [simp]: "distinct xs ==> set(del_list x xs) = set xs - {x}" -apply(induct xs) - apply simp -apply simp -apply blast -done -*) - subsection \Lemmas for @{const del_list}\ diff -r cb595e12451d -r f6b9f528c89c src/HOL/Data_Structures/Lookup2.thy --- a/src/HOL/Data_Structures/Lookup2.thy Tue Nov 17 11:44:10 2015 +0100 +++ b/src/HOL/Data_Structures/Lookup2.thy Tue Nov 17 12:01:19 2015 +0100 @@ -5,17 +5,16 @@ theory Lookup2 imports Tree2 + Cmp Map_by_Ordered begin -fun lookup :: "('a::linorder * 'b, 'c) tree \ 'a \ 'b option" where +fun lookup :: "('a::cmp * 'b, 'c) tree \ 'a \ 'b option" where "lookup Leaf x = None" | "lookup (Node _ l (a,b) r) x = - (if x < a then lookup l x else - if x > a then lookup r x else Some b)" + (case cmp x a of LT \ lookup l x | GT \ lookup r x | EQ \ Some b)" -lemma lookup_eq: - "sorted1(inorder t) \ lookup t x = map_of (inorder t) x" +lemma lookup_eq: "sorted1(inorder t) \ lookup t x = map_of (inorder t) x" by(induction t) (auto simp: map_of_simps split: option.split) end \ No newline at end of file