# HG changeset patch # User nipkow # Date 1678951800 -3600 # Node ID 8a6a79ed5a83199a9ef94fc98f19a4aa22b419e1 # Parent 5cb7fd36223b7e352053d51be973bc6d9c1547a9 unified function update and map update syntaxes diff -r 5cb7fd36223b -r 8a6a79ed5a83 NEWS --- a/NEWS Tue Mar 14 22:00:06 2023 +0100 +++ b/NEWS Thu Mar 16 08:30:00 2023 +0100 @@ -52,7 +52,16 @@ Minor INCOMPATIBILITY. * Theory HOL.Map: - - Map.empty has been demoted to an input abbreviation only + - Map.empty has been demoted to an input abbreviation. + - Map update syntax "_(_ \ _)" now has the same priorities + as function update syntax "_(_ := _)". This means: + 1. "f t(a \ b)" must now be written "(f t)(a \ b)" + 2. "t(a \ b)(c \ d)" must now be written + "t(a \ b, c \ d)" or + "(t(a \ b))(c \ d)". + Moreover, ":=" and "\" can be mixed freely now. + Except in "[...]" maps where ":=" would create a clash with + list update syntax "xs[i := x]". * Theory "HOL.Relation": - Added predicates irrefl_on and irreflp_on and redefined irrefl and diff -r 5cb7fd36223b -r 8a6a79ed5a83 src/HOL/Map.thy --- a/src/HOL/Map.thy Tue Mar 14 22:00:06 2023 +0100 +++ b/src/HOL/Map.thy Thu Mar 16 08:30:00 2023 +0100 @@ -49,43 +49,61 @@ map_le :: "('a \ 'b) \ ('a \ 'b) \ bool" (infix "\\<^sub>m" 50) where "(m\<^sub>1 \\<^sub>m m\<^sub>2) \ (\a \ dom m\<^sub>1. m\<^sub>1 a = m\<^sub>2 a)" -nonterminal maplets and maplet +text \Function update syntax \f(x := y, \)\ is extended with \x \ y\, which is short for +\x := Some y\. \:=\ and \\\ can be mixed freely. +The syntax \[x \ y, \]\ is short for \Map.empty(x \ y, \)\ +but must only contain \\\, not \:=\, because \[x:=y]\ clashes with the list update syntax \xs[i:=x]\.\ + +nonterminal maplet and maplets syntax "_maplet" :: "['a, 'a] \ maplet" ("_ /\/ _") - "_maplets" :: "['a, 'a] \ maplet" ("_ /[\]/ _") + "" :: "maplet \ updbind" ("_") "" :: "maplet \ maplets" ("_") "_Maplets" :: "[maplet, maplets] \ maplets" ("_,/ _") - "_MapUpd" :: "['a \ 'b, maplets] \ 'a \ 'b" ("_/'(_')" [1000, 0] 900) - "_Map" :: "maplets \ 'a \ 'b" ("(1[_])") + "_Map" :: "maplets \ 'a \ 'b" ("(1[_])") +(* Syntax forbids \[\, x := y, \]\ by introducing \maplets\ in addition to \updbinds\ *) syntax (ASCII) "_maplet" :: "['a, 'a] \ maplet" ("_ /|->/ _") - "_maplets" :: "['a, 'a] \ maplet" ("_ /[|->]/ _") translations - "_MapUpd m (_Maplets xy ms)" \ "_MapUpd (_MapUpd m xy) ms" - "_MapUpd m (_maplet x y)" \ "m(x := CONST Some y)" - "_Map ms" \ "_MapUpd (CONST empty) ms" - "_Map ms" \ "_MapUpd (\x. CONST None) ms" \\both are needed\ - "_Map (_Maplets ms1 ms2)" \ "_MapUpd (_Map ms1) ms2" - "_Maplets ms1 (_Maplets ms2 ms3)" \ "_Maplets (_Maplets ms1 ms2) ms3" + "_Update f (_maplet x y)" \ "f(x := CONST Some y)" + "_Maplets m ms" \ "_updbinds m ms" + "_Map ms" \ "_Update (CONST empty) ms" -primrec map_of :: "('a \ 'b) list \ 'a \ 'b" -where +(* Printing must create \_Map\ only for \_maplet\ *) + "_Map (_maplet x y)" \ "_Update (\u. CONST None) (_maplet x y)" + "_Map (_updbinds m (_maplet x y))" \ "_Update (_Map m) (_maplet x y)" + +text \Updating with lists:\ + +primrec map_of :: "('a \ 'b) list \ 'a \ 'b" where "map_of [] = empty" | "map_of (p # ps) = (map_of ps)(fst p \ snd p)" -definition map_upds :: "('a \ 'b) \ 'a list \ 'b list \ 'a \ 'b" - where "map_upds m xs ys = m ++ map_of (rev (zip xs ys))" -translations - "_MapUpd m (_maplets x y)" \ "CONST map_upds m x y" - lemma map_of_Cons_code [code]: "map_of [] k = None" "map_of ((l, v) # ps) k = (if l = k then Some v else map_of ps k)" by simp_all +definition map_upds :: "('a \ 'b) \ 'a list \ 'b list \ 'a \ 'b" where +"map_upds m xs ys = m ++ map_of (rev (zip xs ys))" + +text \There is also the more specialized update syntax \xs [\] ys\ for lists \xs\ and \ys\.\ + +syntax + "_maplets" :: "['a, 'a] \ maplet" ("_ /[\]/ _") + +syntax (ASCII) + "_maplets" :: "['a, 'a] \ maplet" ("_ /[|->]/ _") + +translations + "_Update m (_maplets xs ys)" \ "CONST map_upds m xs ys" + + "_Map (_maplets xs ys)" \ "_Update (\u. CONST None) (_maplets xs ys)" + "_Map (_updbinds m (_maplets xs ys))" \ "_Update (_Map m) (_maplets xs ys)" + subsection \@{term [source] empty}\