tuned Map, renamed lex stuff in List.
--- a/src/HOL/List.ML Sun Apr 10 17:20:03 2005 +0200
+++ b/src/HOL/List.ML Mon Apr 11 12:14:23 2005 +0200
@@ -90,8 +90,8 @@
val length_zip = thm "length_zip";
val lex_conv = thm "lex_conv";
val lex_def = thm "lex_def";
-val lexico_conv = thm "lexico_conv";
-val lexico_def = thm "lexico_def";
+val lenlex_conv = thm "lenlex_conv";
+val lenlex_def = thm "lenlex_def";
val lexn_conv = thm "lexn_conv";
val lexn_length = thm "lexn_length";
val list_all2_Cons = thm "list_all2_Cons";
@@ -222,7 +222,7 @@
val upt_conv_Nil = thm "upt_conv_Nil";
val upt_rec = thm "upt_rec";
val wf_lex = thm "wf_lex";
-val wf_lexico = thm "wf_lexico";
+val wf_lenlex = thm "wf_lenlex";
val wf_lexn = thm "wf_lexn";
val zip_Cons_Cons = thm "zip_Cons_Cons";
val zip_Nil = thm "zip_Nil";
--- a/src/HOL/List.thy Sun Apr 10 17:20:03 2005 +0200
+++ b/src/HOL/List.thy Mon Apr 11 12:14:23 2005 +0200
@@ -1909,8 +1909,8 @@
"lex r == \<Union>n. lexn r n"
--{*Holds only between lists of the same length*}
- lexico :: "('a \<times> 'a) set => ('a list \<times> 'a list) set"
- "lexico r == inv_image (less_than <*lex*> lex r) (%xs. (length xs, xs))"
+ lenlex :: "('a \<times> 'a) set => ('a list \<times> 'a list) set"
+ "lenlex r == inv_image (less_than <*lex*> lex r) (%xs. (length xs, xs))"
--{*Compares lists by their length and then lexicographically*}
@@ -1952,13 +1952,13 @@
(\<exists>xys x y xs' ys'. xs = xys @ x # xs' \<and> ys = xys @ y # ys' \<and> (x, y):r)}"
by (force simp add: lex_def lexn_conv)
-lemma wf_lexico [intro!]: "wf r ==> wf (lexico r)"
-by (unfold lexico_def) blast
-
-lemma lexico_conv:
- "lexico r = {(xs,ys). length xs < length ys |
+lemma wf_lenlex [intro!]: "wf r ==> wf (lenlex r)"
+by (unfold lenlex_def) blast
+
+lemma lenlex_conv:
+ "lenlex r = {(xs,ys). length xs < length ys |
length xs = length ys \<and> (xs, ys) : lex r}"
-by (simp add: lexico_def diag_def lex_prod_def measure_def inv_image_def)
+by (simp add: lenlex_def diag_def lex_prod_def measure_def inv_image_def)
lemma Nil_notin_lex [iff]: "([], ys) \<notin> lex r"
by (simp add: lex_conv)
--- a/src/HOL/Map.thy Sun Apr 10 17:20:03 2005 +0200
+++ b/src/HOL/Map.thy Mon Apr 11 12:14:23 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)" (infixl "|^" 110)
+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,6 @@
"_maplet" :: "['a, 'a] => maplet" ("_ /\<mapsto>/ _")
"_maplets" :: "['a, 'a] => maplet" ("_ /[\<mapsto>]/ _")
- restrict_map :: "('a ~=> 'b) => 'a set => ('a ~=> 'b)" (infixl "\<upharpoonright>" 110) --"requires amssymb!"
map_upd_s :: "('a ~=> 'b) => 'a set => 'b => ('a ~=> 'b)"
("_/'(_/{\<mapsto>}/_')" [900,0,0]900)
map_subst :: "('a ~=> 'b) => 'b => 'b =>
@@ -63,6 +62,9 @@
"@chg_map" :: "('a ~=> 'b) => 'a => ('b => 'b) => ('a ~=> 'b)"
("_/'(_/\<mapsto>\<lambda>_. _')" [900,0,0,0] 900)
+syntax (latex output)
+ restrict_map :: "('a ~=> 'b) => 'a set => ('a ~=> 'b)" ("_\<restriction>\<^bsub>_\<^esub>" [111,110] 110) --"requires amssymb!"
+
translations
"empty" => "_K None"
"empty" <= "%x. None"
@@ -80,7 +82,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 +326,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 \<in> A \<Longrightarrow> (m|^A) x = m x"
+lemma restrict_in [simp]: "x \<in> A \<Longrightarrow> (m|`A) x = m x"
by (auto simp: restrict_map_def)
-lemma restrict_out [simp]: "x \<notin> A \<Longrightarrow> (m|^A) x = None"
+lemma restrict_out [simp]: "x \<notin> A \<Longrightarrow> (m|`A) x = None"
by (auto simp: restrict_map_def)
-lemma ran_restrictD: "y \<in> ran (m|^A) \<Longrightarrow> \<exists>x\<in>A. m x = Some y"
+lemma ran_restrictD: "y \<in> ran (m|`A) \<Longrightarrow> \<exists>x\<in>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 \<inter> A"
+lemma dom_restrict [simp]: "dom (m|`A) = dom m \<inter> A"
by (auto simp: restrict_map_def dom_def split: split_if_asm)
-lemma restrict_upd_same [simp]: "m(x\<mapsto>y)|^(-{x}) = m|^(-{x})"
+lemma restrict_upd_same [simp]: "m(x\<mapsto>y)|`(-{x}) = m|`(-{x})"
by (rule ext, auto simp: restrict_map_def)
-lemma restrict_restrict [simp]: "m|^A|^B = m|^(A\<inter>B)"
+lemma restrict_restrict [simp]: "m|`A|`B = m|`(A\<inter>B)"
by (rule ext, auto simp: restrict_map_def)
lemma restrict_fun_upd[simp]:
- "m(x := y)|^D = (if x \<in> D then (m|^(D-{x}))(x := y) else m|^D)"
+ "m(x := y)|`D = (if x \<in> 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 \<in> D \<Longrightarrow> (m|^D)(x := y) = (m|^(D-{x}))(x := y)"
+ "x \<in> D \<Longrightarrow> (m|`D)(x := y) = (m|`(D-{x}))(x := y)"
by(simp add: restrict_map_def expand_fun_eq)
@@ -432,7 +434,7 @@
lemma restrict_map_upds[simp]: "!!m ys.
\<lbrakk> length xs = length ys; set xs \<subseteq> D \<rbrakk>
- \<Longrightarrow> m(xs [\<mapsto>] ys)|^D = (m|^(D - set xs))(xs [\<mapsto>] ys)"
+ \<Longrightarrow> m(xs [\<mapsto>] ys)|`D = (m|`(D - set xs))(xs [\<mapsto>] ys)"
apply (induct xs, simp)
apply (case_tac ys, simp)
apply(simp add:Diff_insert[symmetric] insert_absorb)