# HG changeset patch # User nipkow # Date 1113146343 -7200 # Node ID 900cf45ff0a611433b47da5d38c57b88559cf05c # Parent 1da2cfd1ca452ae10d347a09d1c33d07f1a2ff2a _(_|_) is now override_on diff -r 1da2cfd1ca45 -r 900cf45ff0a6 src/HOL/Fun.thy --- a/src/HOL/Fun.thy Sun Apr 10 11:42:07 2005 +0200 +++ b/src/HOL/Fun.thy Sun Apr 10 17:19:03 2005 +0200 @@ -39,9 +39,8 @@ *) constdefs - overwrite :: "('a => 'b) => ('a => 'b) => 'a set => ('a => 'b)" - ("_/'(_|/_')" [900,0,0]900) -"f(g|A) == %a. if a : A then g a else f a" + override_on :: "('a => 'b) => ('a => 'b) => 'a set => ('a => 'b)" +"override_on f g A == %a. if a : A then g a else f a" id :: "'a => 'a" "id == %x. x" @@ -392,16 +391,16 @@ "f(x:=y) ` A = (if x \ A then insert y (f ` (A-{x})) else f ` A)" by auto -subsection{* overwrite *} +subsection{* @{text override_on} *} -lemma overwrite_emptyset[simp]: "f(g|{}) = f" -by(simp add:overwrite_def) +lemma override_on_emptyset[simp]: "override_on f g {} = f" +by(simp add:override_on_def) -lemma overwrite_apply_notin[simp]: "a ~: A ==> (f(g|A)) a = f a" -by(simp add:overwrite_def) +lemma override_on_apply_notin[simp]: "a ~: A ==> (override_on f g A) a = f a" +by(simp add:override_on_def) -lemma overwrite_apply_in[simp]: "a : A ==> (f(g|A)) a = g a" -by(simp add:overwrite_def) +lemma override_on_apply_in[simp]: "a : A ==> (override_on f g A) a = g a" +by(simp add:override_on_def) subsection{* swap *} @@ -410,7 +409,7 @@ "swap a b f == f(a := f b, b:= f a)" lemma swap_self: "swap a a f = f" -by (simp add: swap_def) +by (simp add: swap_def) lemma swap_commute: "swap a b f = swap b a f" by (rule ext, simp add: fun_upd_def swap_def) diff -r 1da2cfd1ca45 -r 900cf45ff0a6 src/HOL/Map.thy --- a/src/HOL/Map.thy Sun Apr 10 11:42:07 2005 +0200 +++ b/src/HOL/Map.thy Sun Apr 10 17:19:03 2005 +0200 @@ -18,7 +18,7 @@ consts chg_map :: "('b => 'b) => 'a => ('a ~=> 'b) => ('a ~=> 'b)" map_add :: "('a ~=> 'b) => ('a ~=> 'b) => ('a ~=> 'b)" (infixl "++" 100) -restrict_map :: "('a ~=> 'b) => 'a set => ('a ~=> 'b)" ("_|'__" [90, 91] 90) +restrict_map :: "('a ~=> 'b) => 'a set => ('a ~=> 'b)" (infixl "|^" 110) dom :: "('a ~=> 'b) => 'a set" ran :: "('a ~=> 'b) => 'b set" map_of :: "('a * 'b)list => 'a ~=> 'b" @@ -55,7 +55,7 @@ "_maplet" :: "['a, 'a] => maplet" ("_ /\/ _") "_maplets" :: "['a, 'a] => maplet" ("_ /[\]/ _") - restrict_map :: "('a ~=> 'b) => 'a set => ('a ~=> 'b)" ("_\_" [90, 91] 90) + restrict_map :: "('a ~=> 'b) => 'a set => ('a ~=> 'b)" (infixl "\" 110) --"requires amssymb!" map_upd_s :: "('a ~=> 'b) => 'a set => 'b => ('a ~=> 'b)" ("_/'(_/{\}/_')" [900,0,0]900) map_subst :: "('a ~=> 'b) => 'b => 'b => @@ -80,7 +80,7 @@ chg_map_def: "chg_map f a m == case m a of None => m | Some b => m(a|->f b)" map_add_def: "m1++m2 == %x. case m2 x of None => m1 x | Some y => Some y" -restrict_map_def: "m|_A == %x. if x : A then m x else None" +restrict_map_def: "m|^A == %x. if x : A then m x else None" map_upds_def: "m(xs [|->] ys) == m ++ map_of (rev(zip xs ys))" map_upd_s_def: "m(as{|->}b) == %x. if x : as then Some b else m x" @@ -324,44 +324,44 @@ subsection {* @{term restrict_map} *} -lemma restrict_map_to_empty[simp]: "m\{} = empty" +lemma restrict_map_to_empty[simp]: "m|^{} = empty" by(simp add: restrict_map_def) -lemma restrict_map_empty[simp]: "empty\D = empty" +lemma restrict_map_empty[simp]: "empty|^D = empty" by(simp add: restrict_map_def) -lemma restrict_in [simp]: "x \ A \ (m\A) x = m x" +lemma restrict_in [simp]: "x \ A \ (m|^A) x = m x" by (auto simp: restrict_map_def) -lemma restrict_out [simp]: "x \ A \ (m\A) x = None" +lemma restrict_out [simp]: "x \ A \ (m|^A) x = None" by (auto simp: restrict_map_def) -lemma ran_restrictD: "y \ ran (m\A) \ \x\A. m x = Some y" +lemma ran_restrictD: "y \ ran (m|^A) \ \x\A. m x = Some y" by (auto simp: restrict_map_def ran_def split: split_if_asm) -lemma dom_restrict [simp]: "dom (m\A) = dom m \ A" +lemma dom_restrict [simp]: "dom (m|^A) = dom m \ A" by (auto simp: restrict_map_def dom_def split: split_if_asm) -lemma restrict_upd_same [simp]: "m(x\y)\(-{x}) = m\(-{x})" +lemma restrict_upd_same [simp]: "m(x\y)|^(-{x}) = m|^(-{x})" by (rule ext, auto simp: restrict_map_def) -lemma restrict_restrict [simp]: "m\A\B = m\(A\B)" +lemma restrict_restrict [simp]: "m|^A|^B = m|^(A\B)" by (rule ext, auto simp: restrict_map_def) lemma restrict_fun_upd[simp]: - "m(x := y)\D = (if x \ D then (m\(D-{x}))(x := y) else m\D)" + "m(x := y)|^D = (if x \ D then (m|^(D-{x}))(x := y) else m|^D)" by(simp add: restrict_map_def expand_fun_eq) lemma fun_upd_None_restrict[simp]: - "(m\D)(x := None) = (if x:D then m\(D - {x}) else m\D)" + "(m|^D)(x := None) = (if x:D then m|^(D - {x}) else m|^D)" by(simp add: restrict_map_def expand_fun_eq) lemma fun_upd_restrict: - "(m\D)(x := y) = (m\(D-{x}))(x := y)" + "(m|^D)(x := y) = (m|^(D-{x}))(x := y)" by(simp add: restrict_map_def expand_fun_eq) lemma fun_upd_restrict_conv[simp]: - "x \ D \ (m\D)(x := y) = (m\(D-{x}))(x := y)" + "x \ D \ (m|^D)(x := y) = (m|^(D-{x}))(x := y)" by(simp add: restrict_map_def expand_fun_eq) @@ -432,7 +432,7 @@ lemma restrict_map_upds[simp]: "!!m ys. \ length xs = length ys; set xs \ D \ - \ m(xs [\] ys)\D = (m\(D - set xs))(xs [\] ys)" + \ m(xs [\] ys)|^D = (m|^(D - set xs))(xs [\] ys)" apply (induct xs, simp) apply (case_tac ys, simp) apply(simp add:Diff_insert[symmetric] insert_absorb) @@ -500,9 +500,10 @@ lemma dom_map_add[simp]: "dom(m++n) = dom n Un dom m" by (unfold dom_def, auto) -lemma dom_overwrite[simp]: - "dom(f(g|A)) = (dom f - {a. a : A - dom g}) Un {a. a : A Int dom g}" -by(auto simp add: dom_def overwrite_def) +lemma dom_override_on[simp]: + "dom(override_on f g A) = + (dom f - {a. a : A - dom g}) Un {a. a : A Int dom g}" +by(auto simp add: dom_def override_on_def) lemma map_add_comm: "dom m1 \ dom m2 = {} \ m1++m2 = m2++m1" apply(rule ext) diff -r 1da2cfd1ca45 -r 900cf45ff0a6 src/HOL/document/root.tex --- a/src/HOL/document/root.tex Sun Apr 10 11:42:07 2005 +0200 +++ b/src/HOL/document/root.tex Sun Apr 10 17:19:03 2005 +0200 @@ -3,6 +3,7 @@ \documentclass[11pt,a4paper]{article} \usepackage{graphicx,isabelle,isabellesym,latexsym} +\usepackage{amssymb} \usepackage[only,bigsqcap]{stmaryrd} \usepackage[latin1]{inputenc} \usepackage{pdfsetup}