# HG changeset patch # User wenzelm # Date 1122916825 -7200 # Node ID d3f9abe007124ba7e89597dd9edd6cd4bf366de6 # Parent 968adbfbf93bb6792e0c4473f426df8bb8f1b001 no eq_sym_conv; tuned; diff -r 968adbfbf93b -r d3f9abe00712 src/HOL/Hoare/SepLogHeap.thy --- a/src/HOL/Hoare/SepLogHeap.thy Mon Aug 01 19:20:24 2005 +0200 +++ b/src/HOL/Hoare/SepLogHeap.thy Mon Aug 01 19:20:25 2005 +0200 @@ -11,8 +11,8 @@ types heap = "(nat \ nat option)" -text{* Some means allocated, none means free. Address 0 serves as the -null reference. *} +text{* @{text "Some"} means allocated, @{text "None"} means +free. Address @{text "0"} serves as the null reference. *} subsection "Paths in the heap" @@ -23,24 +23,18 @@ "Path h x (a#as) y = (x\0 \ a=x \ (\b. h x = Some b \ Path h b as y))" lemma [iff]: "Path h 0 xs y = (xs = [] \ y = 0)" -apply(case_tac xs) -apply fastsimp -apply fastsimp -done +by (cases xs) simp_all lemma [simp]: "x\0 \ Path h x as z = (as = [] \ z = x \ (\y bs. as = x#bs \ h x = Some y & Path h y bs z))" -apply(case_tac as) -apply fastsimp -apply fastsimp -done +by (cases as) auto lemma [simp]: "\x. Path f x (as@bs) z = (\y. Path f x as y \ Path f y bs z)" -by(induct as, auto) +by (induct as) auto lemma Path_upd[simp]: "\x. u \ set as \ Path (f(u := v)) x as y = Path f x as y" -by(induct as, simp, simp add:eq_sym_conv) +by (induct as) simp_all subsection "Lists on the heap" @@ -50,34 +44,31 @@ "List h x as == Path h x as 0" lemma [simp]: "List h x [] = (x = 0)" -by(simp add:List_def) +by (simp add: List_def) lemma [simp]: "List h x (a#as) = (x\0 \ a=x \ (\y. h x = Some y \ List h y as))" -by(simp add:List_def) +by (simp add: List_def) lemma [simp]: "List h 0 as = (as = [])" -by(case_tac as, simp_all) +by (cases as) simp_all lemma List_non_null: "a\0 \ List h a as = (\b bs. as = a#bs \ h a = Some b \ List h b bs)" -by(case_tac as, simp_all) +by (cases as) simp_all theorem notin_List_update[simp]: "\x. a \ set as \ List (h(a := y)) x as = List h x as" -apply(induct as) -apply simp -apply(clarsimp simp add:fun_upd_apply) -done +by (induct as) simp_all lemma List_unique: "\x bs. List h x as \ List h x bs \ as = bs" -by(induct as, auto simp add:List_non_null) +by (induct as) (auto simp add:List_non_null) lemma List_unique1: "List h p as \ \!as. List h p as" -by(blast intro:List_unique) +by (blast intro: List_unique) lemma List_app: "\x. List h x (as@bs) = (\y. Path h x as y \ List h y bs)" -by(induct as, auto) +by (induct as) auto lemma List_hd_not_in_tl[simp]: "List h b as \ h a = Some b \ a \ set as" apply (clarsimp simp add:in_set_conv_decomp) @@ -86,19 +77,17 @@ done lemma List_distinct[simp]: "\x. List h x as \ distinct as" -apply(induct as, simp) -apply(fastsimp dest:List_hd_not_in_tl) -done +by (induct as) (auto dest:List_hd_not_in_tl) lemma list_in_heap: "\p. List h p ps \ set ps \ dom h" -by(induct ps, auto) +by (induct ps) auto lemma list_ortho_sum1[simp]: "\p. \ List h1 p ps; dom h1 \ dom h2 = {}\ \ List (h1++h2) p ps" -by(induct ps, auto simp add:map_add_def split:option.split) +by (induct ps) (auto simp add:map_add_def split:option.split) lemma list_ortho_sum2[simp]: "\p. \ List h2 p ps; dom h1 \ dom h2 = {}\ \ List (h1++h2) p ps" -by(induct ps, auto simp add:map_add_def split:option.split) +by (induct ps) (auto simp add:map_add_def split:option.split) end