--- a/NEWS Tue Jul 12 13:55:35 2016 +0200
+++ b/NEWS Tue Jul 12 16:04:19 2016 +0200
@@ -76,6 +76,9 @@
* Highlighting of entity def/ref positions wrt. cursor.
+* Refined folding mode "isabelle" based on Isar syntax: 'next' and 'qed'
+are treated as delimiters for fold structure.
+
* Improved support for indentation according to Isabelle outer syntax.
Action "indent-lines" (shortcut C+i) indents the current line according
to command keywords and some command substructure. Action
--- a/src/HOL/Library/AList.thy Tue Jul 12 13:55:35 2016 +0200
+++ b/src/HOL/Library/AList.thy Tue Jul 12 16:04:19 2016 +0200
@@ -5,7 +5,7 @@
section \<open>Implementation of Association Lists\<close>
theory AList
-imports Main
+ imports Main
begin
context
@@ -21,9 +21,9 @@
subsection \<open>\<open>update\<close> and \<open>updates\<close>\<close>
qualified primrec update :: "'key \<Rightarrow> 'val \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
-where
- "update k v [] = [(k, v)]"
-| "update k v (p # ps) = (if fst p = k then (k, v) # ps else p # update k v ps)"
+ where
+ "update k v [] = [(k, v)]"
+ | "update k v (p # ps) = (if fst p = k then (k, v) # ps else p # update k v ps)"
lemma update_conv': "map_of (update k v al) = (map_of al)(k\<mapsto>v)"
by (induct al) (auto simp add: fun_eq_iff)
@@ -82,12 +82,11 @@
x = k \<and> v = y \<or> x \<noteq> k \<and> map_of al x = Some y"
by (simp add: update_conv' map_upd_Some_unfold)
-lemma image_update [simp]:
- "x \<notin> A \<Longrightarrow> map_of (update x y al) ` A = map_of al ` A"
+lemma image_update [simp]: "x \<notin> A \<Longrightarrow> map_of (update x y al) ` A = map_of al ` A"
by (simp add: update_conv')
-qualified definition
- updates :: "'key list \<Rightarrow> 'val list \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
+qualified definition updates
+ :: "'key list \<Rightarrow> 'val list \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
where "updates ks vs = fold (case_prod update) (zip ks vs)"
lemma updates_simps [simp]:
@@ -217,66 +216,75 @@
subsection \<open>\<open>update_with_aux\<close> and \<open>delete_aux\<close>\<close>
-qualified primrec update_with_aux :: "'val \<Rightarrow> 'key \<Rightarrow> ('val \<Rightarrow> 'val) \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
-where
- "update_with_aux v k f [] = [(k, f v)]"
-| "update_with_aux v k f (p # ps) = (if (fst p = k) then (k, f (snd p)) # ps else p # update_with_aux v k f ps)"
+qualified primrec update_with_aux
+ :: "'val \<Rightarrow> 'key \<Rightarrow> ('val \<Rightarrow> 'val) \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
+ where
+ "update_with_aux v k f [] = [(k, f v)]"
+ | "update_with_aux v k f (p # ps) =
+ (if (fst p = k) then (k, f (snd p)) # ps else p # update_with_aux v k f ps)"
text \<open>
The above @{term "delete"} traverses all the list even if it has found the key.
This one does not have to keep going because is assumes the invariant that keys are distinct.
\<close>
qualified fun delete_aux :: "'key \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
-where
- "delete_aux k [] = []"
-| "delete_aux k ((k', v) # xs) = (if k = k' then xs else (k', v) # delete_aux k xs)"
+ where
+ "delete_aux k [] = []"
+ | "delete_aux k ((k', v) # xs) = (if k = k' then xs else (k', v) # delete_aux k xs)"
lemma map_of_update_with_aux':
- "map_of (update_with_aux v k f ps) k' = ((map_of ps)(k \<mapsto> (case map_of ps k of None \<Rightarrow> f v | Some v \<Rightarrow> f v))) k'"
-by(induct ps) auto
+ "map_of (update_with_aux v k f ps) k' =
+ ((map_of ps)(k \<mapsto> (case map_of ps k of None \<Rightarrow> f v | Some v \<Rightarrow> f v))) k'"
+ by (induct ps) auto
lemma map_of_update_with_aux:
- "map_of (update_with_aux v k f ps) = (map_of ps)(k \<mapsto> (case map_of ps k of None \<Rightarrow> f v | Some v \<Rightarrow> f v))"
-by(simp add: fun_eq_iff map_of_update_with_aux')
+ "map_of (update_with_aux v k f ps) =
+ (map_of ps)(k \<mapsto> (case map_of ps k of None \<Rightarrow> f v | Some v \<Rightarrow> f v))"
+ by (simp add: fun_eq_iff map_of_update_with_aux')
lemma dom_update_with_aux: "fst ` set (update_with_aux v k f ps) = {k} \<union> fst ` set ps"
by (induct ps) auto
lemma distinct_update_with_aux [simp]:
"distinct (map fst (update_with_aux v k f ps)) = distinct (map fst ps)"
-by(induct ps)(auto simp add: dom_update_with_aux)
+ by (induct ps) (auto simp add: dom_update_with_aux)
lemma set_update_with_aux:
- "distinct (map fst xs)
- \<Longrightarrow> set (update_with_aux v k f xs) = (set xs - {k} \<times> UNIV \<union> {(k, f (case map_of xs k of None \<Rightarrow> v | Some v \<Rightarrow> v))})"
-by(induct xs)(auto intro: rev_image_eqI)
+ "distinct (map fst xs) \<Longrightarrow>
+ set (update_with_aux v k f xs) =
+ (set xs - {k} \<times> UNIV \<union> {(k, f (case map_of xs k of None \<Rightarrow> v | Some v \<Rightarrow> v))})"
+ by (induct xs) (auto intro: rev_image_eqI)
lemma set_delete_aux: "distinct (map fst xs) \<Longrightarrow> set (delete_aux k xs) = set xs - {k} \<times> UNIV"
-apply(induct xs)
-apply simp_all
-apply clarsimp
-apply(fastforce intro: rev_image_eqI)
-done
+ apply (induct xs)
+ apply simp_all
+ apply clarsimp
+ apply (fastforce intro: rev_image_eqI)
+ done
lemma dom_delete_aux: "distinct (map fst ps) \<Longrightarrow> fst ` set (delete_aux k ps) = fst ` set ps - {k}"
-by(auto simp add: set_delete_aux)
+ by (auto simp add: set_delete_aux)
-lemma distinct_delete_aux [simp]:
- "distinct (map fst ps) \<Longrightarrow> distinct (map fst (delete_aux k ps))"
-proof(induct ps)
- case Nil thus ?case by simp
+lemma distinct_delete_aux [simp]: "distinct (map fst ps) \<Longrightarrow> distinct (map fst (delete_aux k ps))"
+proof (induct ps)
+ case Nil
+ then show ?case by simp
next
case (Cons a ps)
- obtain k' v where a: "a = (k', v)" by(cases a)
+ obtain k' v where a: "a = (k', v)"
+ by (cases a)
show ?case
- proof(cases "k' = k")
- case True with Cons a show ?thesis by simp
+ proof (cases "k' = k")
+ case True
+ with Cons a show ?thesis by simp
next
case False
- with Cons a have "k' \<notin> fst ` set ps" "distinct (map fst ps)" by simp_all
+ with Cons a have "k' \<notin> fst ` set ps" "distinct (map fst ps)"
+ by simp_all
with False a have "k' \<notin> fst ` set (delete_aux k ps)"
- by(auto dest!: dom_delete_aux[where k=k])
- with Cons a show ?thesis by simp
+ by (auto dest!: dom_delete_aux[where k=k])
+ with Cons a show ?thesis
+ by simp
qed
qed
@@ -290,10 +298,10 @@
lemma map_of_delete_aux:
"distinct (map fst xs) \<Longrightarrow> map_of (delete_aux k xs) k' = ((map_of xs)(k := None)) k'"
-by(simp add: map_of_delete_aux')
+ by (simp add: map_of_delete_aux')
lemma delete_aux_eq_Nil_conv: "delete_aux k ts = [] \<longleftrightarrow> ts = [] \<or> (\<exists>v. ts = [(k, v)])"
-by(cases ts)(auto split: if_split_asm)
+ by (cases ts) (auto split: if_split_asm)
subsection \<open>\<open>restrict\<close>\<close>
@@ -308,16 +316,18 @@
lemma restr_conv': "map_of (restrict A al) = ((map_of al)|` A)"
proof
- fix k
- show "map_of (restrict A al) k = ((map_of al)|` A) k"
- by (induct al) (simp, cases "k \<in> A", auto)
+ show "map_of (restrict A al) k = ((map_of al)|` A) k" for k
+ apply (induct al)
+ apply simp
+ apply (cases "k \<in> A")
+ apply auto
+ done
qed
corollary restr_conv: "map_of (restrict A al) k = ((map_of al)|` A) k"
by (simp add: restr_conv')
-lemma distinct_restr:
- "distinct (map fst al) \<Longrightarrow> distinct (map fst (restrict A al))"
+lemma distinct_restr: "distinct (map fst al) \<Longrightarrow> distinct (map fst (restrict A al))"
by (induct al) (auto simp add: restrict_eq)
lemma restr_empty [simp]:
@@ -341,8 +351,8 @@
by (induct al) (auto simp add: restrict_eq)
lemma restr_update[simp]:
- "map_of (restrict D (update x y al)) =
- map_of ((if x \<in> D then (update x y (restrict (D-{x}) al)) else restrict D al))"
+ "map_of (restrict D (update x y al)) =
+ map_of ((if x \<in> D then (update x y (restrict (D-{x}) al)) else restrict D al))"
by (simp add: restr_conv' update_conv')
lemma restr_delete [simp]:
@@ -350,12 +360,12 @@
apply (simp add: delete_eq restrict_eq)
apply (auto simp add: split_def)
proof -
- have "\<And>y. y \<noteq> x \<longleftrightarrow> x \<noteq> y"
+ have "y \<noteq> x \<longleftrightarrow> x \<noteq> y" for y
by auto
then show "[p \<leftarrow> al. fst p \<in> D \<and> x \<noteq> fst p] = [p \<leftarrow> al. fst p \<in> D \<and> fst p \<noteq> x]"
by simp
assume "x \<notin> D"
- then have "\<And>y. y \<in> D \<longleftrightarrow> y \<in> D \<and> x \<noteq> y"
+ then have "y \<in> D \<longleftrightarrow> y \<in> D \<and> x \<noteq> y" for y
by auto
then show "[p \<leftarrow> al . fst p \<in> D \<and> x \<noteq> fst p] = [p \<leftarrow> al . fst p \<in> D]"
by simp
@@ -383,9 +393,9 @@
subsection \<open>\<open>clearjunk\<close>\<close>
qualified function clearjunk :: "('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
-where
- "clearjunk [] = []"
-| "clearjunk (p#ps) = p # clearjunk (delete (fst p) ps)"
+ where
+ "clearjunk [] = []"
+ | "clearjunk (p#ps) = p # clearjunk (delete (fst p) ps)"
by pat_completeness auto
termination
by (relation "measure length") (simp_all add: less_Suc_eq_le length_delete_le)
@@ -420,7 +430,7 @@
lemma clearjunk_updates: "clearjunk (updates ks vs al) = updates ks vs (clearjunk al)"
proof -
have "clearjunk \<circ> fold (case_prod update) (zip ks vs) =
- fold (case_prod update) (zip ks vs) \<circ> clearjunk"
+ fold (case_prod update) (zip ks vs) \<circ> clearjunk"
by (rule fold_commute) (simp add: clearjunk_update case_prod_beta o_def)
then show ?thesis
by (simp add: updates_def fun_eq_iff)
@@ -506,10 +516,8 @@
lemma dom_merge: "fst ` set (merge xs ys) = fst ` set xs \<union> fst ` set ys"
by (induct ys arbitrary: xs) (auto simp add: dom_update)
-lemma distinct_merge:
- assumes "distinct (map fst xs)"
- shows "distinct (map fst (merge xs ys))"
- using assms by (simp add: merge_updates distinct_updates)
+lemma distinct_merge: "distinct (map fst xs) \<Longrightarrow> distinct (map fst (merge xs ys))"
+ by (simp add: merge_updates distinct_updates)
lemma clearjunk_merge: "clearjunk (merge xs ys) = merge (clearjunk xs) ys"
by (simp add: merge_updates clearjunk_updates)
@@ -542,12 +550,10 @@
lemma merge_find_right [simp]: "map_of n k = Some v \<Longrightarrow> map_of (merge m n) k = Some v"
by (simp add: merge_conv')
-lemma merge_None [iff]:
- "(map_of (merge m n) k = None) = (map_of n k = None \<and> map_of m k = None)"
+lemma merge_None [iff]: "(map_of (merge m n) k = None) = (map_of n k = None \<and> map_of m k = None)"
by (simp add: merge_conv')
-lemma merge_upd [simp]:
- "map_of (merge m (update k v n)) = map_of (update k v (merge m n))"
+lemma merge_upd [simp]: "map_of (merge m (update k v n)) = map_of (update k v (merge m n))"
by (simp add: update_conv' merge_conv')
lemma merge_updatess [simp]:
@@ -561,20 +567,18 @@
subsection \<open>\<open>compose\<close>\<close>
qualified function compose :: "('key \<times> 'a) list \<Rightarrow> ('a \<times> 'b) list \<Rightarrow> ('key \<times> 'b) list"
-where
- "compose [] ys = []"
-| "compose (x # xs) ys =
- (case map_of ys (snd x) of
- None \<Rightarrow> compose (delete (fst x) xs) ys
- | Some v \<Rightarrow> (fst x, v) # compose xs ys)"
+ where
+ "compose [] ys = []"
+ | "compose (x # xs) ys =
+ (case map_of ys (snd x) of
+ None \<Rightarrow> compose (delete (fst x) xs) ys
+ | Some v \<Rightarrow> (fst x, v) # compose xs ys)"
by pat_completeness auto
termination
by (relation "measure (length \<circ> fst)") (simp_all add: less_Suc_eq_le length_delete_le)
-lemma compose_first_None [simp]:
- assumes "map_of xs k = None"
- shows "map_of (compose xs ys) k = None"
- using assms by (induct xs ys rule: compose.induct) (auto split: option.splits if_split_asm)
+lemma compose_first_None [simp]: "map_of xs k = None \<Longrightarrow> map_of (compose xs ys) k = None"
+ by (induct xs ys rule: compose.induct) (auto split: option.splits if_split_asm)
lemma compose_conv: "map_of (compose xs ys) k = (map_of ys \<circ>\<^sub>m map_of xs) k"
proof (induct xs ys rule: compose.induct)
@@ -617,10 +621,8 @@
lemma compose_conv': "map_of (compose xs ys) = (map_of ys \<circ>\<^sub>m map_of xs)"
by (rule ext) (rule compose_conv)
-lemma compose_first_Some [simp]:
- assumes "map_of xs k = Some v"
- shows "map_of (compose xs ys) k = map_of ys v"
- using assms by (simp add: compose_conv)
+lemma compose_first_Some [simp]: "map_of xs k = Some v \<Longrightarrow> map_of (compose xs ys) k = map_of ys v"
+ by (simp add: compose_conv)
lemma dom_compose: "fst ` set (compose xs ys) \<subseteq> fst ` set xs"
proof (induct xs ys rule: compose.induct)
@@ -631,19 +633,15 @@
show ?case
proof (cases "map_of ys (snd x)")
case None
- with "2.hyps"
- have "fst ` set (compose (delete (fst x) xs) ys) \<subseteq> fst ` set (delete (fst x) xs)"
+ with "2.hyps" have "fst ` set (compose (delete (fst x) xs) ys) \<subseteq> fst ` set (delete (fst x) xs)"
by simp
- also
- have "\<dots> \<subseteq> fst ` set xs"
+ also have "\<dots> \<subseteq> fst ` set xs"
by (rule dom_delete_subset)
finally show ?thesis
- using None
- by auto
+ using None by auto
next
case (Some v)
- with "2.hyps"
- have "fst ` set (compose xs ys) \<subseteq> fst ` set xs"
+ with "2.hyps" have "fst ` set (compose xs ys) \<subseteq> fst ` set xs"
by simp
with Some show ?thesis
by auto
@@ -726,10 +724,10 @@
subsection \<open>\<open>map_entry\<close>\<close>
qualified fun map_entry :: "'key \<Rightarrow> ('val \<Rightarrow> 'val) \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
-where
- "map_entry k f [] = []"
-| "map_entry k f (p # ps) =
- (if fst p = k then (k, f (snd p)) # ps else p # map_entry k f ps)"
+ where
+ "map_entry k f [] = []"
+ | "map_entry k f (p # ps) =
+ (if fst p = k then (k, f (snd p)) # ps else p # map_entry k f ps)"
lemma map_of_map_entry:
"map_of (map_entry k f xs) =
@@ -748,10 +746,10 @@
subsection \<open>\<open>map_default\<close>\<close>
fun map_default :: "'key \<Rightarrow> 'val \<Rightarrow> ('val \<Rightarrow> 'val) \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
-where
- "map_default k v f [] = [(k, v)]"
-| "map_default k v f (p # ps) =
- (if fst p = k then (k, f (snd p)) # ps else p # map_default k v f ps)"
+ where
+ "map_default k v f [] = [(k, v)]"
+ | "map_default k v f (p # ps) =
+ (if fst p = k then (k, f (snd p)) # ps else p # map_default k v f ps)"
lemma map_of_map_default:
"map_of (map_default k v f xs) =
--- a/src/HOL/Library/AList_Mapping.thy Tue Jul 12 13:55:35 2016 +0200
+++ b/src/HOL/Library/AList_Mapping.thy Tue Jul 12 16:04:19 2016 +0200
@@ -1,51 +1,43 @@
-(* Title: HOL/Library/AList_Mapping.thy
- Author: Florian Haftmann, TU Muenchen
+(* Title: HOL/Library/AList_Mapping.thy
+ Author: Florian Haftmann, TU Muenchen
*)
section \<open>Implementation of mappings with Association Lists\<close>
theory AList_Mapping
-imports AList Mapping
+ imports AList Mapping
begin
lift_definition Mapping :: "('a \<times> 'b) list \<Rightarrow> ('a, 'b) mapping" is map_of .
code_datatype Mapping
-lemma lookup_Mapping [simp, code]:
- "Mapping.lookup (Mapping xs) = map_of xs"
+lemma lookup_Mapping [simp, code]: "Mapping.lookup (Mapping xs) = map_of xs"
by transfer rule
-lemma keys_Mapping [simp, code]:
- "Mapping.keys (Mapping xs) = set (map fst xs)"
+lemma keys_Mapping [simp, code]: "Mapping.keys (Mapping xs) = set (map fst xs)"
by transfer (simp add: dom_map_of_conv_image_fst)
-lemma empty_Mapping [code]:
- "Mapping.empty = Mapping []"
+lemma empty_Mapping [code]: "Mapping.empty = Mapping []"
by transfer simp
-lemma is_empty_Mapping [code]:
- "Mapping.is_empty (Mapping xs) \<longleftrightarrow> List.null xs"
+lemma is_empty_Mapping [code]: "Mapping.is_empty (Mapping xs) \<longleftrightarrow> List.null xs"
by (case_tac xs) (simp_all add: is_empty_def null_def)
-lemma update_Mapping [code]:
- "Mapping.update k v (Mapping xs) = Mapping (AList.update k v xs)"
+lemma update_Mapping [code]: "Mapping.update k v (Mapping xs) = Mapping (AList.update k v xs)"
by transfer (simp add: update_conv')
-lemma delete_Mapping [code]:
- "Mapping.delete k (Mapping xs) = Mapping (AList.delete k xs)"
+lemma delete_Mapping [code]: "Mapping.delete k (Mapping xs) = Mapping (AList.delete k xs)"
by transfer (simp add: delete_conv')
lemma ordered_keys_Mapping [code]:
"Mapping.ordered_keys (Mapping xs) = sort (remdups (map fst xs))"
by (simp only: ordered_keys_def keys_Mapping sorted_list_of_set_sort_remdups) simp
-lemma size_Mapping [code]:
- "Mapping.size (Mapping xs) = length (remdups (map fst xs))"
+lemma size_Mapping [code]: "Mapping.size (Mapping xs) = length (remdups (map fst xs))"
by (simp add: size_def length_remdups_card_conv dom_map_of_conv_image_fst)
-lemma tabulate_Mapping [code]:
- "Mapping.tabulate ks f = Mapping (map (\<lambda>k. (k, f k)) ks)"
+lemma tabulate_Mapping [code]: "Mapping.tabulate ks f = Mapping (map (\<lambda>k. (k, f k)) ks)"
by transfer (simp add: map_of_map_restrict)
lemma bulkload_Mapping [code]:
@@ -55,63 +47,69 @@
lemma equal_Mapping [code]:
"HOL.equal (Mapping xs) (Mapping ys) \<longleftrightarrow>
(let ks = map fst xs; ls = map fst ys
- in (\<forall>l\<in>set ls. l \<in> set ks) \<and> (\<forall>k\<in>set ks. k \<in> set ls \<and> map_of xs k = map_of ys k))"
+ in (\<forall>l\<in>set ls. l \<in> set ks) \<and> (\<forall>k\<in>set ks. k \<in> set ls \<and> map_of xs k = map_of ys k))"
proof -
- have aux: "\<And>a b xs. (a, b) \<in> set xs \<Longrightarrow> a \<in> fst ` set xs"
+ have *: "(a, b) \<in> set xs \<Longrightarrow> a \<in> fst ` set xs" for a b xs
by (auto simp add: image_def intro!: bexI)
show ?thesis apply transfer
- by (auto intro!: map_of_eqI) (auto dest!: map_of_eq_dom intro: aux)
+ by (auto intro!: map_of_eqI) (auto dest!: map_of_eq_dom intro: *)
qed
lemma map_values_Mapping [code]:
- fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> 'b" and xs :: "('c \<times> 'a) list"
- shows "Mapping.map_values f (Mapping xs) = Mapping (map (\<lambda>(x,y). (x, f x y)) xs)"
-proof (transfer, rule ext, goal_cases)
- case (1 f xs x)
- thus ?case by (induction xs) auto
-qed
+ "Mapping.map_values f (Mapping xs) = Mapping (map (\<lambda>(x,y). (x, f x y)) xs)"
+ for f :: "'c \<Rightarrow> 'a \<Rightarrow> 'b" and xs :: "('c \<times> 'a) list"
+ apply transfer
+ apply (rule ext)
+ subgoal for f xs x by (induct xs) auto
+ done
-lemma combine_with_key_code [code]:
+lemma combine_with_key_code [code]:
"Mapping.combine_with_key f (Mapping xs) (Mapping ys) =
- Mapping.tabulate (remdups (map fst xs @ map fst ys))
+ Mapping.tabulate (remdups (map fst xs @ map fst ys))
(\<lambda>x. the (combine_options (f x) (map_of xs x) (map_of ys x)))"
-proof (transfer, rule ext, rule sym, goal_cases)
- case (1 f xs ys x)
- show ?case
- by (cases "map_of xs x"; cases "map_of ys x"; simp)
- (force simp: map_of_eq_None_iff combine_options_def option.the_def o_def image_iff
- dest: map_of_SomeD split: option.splits)+
-qed
+ apply transfer
+ apply (rule ext)
+ apply (rule sym)
+ subgoal for f xs ys x
+ apply (cases "map_of xs x"; cases "map_of ys x"; simp)
+ apply (force simp: map_of_eq_None_iff combine_options_def option.the_def o_def image_iff
+ dest: map_of_SomeD split: option.splits)+
+ done
+ done
-lemma combine_code [code]:
+lemma combine_code [code]:
"Mapping.combine f (Mapping xs) (Mapping ys) =
- Mapping.tabulate (remdups (map fst xs @ map fst ys))
+ Mapping.tabulate (remdups (map fst xs @ map fst ys))
(\<lambda>x. the (combine_options f (map_of xs x) (map_of ys x)))"
-proof (transfer, rule ext, rule sym, goal_cases)
- case (1 f xs ys x)
- show ?case
- by (cases "map_of xs x"; cases "map_of ys x"; simp)
- (force simp: map_of_eq_None_iff combine_options_def option.the_def o_def image_iff
- dest: map_of_SomeD split: option.splits)+
-qed
+ apply transfer
+ apply (rule ext)
+ apply (rule sym)
+ subgoal for f xs ys x
+ apply (cases "map_of xs x"; cases "map_of ys x"; simp)
+ apply (force simp: map_of_eq_None_iff combine_options_def option.the_def o_def image_iff
+ dest: map_of_SomeD split: option.splits)+
+ done
+ done
-(* TODO: Move? *)
-lemma map_of_filter_distinct:
+lemma map_of_filter_distinct: (* TODO: move? *)
assumes "distinct (map fst xs)"
- shows "map_of (filter P xs) x =
- (case map_of xs x of None \<Rightarrow> None | Some y \<Rightarrow> if P (x,y) then Some y else None)"
+ shows "map_of (filter P xs) x =
+ (case map_of xs x of
+ None \<Rightarrow> None
+ | Some y \<Rightarrow> if P (x,y) then Some y else None)"
using assms
by (auto simp: map_of_eq_None_iff filter_map distinct_map_filter dest: map_of_SomeD
- simp del: map_of_eq_Some_iff intro!: map_of_is_SomeI split: option.splits)
-(* END TODO *)
-
+ simp del: map_of_eq_Some_iff intro!: map_of_is_SomeI split: option.splits)
+
lemma filter_Mapping [code]:
"Mapping.filter P (Mapping xs) = Mapping (filter (\<lambda>(k,v). P k v) (AList.clearjunk xs))"
- by (transfer, rule ext)
- (subst map_of_filter_distinct, simp_all add: map_of_clearjunk split: option.split)
+ apply transfer
+ apply (rule ext)
+ apply (subst map_of_filter_distinct)
+ apply (simp_all add: map_of_clearjunk split: option.split)
+ done
-lemma [code nbe]:
- "HOL.equal (x :: ('a, 'b) mapping) x \<longleftrightarrow> True"
+lemma [code nbe]: "HOL.equal (x :: ('a, 'b) mapping) x \<longleftrightarrow> True"
by (fact equal_refl)
end
--- a/src/HOL/Library/BigO.thy Tue Jul 12 13:55:35 2016 +0200
+++ b/src/HOL/Library/BigO.thy Tue Jul 12 16:04:19 2016 +0200
@@ -5,7 +5,7 @@
section \<open>Big O notation\<close>
theory BigO
-imports Complex_Main Function_Algebras Set_Algebras
+ imports Complex_Main Function_Algebras Set_Algebras
begin
text \<open>
--- a/src/HOL/Library/Bit.thy Tue Jul 12 13:55:35 2016 +0200
+++ b/src/HOL/Library/Bit.thy Tue Jul 12 16:04:19 2016 +0200
@@ -11,17 +11,14 @@
subsection \<open>Bits as a datatype\<close>
typedef bit = "UNIV :: bool set"
- morphisms set Bit
- ..
+ morphisms set Bit ..
instantiation bit :: "{zero, one}"
begin
-definition zero_bit_def:
- "0 = Bit False"
+definition zero_bit_def: "0 = Bit False"
-definition one_bit_def:
- "1 = Bit True"
+definition one_bit_def: "1 = Bit True"
instance ..
@@ -29,8 +26,9 @@
old_rep_datatype "0::bit" "1::bit"
proof -
- fix P and x :: bit
- assume "P (0::bit)" and "P (1::bit)"
+ fix P :: "bit \<Rightarrow> bool"
+ fix x :: bit
+ assume "P 0" and "P 1"
then have "\<forall>b. P (Bit b)"
unfolding zero_bit_def one_bit_def
by (simp add: all_bool_eq)
@@ -42,21 +40,17 @@
by (simp add: Bit_inject)
qed
-lemma Bit_set_eq [simp]:
- "Bit (set b) = b"
+lemma Bit_set_eq [simp]: "Bit (set b) = b"
by (fact set_inverse)
-
-lemma set_Bit_eq [simp]:
- "set (Bit P) = P"
+
+lemma set_Bit_eq [simp]: "set (Bit P) = P"
by (rule Bit_inverse) rule
-lemma bit_eq_iff:
- "x = y \<longleftrightarrow> (set x \<longleftrightarrow> set y)"
+lemma bit_eq_iff: "x = y \<longleftrightarrow> (set x \<longleftrightarrow> set y)"
by (auto simp add: set_inject)
-lemma Bit_inject [simp]:
- "Bit P = Bit Q \<longleftrightarrow> (P \<longleftrightarrow> Q)"
- by (auto simp add: Bit_inject)
+lemma Bit_inject [simp]: "Bit P = Bit Q \<longleftrightarrow> (P \<longleftrightarrow> Q)"
+ by (auto simp add: Bit_inject)
lemma set [iff]:
"\<not> set 0"
@@ -68,8 +62,7 @@
"set 1 \<longleftrightarrow> True"
by simp_all
-lemma set_iff:
- "set b \<longleftrightarrow> b = 1"
+lemma set_iff: "set b \<longleftrightarrow> b = 1"
by (cases b) simp_all
lemma bit_eq_iff_set:
@@ -82,42 +75,34 @@
"Bit True = 1"
by (simp_all add: zero_bit_def one_bit_def)
-lemma bit_not_0_iff [iff]:
- "(x::bit) \<noteq> 0 \<longleftrightarrow> x = 1"
+lemma bit_not_0_iff [iff]: "x \<noteq> 0 \<longleftrightarrow> x = 1" for x :: bit
by (simp add: bit_eq_iff)
-lemma bit_not_1_iff [iff]:
- "(x::bit) \<noteq> 1 \<longleftrightarrow> x = 0"
+lemma bit_not_1_iff [iff]: "x \<noteq> 1 \<longleftrightarrow> x = 0" for x :: bit
by (simp add: bit_eq_iff)
lemma [code]:
"HOL.equal 0 b \<longleftrightarrow> \<not> set b"
"HOL.equal 1 b \<longleftrightarrow> set b"
- by (simp_all add: equal set_iff)
+ by (simp_all add: equal set_iff)
-
+
subsection \<open>Type @{typ bit} forms a field\<close>
instantiation bit :: field
begin
-definition plus_bit_def:
- "x + y = case_bit y (case_bit 1 0 y) x"
+definition plus_bit_def: "x + y = case_bit y (case_bit 1 0 y) x"
-definition times_bit_def:
- "x * y = case_bit 0 y x"
+definition times_bit_def: "x * y = case_bit 0 y x"
-definition uminus_bit_def [simp]:
- "- x = (x :: bit)"
+definition uminus_bit_def [simp]: "- x = x" for x :: bit
-definition minus_bit_def [simp]:
- "x - y = (x + y :: bit)"
+definition minus_bit_def [simp]: "x - y = x + y" for x y :: bit
-definition inverse_bit_def [simp]:
- "inverse x = (x :: bit)"
+definition inverse_bit_def [simp]: "inverse x = x" for x :: bit
-definition divide_bit_def [simp]:
- "x div y = (x * y :: bit)"
+definition divide_bit_def [simp]: "x div y = x * y" for x y :: bit
lemmas field_bit_defs =
plus_bit_def times_bit_def minus_bit_def uminus_bit_def
@@ -128,18 +113,18 @@
end
-lemma bit_add_self: "x + x = (0 :: bit)"
+lemma bit_add_self: "x + x = 0" for x :: bit
unfolding plus_bit_def by (simp split: bit.split)
-lemma bit_mult_eq_1_iff [simp]: "x * y = (1 :: bit) \<longleftrightarrow> x = 1 \<and> y = 1"
+lemma bit_mult_eq_1_iff [simp]: "x * y = 1 \<longleftrightarrow> x = 1 \<and> y = 1" for x y :: bit
unfolding times_bit_def by (simp split: bit.split)
text \<open>Not sure whether the next two should be simp rules.\<close>
-lemma bit_add_eq_0_iff: "x + y = (0 :: bit) \<longleftrightarrow> x = y"
+lemma bit_add_eq_0_iff: "x + y = 0 \<longleftrightarrow> x = y" for x y :: bit
unfolding plus_bit_def by (simp split: bit.split)
-lemma bit_add_eq_1_iff: "x + y = (1 :: bit) \<longleftrightarrow> x \<noteq> y"
+lemma bit_add_eq_1_iff: "x + y = 1 \<longleftrightarrow> x \<noteq> y" for x y :: bit
unfolding plus_bit_def by (simp split: bit.split)
@@ -166,37 +151,23 @@
begin
definition of_bit :: "bit \<Rightarrow> 'a"
-where
- "of_bit b = case_bit 0 1 b"
+ where "of_bit b = case_bit 0 1 b"
lemma of_bit_eq [simp, code]:
"of_bit 0 = 0"
"of_bit 1 = 1"
by (simp_all add: of_bit_def)
-lemma of_bit_eq_iff:
- "of_bit x = of_bit y \<longleftrightarrow> x = y"
- by (cases x) (cases y, simp_all)+
-
-end
-
-context semiring_1
-begin
-
-lemma of_nat_of_bit_eq:
- "of_nat (of_bit b) = of_bit b"
- by (cases b) simp_all
+lemma of_bit_eq_iff: "of_bit x = of_bit y \<longleftrightarrow> x = y"
+ by (cases x) (cases y; simp)+
end
-context ring_1
-begin
-
-lemma of_int_of_bit_eq:
- "of_int (of_bit b) = of_bit b"
+lemma (in semiring_1) of_nat_of_bit_eq: "of_nat (of_bit b) = of_bit b"
by (cases b) simp_all
-end
+lemma (in ring_1) of_int_of_bit_eq: "of_int (of_bit b) = of_bit b"
+ by (cases b) simp_all
hide_const (open) set
--- a/src/HOL/Library/Boolean_Algebra.thy Tue Jul 12 13:55:35 2016 +0200
+++ b/src/HOL/Library/Boolean_Algebra.thy Tue Jul 12 16:04:19 2016 +0200
@@ -5,7 +5,7 @@
section \<open>Boolean Algebras\<close>
theory Boolean_Algebra
-imports Main
+ imports Main
begin
locale boolean =
@@ -40,18 +40,18 @@
lemmas disj_ac = disj.assoc disj.commute disj.left_commute
lemma dual: "boolean disj conj compl one zero"
-apply (rule boolean.intro)
-apply (rule disj_assoc)
-apply (rule conj_assoc)
-apply (rule disj_commute)
-apply (rule conj_commute)
-apply (rule disj_conj_distrib)
-apply (rule conj_disj_distrib)
-apply (rule disj_zero_right)
-apply (rule conj_one_right)
-apply (rule disj_cancel_right)
-apply (rule conj_cancel_right)
-done
+ apply (rule boolean.intro)
+ apply (rule disj_assoc)
+ apply (rule conj_assoc)
+ apply (rule disj_commute)
+ apply (rule conj_commute)
+ apply (rule disj_conj_distrib)
+ apply (rule conj_disj_distrib)
+ apply (rule disj_zero_right)
+ apply (rule conj_one_right)
+ apply (rule disj_cancel_right)
+ apply (rule conj_cancel_right)
+ done
subsection \<open>Complement\<close>
@@ -63,99 +63,111 @@
assumes 4: "a \<squnion> y = \<one>"
shows "x = y"
proof -
- have "(a \<sqinter> x) \<squnion> (x \<sqinter> y) = (a \<sqinter> y) \<squnion> (x \<sqinter> y)" using 1 3 by simp
- hence "(x \<sqinter> a) \<squnion> (x \<sqinter> y) = (y \<sqinter> a) \<squnion> (y \<sqinter> x)" using conj_commute by simp
- hence "x \<sqinter> (a \<squnion> y) = y \<sqinter> (a \<squnion> x)" using conj_disj_distrib by simp
- hence "x \<sqinter> \<one> = y \<sqinter> \<one>" using 2 4 by simp
- thus "x = y" using conj_one_right by simp
+ have "(a \<sqinter> x) \<squnion> (x \<sqinter> y) = (a \<sqinter> y) \<squnion> (x \<sqinter> y)"
+ using 1 3 by simp
+ then have "(x \<sqinter> a) \<squnion> (x \<sqinter> y) = (y \<sqinter> a) \<squnion> (y \<sqinter> x)"
+ using conj_commute by simp
+ then have "x \<sqinter> (a \<squnion> y) = y \<sqinter> (a \<squnion> x)"
+ using conj_disj_distrib by simp
+ then have "x \<sqinter> \<one> = y \<sqinter> \<one>"
+ using 2 4 by simp
+ then show "x = y"
+ using conj_one_right by simp
qed
-lemma compl_unique: "\<lbrakk>x \<sqinter> y = \<zero>; x \<squnion> y = \<one>\<rbrakk> \<Longrightarrow> \<sim> x = y"
-by (rule complement_unique [OF conj_cancel_right disj_cancel_right])
+lemma compl_unique: "x \<sqinter> y = \<zero> \<Longrightarrow> x \<squnion> y = \<one> \<Longrightarrow> \<sim> x = y"
+ by (rule complement_unique [OF conj_cancel_right disj_cancel_right])
lemma double_compl [simp]: "\<sim> (\<sim> x) = x"
proof (rule compl_unique)
- from conj_cancel_right show "\<sim> x \<sqinter> x = \<zero>" by (simp only: conj_commute)
- from disj_cancel_right show "\<sim> x \<squnion> x = \<one>" by (simp only: disj_commute)
+ from conj_cancel_right show "\<sim> x \<sqinter> x = \<zero>"
+ by (simp only: conj_commute)
+ from disj_cancel_right show "\<sim> x \<squnion> x = \<one>"
+ by (simp only: disj_commute)
qed
-lemma compl_eq_compl_iff [simp]: "(\<sim> x = \<sim> y) = (x = y)"
-by (rule inj_eq [OF inj_on_inverseI], rule double_compl)
+lemma compl_eq_compl_iff [simp]: "\<sim> x = \<sim> y \<longleftrightarrow> x = y"
+ by (rule inj_eq [OF inj_on_inverseI]) (rule double_compl)
subsection \<open>Conjunction\<close>
lemma conj_absorb [simp]: "x \<sqinter> x = x"
proof -
- have "x \<sqinter> x = (x \<sqinter> x) \<squnion> \<zero>" using disj_zero_right by simp
- also have "... = (x \<sqinter> x) \<squnion> (x \<sqinter> \<sim> x)" using conj_cancel_right by simp
- also have "... = x \<sqinter> (x \<squnion> \<sim> x)" using conj_disj_distrib by (simp only:)
- also have "... = x \<sqinter> \<one>" using disj_cancel_right by simp
- also have "... = x" using conj_one_right by simp
+ have "x \<sqinter> x = (x \<sqinter> x) \<squnion> \<zero>"
+ using disj_zero_right by simp
+ also have "... = (x \<sqinter> x) \<squnion> (x \<sqinter> \<sim> x)"
+ using conj_cancel_right by simp
+ also have "... = x \<sqinter> (x \<squnion> \<sim> x)"
+ using conj_disj_distrib by (simp only:)
+ also have "... = x \<sqinter> \<one>"
+ using disj_cancel_right by simp
+ also have "... = x"
+ using conj_one_right by simp
finally show ?thesis .
qed
lemma conj_zero_right [simp]: "x \<sqinter> \<zero> = \<zero>"
proof -
- have "x \<sqinter> \<zero> = x \<sqinter> (x \<sqinter> \<sim> x)" using conj_cancel_right by simp
- also have "... = (x \<sqinter> x) \<sqinter> \<sim> x" using conj_assoc by (simp only:)
- also have "... = x \<sqinter> \<sim> x" using conj_absorb by simp
- also have "... = \<zero>" using conj_cancel_right by simp
+ have "x \<sqinter> \<zero> = x \<sqinter> (x \<sqinter> \<sim> x)"
+ using conj_cancel_right by simp
+ also have "... = (x \<sqinter> x) \<sqinter> \<sim> x"
+ using conj_assoc by (simp only:)
+ also have "... = x \<sqinter> \<sim> x"
+ using conj_absorb by simp
+ also have "... = \<zero>"
+ using conj_cancel_right by simp
finally show ?thesis .
qed
lemma compl_one [simp]: "\<sim> \<one> = \<zero>"
-by (rule compl_unique [OF conj_zero_right disj_zero_right])
+ by (rule compl_unique [OF conj_zero_right disj_zero_right])
lemma conj_zero_left [simp]: "\<zero> \<sqinter> x = \<zero>"
-by (subst conj_commute) (rule conj_zero_right)
+ by (subst conj_commute) (rule conj_zero_right)
lemma conj_one_left [simp]: "\<one> \<sqinter> x = x"
-by (subst conj_commute) (rule conj_one_right)
+ by (subst conj_commute) (rule conj_one_right)
lemma conj_cancel_left [simp]: "\<sim> x \<sqinter> x = \<zero>"
-by (subst conj_commute) (rule conj_cancel_right)
+ by (subst conj_commute) (rule conj_cancel_right)
lemma conj_left_absorb [simp]: "x \<sqinter> (x \<sqinter> y) = x \<sqinter> y"
-by (simp only: conj_assoc [symmetric] conj_absorb)
+ by (simp only: conj_assoc [symmetric] conj_absorb)
-lemma conj_disj_distrib2:
- "(y \<squnion> z) \<sqinter> x = (y \<sqinter> x) \<squnion> (z \<sqinter> x)"
-by (simp only: conj_commute conj_disj_distrib)
+lemma conj_disj_distrib2: "(y \<squnion> z) \<sqinter> x = (y \<sqinter> x) \<squnion> (z \<sqinter> x)"
+ by (simp only: conj_commute conj_disj_distrib)
-lemmas conj_disj_distribs =
- conj_disj_distrib conj_disj_distrib2
+lemmas conj_disj_distribs = conj_disj_distrib conj_disj_distrib2
subsection \<open>Disjunction\<close>
lemma disj_absorb [simp]: "x \<squnion> x = x"
-by (rule boolean.conj_absorb [OF dual])
+ by (rule boolean.conj_absorb [OF dual])
lemma disj_one_right [simp]: "x \<squnion> \<one> = \<one>"
-by (rule boolean.conj_zero_right [OF dual])
+ by (rule boolean.conj_zero_right [OF dual])
lemma compl_zero [simp]: "\<sim> \<zero> = \<one>"
-by (rule boolean.compl_one [OF dual])
+ by (rule boolean.compl_one [OF dual])
lemma disj_zero_left [simp]: "\<zero> \<squnion> x = x"
-by (rule boolean.conj_one_left [OF dual])
+ by (rule boolean.conj_one_left [OF dual])
lemma disj_one_left [simp]: "\<one> \<squnion> x = \<one>"
-by (rule boolean.conj_zero_left [OF dual])
+ by (rule boolean.conj_zero_left [OF dual])
lemma disj_cancel_left [simp]: "\<sim> x \<squnion> x = \<one>"
-by (rule boolean.conj_cancel_left [OF dual])
+ by (rule boolean.conj_cancel_left [OF dual])
lemma disj_left_absorb [simp]: "x \<squnion> (x \<squnion> y) = x \<squnion> y"
-by (rule boolean.conj_left_absorb [OF dual])
+ by (rule boolean.conj_left_absorb [OF dual])
-lemma disj_conj_distrib2:
- "(y \<sqinter> z) \<squnion> x = (y \<squnion> x) \<sqinter> (z \<squnion> x)"
-by (rule boolean.conj_disj_distrib2 [OF dual])
+lemma disj_conj_distrib2: "(y \<sqinter> z) \<squnion> x = (y \<squnion> x) \<sqinter> (z \<squnion> x)"
+ by (rule boolean.conj_disj_distrib2 [OF dual])
-lemmas disj_conj_distribs =
- disj_conj_distrib disj_conj_distrib2
+lemmas disj_conj_distribs = disj_conj_distrib disj_conj_distrib2
subsection \<open>De Morgan's Laws\<close>
@@ -178,7 +190,7 @@
qed
lemma de_Morgan_disj [simp]: "\<sim> (x \<squnion> y) = \<sim> x \<sqinter> \<sim> y"
-by (rule boolean.de_Morgan_conj [OF dual])
+ by (rule boolean.de_Morgan_conj [OF dual])
end
@@ -198,7 +210,7 @@
have "?t \<squnion> (z \<sqinter> x \<sqinter> \<sim> x) \<squnion> (z \<sqinter> y \<sqinter> \<sim> y) =
?t \<squnion> (x \<sqinter> y \<sqinter> \<sim> y) \<squnion> (x \<sqinter> z \<sqinter> \<sim> z)"
by (simp only: conj_cancel_right conj_zero_right)
- thus "(x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)"
+ then show "(x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)"
apply (simp only: xor_def de_Morgan_disj de_Morgan_conj double_compl)
apply (simp only: conj_disj_distribs conj_ac disj_ac)
done
@@ -212,57 +224,55 @@
lemmas xor_ac = xor.assoc xor.commute xor.left_commute
-lemma xor_def2:
- "x \<oplus> y = (x \<squnion> y) \<sqinter> (\<sim> x \<squnion> \<sim> y)"
-by (simp only: xor_def conj_disj_distribs
- disj_ac conj_ac conj_cancel_right disj_zero_left)
+lemma xor_def2: "x \<oplus> y = (x \<squnion> y) \<sqinter> (\<sim> x \<squnion> \<sim> y)"
+ by (simp only: xor_def conj_disj_distribs disj_ac conj_ac conj_cancel_right disj_zero_left)
lemma xor_zero_right [simp]: "x \<oplus> \<zero> = x"
-by (simp only: xor_def compl_zero conj_one_right conj_zero_right disj_zero_right)
+ by (simp only: xor_def compl_zero conj_one_right conj_zero_right disj_zero_right)
lemma xor_zero_left [simp]: "\<zero> \<oplus> x = x"
-by (subst xor_commute) (rule xor_zero_right)
+ by (subst xor_commute) (rule xor_zero_right)
lemma xor_one_right [simp]: "x \<oplus> \<one> = \<sim> x"
-by (simp only: xor_def compl_one conj_zero_right conj_one_right disj_zero_left)
+ by (simp only: xor_def compl_one conj_zero_right conj_one_right disj_zero_left)
lemma xor_one_left [simp]: "\<one> \<oplus> x = \<sim> x"
-by (subst xor_commute) (rule xor_one_right)
+ by (subst xor_commute) (rule xor_one_right)
lemma xor_self [simp]: "x \<oplus> x = \<zero>"
-by (simp only: xor_def conj_cancel_right conj_cancel_left disj_zero_right)
+ by (simp only: xor_def conj_cancel_right conj_cancel_left disj_zero_right)
lemma xor_left_self [simp]: "x \<oplus> (x \<oplus> y) = y"
-by (simp only: xor_assoc [symmetric] xor_self xor_zero_left)
+ by (simp only: xor_assoc [symmetric] xor_self xor_zero_left)
lemma xor_compl_left [simp]: "\<sim> x \<oplus> y = \<sim> (x \<oplus> y)"
-apply (simp only: xor_def de_Morgan_disj de_Morgan_conj double_compl)
-apply (simp only: conj_disj_distribs)
-apply (simp only: conj_cancel_right conj_cancel_left)
-apply (simp only: disj_zero_left disj_zero_right)
-apply (simp only: disj_ac conj_ac)
-done
+ apply (simp only: xor_def de_Morgan_disj de_Morgan_conj double_compl)
+ apply (simp only: conj_disj_distribs)
+ apply (simp only: conj_cancel_right conj_cancel_left)
+ apply (simp only: disj_zero_left disj_zero_right)
+ apply (simp only: disj_ac conj_ac)
+ done
lemma xor_compl_right [simp]: "x \<oplus> \<sim> y = \<sim> (x \<oplus> y)"
-apply (simp only: xor_def de_Morgan_disj de_Morgan_conj double_compl)
-apply (simp only: conj_disj_distribs)
-apply (simp only: conj_cancel_right conj_cancel_left)
-apply (simp only: disj_zero_left disj_zero_right)
-apply (simp only: disj_ac conj_ac)
-done
+ apply (simp only: xor_def de_Morgan_disj de_Morgan_conj double_compl)
+ apply (simp only: conj_disj_distribs)
+ apply (simp only: conj_cancel_right conj_cancel_left)
+ apply (simp only: disj_zero_left disj_zero_right)
+ apply (simp only: disj_ac conj_ac)
+ done
lemma xor_cancel_right: "x \<oplus> \<sim> x = \<one>"
-by (simp only: xor_compl_right xor_self compl_zero)
+ by (simp only: xor_compl_right xor_self compl_zero)
lemma xor_cancel_left: "\<sim> x \<oplus> x = \<one>"
-by (simp only: xor_compl_left xor_self compl_zero)
+ by (simp only: xor_compl_left xor_self compl_zero)
lemma conj_xor_distrib: "x \<sqinter> (y \<oplus> z) = (x \<sqinter> y) \<oplus> (x \<sqinter> z)"
proof -
- have "(x \<sqinter> y \<sqinter> \<sim> z) \<squnion> (x \<sqinter> \<sim> y \<sqinter> z) =
+ have *: "(x \<sqinter> y \<sqinter> \<sim> z) \<squnion> (x \<sqinter> \<sim> y \<sqinter> z) =
(y \<sqinter> x \<sqinter> \<sim> x) \<squnion> (z \<sqinter> x \<sqinter> \<sim> x) \<squnion> (x \<sqinter> y \<sqinter> \<sim> z) \<squnion> (x \<sqinter> \<sim> y \<sqinter> z)"
by (simp only: conj_cancel_right conj_zero_right disj_zero_left)
- thus "x \<sqinter> (y \<oplus> z) = (x \<sqinter> y) \<oplus> (x \<sqinter> z)"
+ then show "x \<sqinter> (y \<oplus> z) = (x \<sqinter> y) \<oplus> (x \<sqinter> z)"
by (simp (no_asm_use) only:
xor_def de_Morgan_disj de_Morgan_conj double_compl
conj_disj_distribs conj_ac disj_ac)
@@ -272,7 +282,7 @@
proof -
have "x \<sqinter> (y \<oplus> z) = (x \<sqinter> y) \<oplus> (x \<sqinter> z)"
by (rule conj_xor_distrib)
- thus "(y \<oplus> z) \<sqinter> x = (y \<sqinter> x) \<oplus> (z \<sqinter> x)"
+ then show "(y \<oplus> z) \<sqinter> x = (y \<sqinter> x) \<oplus> (z \<sqinter> x)"
by (simp only: conj_commute)
qed
--- a/src/HOL/Library/Char_ord.thy Tue Jul 12 13:55:35 2016 +0200
+++ b/src/HOL/Library/Char_ord.thy Tue Jul 12 16:04:19 2016 +0200
@@ -5,7 +5,7 @@
section \<open>Order on characters\<close>
theory Char_ord
-imports Main
+ imports Main
begin
instantiation char :: linorder
@@ -20,17 +20,19 @@
end
lemma less_eq_char_simps:
- "(0 :: char) \<le> c"
+ "0 \<le> c"
"Char k \<le> 0 \<longleftrightarrow> numeral k mod 256 = (0 :: nat)"
"Char k \<le> Char l \<longleftrightarrow> numeral k mod 256 \<le> (numeral l mod 256 :: nat)"
+ for c :: char
by (simp_all add: Char_def less_eq_char_def)
lemma less_char_simps:
- "\<not> c < (0 :: char)"
+ "\<not> c < 0"
"0 < Char k \<longleftrightarrow> (0 :: nat) < numeral k mod 256"
"Char k < Char l \<longleftrightarrow> numeral k mod 256 < (numeral l mod 256 :: nat)"
+ for c :: char
by (simp_all add: Char_def less_char_def)
-
+
instantiation char :: distrib_lattice
begin
@@ -45,28 +47,34 @@
instantiation String.literal :: linorder
begin
-context includes literal.lifting begin
-lift_definition less_literal :: "String.literal \<Rightarrow> String.literal \<Rightarrow> bool" is "ord.lexordp op <" .
-lift_definition less_eq_literal :: "String.literal \<Rightarrow> String.literal \<Rightarrow> bool" is "ord.lexordp_eq op <" .
+context includes literal.lifting
+begin
+
+lift_definition less_literal :: "String.literal \<Rightarrow> String.literal \<Rightarrow> bool"
+ is "ord.lexordp op <" .
+
+lift_definition less_eq_literal :: "String.literal \<Rightarrow> String.literal \<Rightarrow> bool"
+ is "ord.lexordp_eq op <" .
instance
proof -
interpret linorder "ord.lexordp_eq op <" "ord.lexordp op < :: string \<Rightarrow> string \<Rightarrow> bool"
- by(rule linorder.lexordp_linorder[where less_eq="op \<le>"])(unfold_locales)
+ by (rule linorder.lexordp_linorder[where less_eq="op \<le>"]) unfold_locales
show "PROP ?thesis"
- by(intro_classes)(transfer, simp add: less_le_not_le linear)+
+ by intro_classes (transfer, simp add: less_le_not_le linear)+
qed
end
+
end
-lemma less_literal_code [code]:
+lemma less_literal_code [code]:
"op < = (\<lambda>xs ys. ord.lexordp op < (String.explode xs) (String.explode ys))"
-by(simp add: less_literal.rep_eq fun_eq_iff)
+ by (simp add: less_literal.rep_eq fun_eq_iff)
lemma less_eq_literal_code [code]:
"op \<le> = (\<lambda>xs ys. ord.lexordp_eq op < (String.explode xs) (String.explode ys))"
-by(simp add: less_eq_literal.rep_eq fun_eq_iff)
+ by (simp add: less_eq_literal.rep_eq fun_eq_iff)
lifting_update literal.lifting
lifting_forget literal.lifting
--- a/src/HOL/Library/Combine_PER.thy Tue Jul 12 13:55:35 2016 +0200
+++ b/src/HOL/Library/Combine_PER.thy Tue Jul 12 16:04:19 2016 +0200
@@ -1,47 +1,37 @@
-(* Author: Florian Haftmann, TU Muenchen *)
+(* Author: Florian Haftmann, TU Muenchen *)
section \<open>A combinator to build partial equivalence relations from a predicate and an equivalence relation\<close>
theory Combine_PER
-imports Main "~~/src/HOL/Library/Lattice_Syntax"
+ imports Main "~~/src/HOL/Library/Lattice_Syntax"
begin
definition combine_per :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
-where
- "combine_per P R = (\<lambda>x y. P x \<and> P y) \<sqinter> R"
+ where "combine_per P R = (\<lambda>x y. P x \<and> P y) \<sqinter> R"
lemma combine_per_simp [simp]:
- fixes R (infixl "\<approx>" 50)
- shows "combine_per P R x y \<longleftrightarrow> P x \<and> P y \<and> x \<approx> y"
+ "combine_per P R x y \<longleftrightarrow> P x \<and> P y \<and> x \<approx> y" for R (infixl "\<approx>" 50)
by (simp add: combine_per_def)
-lemma combine_per_top [simp]:
- "combine_per \<top> R = R"
+lemma combine_per_top [simp]: "combine_per \<top> R = R"
by (simp add: fun_eq_iff)
-lemma combine_per_eq [simp]:
- "combine_per P HOL.eq = HOL.eq \<sqinter> (\<lambda>x y. P x)"
+lemma combine_per_eq [simp]: "combine_per P HOL.eq = HOL.eq \<sqinter> (\<lambda>x y. P x)"
by (auto simp add: fun_eq_iff)
-lemma symp_combine_per:
- "symp R \<Longrightarrow> symp (combine_per P R)"
+lemma symp_combine_per: "symp R \<Longrightarrow> symp (combine_per P R)"
by (auto simp add: symp_def sym_def combine_per_def)
-lemma transp_combine_per:
- "transp R \<Longrightarrow> transp (combine_per P R)"
+lemma transp_combine_per: "transp R \<Longrightarrow> transp (combine_per P R)"
by (auto simp add: transp_def trans_def combine_per_def)
-lemma combine_perI:
- fixes R (infixl "\<approx>" 50)
- shows "P x \<Longrightarrow> P y \<Longrightarrow> x \<approx> y \<Longrightarrow> combine_per P R x y"
+lemma combine_perI: "P x \<Longrightarrow> P y \<Longrightarrow> x \<approx> y \<Longrightarrow> combine_per P R x y" for R (infixl "\<approx>" 50)
by (simp add: combine_per_def)
-lemma symp_combine_per_symp:
- "symp R \<Longrightarrow> symp (combine_per P R)"
+lemma symp_combine_per_symp: "symp R \<Longrightarrow> symp (combine_per P R)"
by (auto intro!: sympI elim: sympE)
-lemma transp_combine_per_transp:
- "transp R \<Longrightarrow> transp (combine_per P R)"
+lemma transp_combine_per_transp: "transp R \<Longrightarrow> transp (combine_per P R)"
by (auto intro!: transpI elim: transpE)
lemma equivp_combine_per_part_equivp [intro?]:
@@ -50,8 +40,10 @@
shows "part_equivp (combine_per P R)"
proof -
from \<open>\<exists>x. P x\<close> obtain x where "P x" ..
- moreover from \<open>equivp R\<close> have "x \<approx> x" by (rule equivp_reflp)
- ultimately have "\<exists>x. P x \<and> x \<approx> x" by blast
+ moreover from \<open>equivp R\<close> have "x \<approx> x"
+ by (rule equivp_reflp)
+ ultimately have "\<exists>x. P x \<and> x \<approx> x"
+ by blast
with \<open>equivp R\<close> show ?thesis
by (auto intro!: part_equivpI symp_combine_per_symp transp_combine_per_transp
elim: equivpE)
--- a/src/HOL/Library/Mapping.thy Tue Jul 12 13:55:35 2016 +0200
+++ b/src/HOL/Library/Mapping.thy Tue Jul 12 16:04:19 2016 +0200
@@ -10,15 +10,13 @@
subsection \<open>Parametricity transfer rules\<close>
-lemma map_of_foldr: \<comment> \<open>FIXME move\<close>
- "map_of xs = foldr (\<lambda>(k, v) m. m(k \<mapsto> v)) xs Map.empty"
+lemma map_of_foldr: "map_of xs = foldr (\<lambda>(k, v) m. m(k \<mapsto> v)) xs Map.empty" (* FIXME move *)
using map_add_map_of_foldr [of Map.empty] by auto
context includes lifting_syntax
begin
-lemma empty_parametric:
- "(A ===> rel_option B) Map.empty Map.empty"
+lemma empty_parametric: "(A ===> rel_option B) Map.empty Map.empty"
by transfer_prover
lemma lookup_parametric: "((A ===> B) ===> A ===> B) (\<lambda>m k. m k) (\<lambda>m k. m k)"
@@ -32,7 +30,7 @@
lemma delete_parametric:
assumes [transfer_rule]: "bi_unique A"
- shows "(A ===> (A ===> rel_option B) ===> A ===> rel_option B)
+ shows "(A ===> (A ===> rel_option B) ===> A ===> rel_option B)
(\<lambda>k m. m(k := None)) (\<lambda>k m. m(k := None))"
by transfer_prover
@@ -42,7 +40,7 @@
lemma dom_parametric:
assumes [transfer_rule]: "bi_total A"
- shows "((A ===> rel_option B) ===> rel_set A) dom dom"
+ shows "((A ===> rel_option B) ===> rel_set A) dom dom"
unfolding dom_def [abs_def] Option.is_none_def [symmetric] by transfer_prover
lemma map_of_parametric [transfer_rule]:
@@ -52,51 +50,53 @@
lemma map_entry_parametric [transfer_rule]:
assumes [transfer_rule]: "bi_unique A"
- shows "(A ===> (B ===> B) ===> (A ===> rel_option B) ===> A ===> rel_option B)
+ shows "(A ===> (B ===> B) ===> (A ===> rel_option B) ===> A ===> rel_option B)
(\<lambda>k f m. (case m k of None \<Rightarrow> m
| Some v \<Rightarrow> m (k \<mapsto> (f v)))) (\<lambda>k f m. (case m k of None \<Rightarrow> m
| Some v \<Rightarrow> m (k \<mapsto> (f v))))"
by transfer_prover
-lemma tabulate_parametric:
+lemma tabulate_parametric:
assumes [transfer_rule]: "bi_unique A"
- shows "(list_all2 A ===> (A ===> B) ===> A ===> rel_option B)
+ shows "(list_all2 A ===> (A ===> B) ===> A ===> rel_option B)
(\<lambda>ks f. (map_of (map (\<lambda>k. (k, f k)) ks))) (\<lambda>ks f. (map_of (map (\<lambda>k. (k, f k)) ks)))"
by transfer_prover
-lemma bulkload_parametric:
- "(list_all2 A ===> HOL.eq ===> rel_option A)
- (\<lambda>xs k. if k < length xs then Some (xs ! k) else None) (\<lambda>xs k. if k < length xs then Some (xs ! k) else None)"
+lemma bulkload_parametric:
+ "(list_all2 A ===> HOL.eq ===> rel_option A)
+ (\<lambda>xs k. if k < length xs then Some (xs ! k) else None)
+ (\<lambda>xs k. if k < length xs then Some (xs ! k) else None)"
proof
fix xs ys
assume "list_all2 A xs ys"
- then show "(HOL.eq ===> rel_option A)
- (\<lambda>k. if k < length xs then Some (xs ! k) else None)
- (\<lambda>k. if k < length ys then Some (ys ! k) else None)"
+ then show
+ "(HOL.eq ===> rel_option A)
+ (\<lambda>k. if k < length xs then Some (xs ! k) else None)
+ (\<lambda>k. if k < length ys then Some (ys ! k) else None)"
apply induct
apply auto
unfolding rel_fun_def
- apply clarsimp
- apply (case_tac xa)
+ apply clarsimp
+ apply (case_tac xa)
apply (auto dest: list_all2_lengthD list_all2_nthD)
done
qed
-lemma map_parametric:
- "((A ===> B) ===> (C ===> D) ===> (B ===> rel_option C) ===> A ===> rel_option D)
+lemma map_parametric:
+ "((A ===> B) ===> (C ===> D) ===> (B ===> rel_option C) ===> A ===> rel_option D)
(\<lambda>f g m. (map_option g \<circ> m \<circ> f)) (\<lambda>f g m. (map_option g \<circ> m \<circ> f))"
by transfer_prover
-
-lemma combine_with_key_parametric:
- shows "((A ===> B ===> B ===> B) ===> (A ===> rel_option B) ===> (A ===> rel_option B) ===>
- (A ===> rel_option B)) (\<lambda>f m1 m2 x. combine_options (f x) (m1 x) (m2 x))
- (\<lambda>f m1 m2 x. combine_options (f x) (m1 x) (m2 x))"
+
+lemma combine_with_key_parametric:
+ "((A ===> B ===> B ===> B) ===> (A ===> rel_option B) ===> (A ===> rel_option B) ===>
+ (A ===> rel_option B)) (\<lambda>f m1 m2 x. combine_options (f x) (m1 x) (m2 x))
+ (\<lambda>f m1 m2 x. combine_options (f x) (m1 x) (m2 x))"
unfolding combine_options_def by transfer_prover
-
-lemma combine_parametric:
- shows "((B ===> B ===> B) ===> (A ===> rel_option B) ===> (A ===> rel_option B) ===>
- (A ===> rel_option B)) (\<lambda>f m1 m2 x. combine_options f (m1 x) (m2 x))
- (\<lambda>f m1 m2 x. combine_options f (m1 x) (m2 x))"
+
+lemma combine_parametric:
+ "((B ===> B ===> B) ===> (A ===> rel_option B) ===> (A ===> rel_option B) ===>
+ (A ===> rel_option B)) (\<lambda>f m1 m2 x. combine_options f (m1 x) (m2 x))
+ (\<lambda>f m1 m2 x. combine_options f (m1 x) (m2 x))"
unfolding combine_options_def by transfer_prover
end
@@ -105,8 +105,7 @@
subsection \<open>Type definition and primitive operations\<close>
typedef ('a, 'b) mapping = "UNIV :: ('a \<rightharpoonup> 'b) set"
- morphisms rep Mapping
- ..
+ morphisms rep Mapping ..
setup_lifting type_definition_mapping
@@ -119,7 +118,7 @@
definition "lookup_default d m k = (case Mapping.lookup m k of None \<Rightarrow> d | Some v \<Rightarrow> v)"
declare [[code drop: Mapping.lookup]]
-setup \<open>Code.add_eqn (Code.Equation, true) @{thm Mapping.lookup.abs_eq}\<close> \<comment> \<open>FIXME lifting\<close>
+setup \<open>Code.add_eqn (Code.Equation, true) @{thm Mapping.lookup.abs_eq}\<close> (* FIXME lifting *)
lift_definition update :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping"
is "\<lambda>k v m. m(k \<mapsto> v)" parametric update_parametric .
@@ -128,7 +127,7 @@
is "\<lambda>k m. m(k := None)" parametric delete_parametric .
lift_definition filter :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping"
- is "\<lambda>P m k. case m k of None \<Rightarrow> None | Some v \<Rightarrow> if P k v then Some v else None" .
+ is "\<lambda>P m k. case m k of None \<Rightarrow> None | Some v \<Rightarrow> if P k v then Some v else None" .
lift_definition keys :: "('a, 'b) mapping \<Rightarrow> 'a set"
is dom parametric dom_parametric .
@@ -141,20 +140,20 @@
lift_definition map :: "('c \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> 'd) \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('c, 'd) mapping"
is "\<lambda>f g m. (map_option g \<circ> m \<circ> f)" parametric map_parametric .
-
+
lift_definition map_values :: "('c \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> ('c, 'a) mapping \<Rightarrow> ('c, 'b) mapping"
- is "\<lambda>f m x. map_option (f x) (m x)" .
+ is "\<lambda>f m x. map_option (f x) (m x)" .
-lift_definition combine_with_key ::
+lift_definition combine_with_key ::
"('a \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a,'b) mapping \<Rightarrow> ('a,'b) mapping \<Rightarrow> ('a,'b) mapping"
is "\<lambda>f m1 m2 x. combine_options (f x) (m1 x) (m2 x)" parametric combine_with_key_parametric .
-lift_definition combine ::
+lift_definition combine ::
"('b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a,'b) mapping \<Rightarrow> ('a,'b) mapping \<Rightarrow> ('a,'b) mapping"
is "\<lambda>f m1 m2 x. combine_options f (m1 x) (m2 x)" parametric combine_parametric .
-definition All_mapping where
- "All_mapping m P \<longleftrightarrow> (\<forall>x. case Mapping.lookup m x of None \<Rightarrow> True | Some y \<Rightarrow> P x y)"
+definition "All_mapping m P \<longleftrightarrow>
+ (\<forall>x. case Mapping.lookup m x of None \<Rightarrow> True | Some y \<Rightarrow> P x y)"
declare [[code drop: map]]
@@ -168,52 +167,52 @@
subsection \<open>Derived operations\<close>
definition ordered_keys :: "('a::linorder, 'b) mapping \<Rightarrow> 'a list"
-where
- "ordered_keys m = (if finite (keys m) then sorted_list_of_set (keys m) else [])"
+ where "ordered_keys m = (if finite (keys m) then sorted_list_of_set (keys m) else [])"
definition is_empty :: "('a, 'b) mapping \<Rightarrow> bool"
-where
- "is_empty m \<longleftrightarrow> keys m = {}"
+ where "is_empty m \<longleftrightarrow> keys m = {}"
definition size :: "('a, 'b) mapping \<Rightarrow> nat"
-where
- "size m = (if finite (keys m) then card (keys m) else 0)"
+ where "size m = (if finite (keys m) then card (keys m) else 0)"
definition replace :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping"
-where
- "replace k v m = (if k \<in> keys m then update k v m else m)"
+ where "replace k v m = (if k \<in> keys m then update k v m else m)"
definition default :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping"
-where
- "default k v m = (if k \<in> keys m then m else update k v m)"
+ where "default k v m = (if k \<in> keys m then m else update k v m)"
text \<open>Manual derivation of transfer rule is non-trivial\<close>
lift_definition map_entry :: "'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" is
- "\<lambda>k f m. (case m k of None \<Rightarrow> m
+ "\<lambda>k f m.
+ (case m k of
+ None \<Rightarrow> m
| Some v \<Rightarrow> m (k \<mapsto> (f v)))" parametric map_entry_parametric .
lemma map_entry_code [code]:
- "map_entry k f m = (case lookup m k of None \<Rightarrow> m
+ "map_entry k f m =
+ (case lookup m k of
+ None \<Rightarrow> m
| Some v \<Rightarrow> update k (f v) m)"
by transfer rule
definition map_default :: "'a \<Rightarrow> 'b \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping"
-where
- "map_default k v f m = map_entry k f (default k v m)"
+ where "map_default k v f m = map_entry k f (default k v m)"
definition of_alist :: "('k \<times> 'v) list \<Rightarrow> ('k, 'v) mapping"
-where
- "of_alist xs = foldr (\<lambda>(k, v) m. update k v m) xs empty"
+ where "of_alist xs = foldr (\<lambda>(k, v) m. update k v m) xs empty"
instantiation mapping :: (type, type) equal
begin
-definition
- "HOL.equal m1 m2 \<longleftrightarrow> (\<forall>k. lookup m1 k = lookup m2 k)"
+definition "HOL.equal m1 m2 \<longleftrightarrow> (\<forall>k. lookup m1 k = lookup m2 k)"
instance
- by standard (unfold equal_mapping_def, transfer, auto)
+ apply standard
+ unfolding equal_mapping_def
+ apply transfer
+ apply auto
+ done
end
@@ -222,9 +221,9 @@
lemma [transfer_rule]:
assumes [transfer_rule]: "bi_total A"
- assumes [transfer_rule]: "bi_unique B"
+ and [transfer_rule]: "bi_unique B"
shows "(pcr_mapping A B ===> pcr_mapping A B ===> op=) HOL.eq HOL.equal"
- by (unfold equal) transfer_prover
+ unfolding equal by transfer_prover
lemma of_alist_transfer [transfer_rule]:
assumes [transfer_rule]: "bi_unique R1"
@@ -236,195 +235,186 @@
subsection \<open>Properties\<close>
-lemma mapping_eqI:
- "(\<And>x. lookup m x = lookup m' x) \<Longrightarrow> m = m'"
+lemma mapping_eqI: "(\<And>x. lookup m x = lookup m' x) \<Longrightarrow> m = m'"
by transfer (simp add: fun_eq_iff)
-lemma mapping_eqI':
- assumes "\<And>x. x \<in> Mapping.keys m \<Longrightarrow> Mapping.lookup_default d m x = Mapping.lookup_default d m' x"
- and "Mapping.keys m = Mapping.keys m'"
- shows "m = m'"
+lemma mapping_eqI':
+ assumes "\<And>x. x \<in> Mapping.keys m \<Longrightarrow> Mapping.lookup_default d m x = Mapping.lookup_default d m' x"
+ and "Mapping.keys m = Mapping.keys m'"
+ shows "m = m'"
proof (intro mapping_eqI)
- fix x
- show "Mapping.lookup m x = Mapping.lookup m' x"
+ show "Mapping.lookup m x = Mapping.lookup m' x" for x
proof (cases "Mapping.lookup m x")
case None
- hence "x \<notin> Mapping.keys m" by transfer (simp add: dom_def)
- hence "x \<notin> Mapping.keys m'" by (simp add: assms)
- hence "Mapping.lookup m' x = None" by transfer (simp add: dom_def)
- with None show ?thesis by simp
+ then have "x \<notin> Mapping.keys m"
+ by transfer (simp add: dom_def)
+ then have "x \<notin> Mapping.keys m'"
+ by (simp add: assms)
+ then have "Mapping.lookup m' x = None"
+ by transfer (simp add: dom_def)
+ with None show ?thesis
+ by simp
next
case (Some y)
- hence A: "x \<in> Mapping.keys m" by transfer (simp add: dom_def)
- hence "x \<in> Mapping.keys m'" by (simp add: assms)
- hence "\<exists>y'. Mapping.lookup m' x = Some y'" by transfer (simp add: dom_def)
- with Some assms(1)[OF A] show ?thesis by (auto simp add: lookup_default_def)
+ then have A: "x \<in> Mapping.keys m"
+ by transfer (simp add: dom_def)
+ then have "x \<in> Mapping.keys m'"
+ by (simp add: assms)
+ then have "\<exists>y'. Mapping.lookup m' x = Some y'"
+ by transfer (simp add: dom_def)
+ with Some assms(1)[OF A] show ?thesis
+ by (auto simp add: lookup_default_def)
qed
qed
-lemma lookup_update:
- "lookup (update k v m) k = Some v"
+lemma lookup_update: "lookup (update k v m) k = Some v"
by transfer simp
-lemma lookup_update_neq:
- "k \<noteq> k' \<Longrightarrow> lookup (update k v m) k' = lookup m k'"
+lemma lookup_update_neq: "k \<noteq> k' \<Longrightarrow> lookup (update k v m) k' = lookup m k'"
by transfer simp
-lemma lookup_update':
- "Mapping.lookup (update k v m) k' = (if k = k' then Some v else lookup m k')"
+lemma lookup_update': "Mapping.lookup (update k v m) k' = (if k = k' then Some v else lookup m k')"
by (auto simp: lookup_update lookup_update_neq)
-lemma lookup_empty:
- "lookup empty k = None"
+lemma lookup_empty: "lookup empty k = None"
by transfer simp
lemma lookup_filter:
- "lookup (filter P m) k =
- (case lookup m k of None \<Rightarrow> None | Some v \<Rightarrow> if P k v then Some v else None)"
+ "lookup (filter P m) k =
+ (case lookup m k of
+ None \<Rightarrow> None
+ | Some v \<Rightarrow> if P k v then Some v else None)"
by transfer simp_all
-lemma lookup_map_values:
- "lookup (map_values f m) k = map_option (f k) (lookup m k)"
+lemma lookup_map_values: "lookup (map_values f m) k = map_option (f k) (lookup m k)"
by transfer simp_all
lemma lookup_default_empty: "lookup_default d empty k = d"
by (simp add: lookup_default_def lookup_empty)
-lemma lookup_default_update:
- "lookup_default d (update k v m) k = v"
+lemma lookup_default_update: "lookup_default d (update k v m) k = v"
by (simp add: lookup_default_def lookup_update)
lemma lookup_default_update_neq:
- "k \<noteq> k' \<Longrightarrow> lookup_default d (update k v m) k' = lookup_default d m k'"
+ "k \<noteq> k' \<Longrightarrow> lookup_default d (update k v m) k' = lookup_default d m k'"
by (simp add: lookup_default_def lookup_update_neq)
-lemma lookup_default_update':
+lemma lookup_default_update':
"lookup_default d (update k v m) k' = (if k = k' then v else lookup_default d m k')"
by (auto simp: lookup_default_update lookup_default_update_neq)
lemma lookup_default_filter:
- "lookup_default d (filter P m) k =
+ "lookup_default d (filter P m) k =
(if P k (lookup_default d m k) then lookup_default d m k else d)"
by (simp add: lookup_default_def lookup_filter split: option.splits)
lemma lookup_default_map_values:
"lookup_default (f k d) (map_values f m) k = f k (lookup_default d m k)"
- by (simp add: lookup_default_def lookup_map_values split: option.splits)
+ by (simp add: lookup_default_def lookup_map_values split: option.splits)
lemma lookup_combine_with_key:
- "Mapping.lookup (combine_with_key f m1 m2) x =
- combine_options (f x) (Mapping.lookup m1 x) (Mapping.lookup m2 x)"
+ "Mapping.lookup (combine_with_key f m1 m2) x =
+ combine_options (f x) (Mapping.lookup m1 x) (Mapping.lookup m2 x)"
by transfer (auto split: option.splits)
-
+
lemma combine_altdef: "combine f m1 m2 = combine_with_key (\<lambda>_. f) m1 m2"
by transfer' (rule refl)
lemma lookup_combine:
- "Mapping.lookup (combine f m1 m2) x =
+ "Mapping.lookup (combine f m1 m2) x =
combine_options f (Mapping.lookup m1 x) (Mapping.lookup m2 x)"
by transfer (auto split: option.splits)
-
-lemma lookup_default_neutral_combine_with_key:
+
+lemma lookup_default_neutral_combine_with_key:
assumes "\<And>x. f k d x = x" "\<And>x. f k x d = x"
- shows "Mapping.lookup_default d (combine_with_key f m1 m2) k =
- f k (Mapping.lookup_default d m1 k) (Mapping.lookup_default d m2 k)"
+ shows "Mapping.lookup_default d (combine_with_key f m1 m2) k =
+ f k (Mapping.lookup_default d m1 k) (Mapping.lookup_default d m2 k)"
by (auto simp: lookup_default_def lookup_combine_with_key assms split: option.splits)
-
-lemma lookup_default_neutral_combine:
+
+lemma lookup_default_neutral_combine:
assumes "\<And>x. f d x = x" "\<And>x. f x d = x"
- shows "Mapping.lookup_default d (combine f m1 m2) x =
- f (Mapping.lookup_default d m1 x) (Mapping.lookup_default d m2 x)"
+ shows "Mapping.lookup_default d (combine f m1 m2) x =
+ f (Mapping.lookup_default d m1 x) (Mapping.lookup_default d m2 x)"
by (auto simp: lookup_default_def lookup_combine assms split: option.splits)
-lemma lookup_map_entry:
- "lookup (map_entry x f m) x = map_option f (lookup m x)"
+lemma lookup_map_entry: "lookup (map_entry x f m) x = map_option f (lookup m x)"
by transfer (auto split: option.splits)
-lemma lookup_map_entry_neq:
- "x \<noteq> y \<Longrightarrow> lookup (map_entry x f m) y = lookup m y"
+lemma lookup_map_entry_neq: "x \<noteq> y \<Longrightarrow> lookup (map_entry x f m) y = lookup m y"
by transfer (auto split: option.splits)
lemma lookup_map_entry':
- "lookup (map_entry x f m) y =
+ "lookup (map_entry x f m) y =
(if x = y then map_option f (lookup m y) else lookup m y)"
by transfer (auto split: option.splits)
-
-lemma lookup_default:
- "lookup (default x d m) x = Some (lookup_default d m x)"
- unfolding lookup_default_def default_def
- by transfer (auto split: option.splits)
-lemma lookup_default_neq:
- "x \<noteq> y \<Longrightarrow> lookup (default x d m) y = lookup m y"
- unfolding lookup_default_def default_def
- by transfer (auto split: option.splits)
+lemma lookup_default: "lookup (default x d m) x = Some (lookup_default d m x)"
+ unfolding lookup_default_def default_def
+ by transfer (auto split: option.splits)
+
+lemma lookup_default_neq: "x \<noteq> y \<Longrightarrow> lookup (default x d m) y = lookup m y"
+ unfolding lookup_default_def default_def
+ by transfer (auto split: option.splits)
lemma lookup_default':
- "lookup (default x d m) y =
- (if x = y then Some (lookup_default d m x) else lookup m y)"
+ "lookup (default x d m) y =
+ (if x = y then Some (lookup_default d m x) else lookup m y)"
unfolding lookup_default_def default_def
by transfer (auto split: option.splits)
-
-lemma lookup_map_default:
- "lookup (map_default x d f m) x = Some (f (lookup_default d m x))"
- unfolding lookup_default_def default_def
- by (simp add: map_default_def lookup_map_entry lookup_default lookup_default_def)
-lemma lookup_map_default_neq:
- "x \<noteq> y \<Longrightarrow> lookup (map_default x d f m) y = lookup m y"
- unfolding lookup_default_def default_def
- by (simp add: map_default_def lookup_map_entry_neq lookup_default_neq)
+lemma lookup_map_default: "lookup (map_default x d f m) x = Some (f (lookup_default d m x))"
+ unfolding lookup_default_def default_def
+ by (simp add: map_default_def lookup_map_entry lookup_default lookup_default_def)
+
+lemma lookup_map_default_neq: "x \<noteq> y \<Longrightarrow> lookup (map_default x d f m) y = lookup m y"
+ unfolding lookup_default_def default_def
+ by (simp add: map_default_def lookup_map_entry_neq lookup_default_neq)
lemma lookup_map_default':
- "lookup (map_default x d f m) y =
- (if x = y then Some (f (lookup_default d m x)) else lookup m y)"
- unfolding lookup_default_def default_def
- by (simp add: map_default_def lookup_map_entry' lookup_default' lookup_default_def)
+ "lookup (map_default x d f m) y =
+ (if x = y then Some (f (lookup_default d m x)) else lookup m y)"
+ unfolding lookup_default_def default_def
+ by (simp add: map_default_def lookup_map_entry' lookup_default' lookup_default_def)
-lemma lookup_tabulate:
+lemma lookup_tabulate:
assumes "distinct xs"
- shows "Mapping.lookup (Mapping.tabulate xs f) x = (if x \<in> set xs then Some (f x) else None)"
+ shows "Mapping.lookup (Mapping.tabulate xs f) x = (if x \<in> set xs then Some (f x) else None)"
using assms by transfer (auto simp: map_of_eq_None_iff o_def dest!: map_of_SomeD)
lemma lookup_of_alist: "Mapping.lookup (Mapping.of_alist xs) k = map_of xs k"
by transfer simp_all
-lemma keys_is_none_rep [code_unfold]:
- "k \<in> keys m \<longleftrightarrow> \<not> (Option.is_none (lookup m k))"
+lemma keys_is_none_rep [code_unfold]: "k \<in> keys m \<longleftrightarrow> \<not> (Option.is_none (lookup m k))"
by transfer (auto simp add: Option.is_none_def)
lemma update_update:
"update k v (update k w m) = update k v m"
"k \<noteq> l \<Longrightarrow> update k v (update l w m) = update l w (update k v m)"
- by (transfer, simp add: fun_upd_twist)+
+ by (transfer; simp add: fun_upd_twist)+
-lemma update_delete [simp]:
- "update k v (delete k m) = update k v m"
+lemma update_delete [simp]: "update k v (delete k m) = update k v m"
by transfer simp
lemma delete_update:
"delete k (update k v m) = delete k m"
"k \<noteq> l \<Longrightarrow> delete k (update l v m) = update l v (delete k m)"
- by (transfer, simp add: fun_upd_twist)+
+ by (transfer; simp add: fun_upd_twist)+
-lemma delete_empty [simp]:
- "delete k empty = empty"
+lemma delete_empty [simp]: "delete k empty = empty"
by transfer simp
lemma replace_update:
"k \<notin> keys m \<Longrightarrow> replace k v m = m"
"k \<in> keys m \<Longrightarrow> replace k v m = update k v m"
- by (transfer, auto simp add: replace_def fun_upd_twist)+
-
+ by (transfer; auto simp add: replace_def fun_upd_twist)+
+
lemma map_values_update: "map_values f (update k v m) = update k (f k v) (map_values f m)"
by transfer (simp_all add: fun_eq_iff)
-
-lemma size_mono:
- "finite (keys m') \<Longrightarrow> keys m \<subseteq> keys m' \<Longrightarrow> size m \<le> size m'"
+
+lemma size_mono: "finite (keys m') \<Longrightarrow> keys m \<subseteq> keys m' \<Longrightarrow> size m \<le> size m'"
unfolding size_def by (auto intro: card_mono)
-lemma size_empty [simp]:
- "size empty = 0"
+lemma size_empty [simp]: "size empty = 0"
unfolding size_def by transfer simp
lemma size_update:
@@ -432,13 +422,11 @@
(if k \<in> keys m then size m else Suc (size m))"
unfolding size_def by transfer (auto simp add: insert_dom)
-lemma size_delete:
- "size (delete k m) = (if k \<in> keys m then size m - 1 else size m)"
+lemma size_delete: "size (delete k m) = (if k \<in> keys m then size m - 1 else size m)"
unfolding size_def by transfer simp
-lemma size_tabulate [simp]:
- "size (tabulate ks f) = length (remdups ks)"
- unfolding size_def by transfer (auto simp add: map_of_map_restrict card_set comp_def)
+lemma size_tabulate [simp]: "size (tabulate ks f) = length (remdups ks)"
+ unfolding size_def by transfer (auto simp add: map_of_map_restrict card_set comp_def)
lemma keys_filter: "keys (filter P m) \<subseteq> keys m"
by transfer (auto split: option.splits)
@@ -446,131 +434,113 @@
lemma size_filter: "finite (keys m) \<Longrightarrow> size (filter P m) \<le> size m"
by (intro size_mono keys_filter)
-
-lemma bulkload_tabulate:
- "bulkload xs = tabulate [0..<length xs] (nth xs)"
+lemma bulkload_tabulate: "bulkload xs = tabulate [0..<length xs] (nth xs)"
by transfer (auto simp add: map_of_map_restrict)
-lemma is_empty_empty [simp]:
- "is_empty empty"
- unfolding is_empty_def by transfer simp
-
-lemma is_empty_update [simp]:
- "\<not> is_empty (update k v m)"
+lemma is_empty_empty [simp]: "is_empty empty"
unfolding is_empty_def by transfer simp
-lemma is_empty_delete:
- "is_empty (delete k m) \<longleftrightarrow> is_empty m \<or> keys m = {k}"
+lemma is_empty_update [simp]: "\<not> is_empty (update k v m)"
+ unfolding is_empty_def by transfer simp
+
+lemma is_empty_delete: "is_empty (delete k m) \<longleftrightarrow> is_empty m \<or> keys m = {k}"
unfolding is_empty_def by transfer (auto simp del: dom_eq_empty_conv)
-lemma is_empty_replace [simp]:
- "is_empty (replace k v m) \<longleftrightarrow> is_empty m"
+lemma is_empty_replace [simp]: "is_empty (replace k v m) \<longleftrightarrow> is_empty m"
unfolding is_empty_def replace_def by transfer auto
-lemma is_empty_default [simp]:
- "\<not> is_empty (default k v m)"
+lemma is_empty_default [simp]: "\<not> is_empty (default k v m)"
unfolding is_empty_def default_def by transfer auto
-lemma is_empty_map_entry [simp]:
- "is_empty (map_entry k f m) \<longleftrightarrow> is_empty m"
+lemma is_empty_map_entry [simp]: "is_empty (map_entry k f m) \<longleftrightarrow> is_empty m"
unfolding is_empty_def by transfer (auto split: option.split)
-lemma is_empty_map_values [simp]:
- "is_empty (map_values f m) \<longleftrightarrow> is_empty m"
+lemma is_empty_map_values [simp]: "is_empty (map_values f m) \<longleftrightarrow> is_empty m"
unfolding is_empty_def by transfer (auto simp: fun_eq_iff)
-lemma is_empty_map_default [simp]:
- "\<not> is_empty (map_default k v f m)"
+lemma is_empty_map_default [simp]: "\<not> is_empty (map_default k v f m)"
by (simp add: map_default_def)
-lemma keys_dom_lookup:
- "keys m = dom (Mapping.lookup m)"
+lemma keys_dom_lookup: "keys m = dom (Mapping.lookup m)"
by transfer rule
-lemma keys_empty [simp]:
- "keys empty = {}"
+lemma keys_empty [simp]: "keys empty = {}"
by transfer simp
-lemma keys_update [simp]:
- "keys (update k v m) = insert k (keys m)"
+lemma keys_update [simp]: "keys (update k v m) = insert k (keys m)"
by transfer simp
-lemma keys_delete [simp]:
- "keys (delete k m) = keys m - {k}"
+lemma keys_delete [simp]: "keys (delete k m) = keys m - {k}"
by transfer simp
-lemma keys_replace [simp]:
- "keys (replace k v m) = keys m"
+lemma keys_replace [simp]: "keys (replace k v m) = keys m"
unfolding replace_def by transfer (simp add: insert_absorb)
-lemma keys_default [simp]:
- "keys (default k v m) = insert k (keys m)"
+lemma keys_default [simp]: "keys (default k v m) = insert k (keys m)"
unfolding default_def by transfer (simp add: insert_absorb)
-lemma keys_map_entry [simp]:
- "keys (map_entry k f m) = keys m"
+lemma keys_map_entry [simp]: "keys (map_entry k f m) = keys m"
by transfer (auto split: option.split)
-lemma keys_map_default [simp]:
- "keys (map_default k v f m) = insert k (keys m)"
+lemma keys_map_default [simp]: "keys (map_default k v f m) = insert k (keys m)"
by (simp add: map_default_def)
-lemma keys_map_values [simp]:
- "keys (map_values f m) = keys m"
+lemma keys_map_values [simp]: "keys (map_values f m) = keys m"
by transfer (simp_all add: dom_def)
-lemma keys_combine_with_key [simp]:
+lemma keys_combine_with_key [simp]:
"Mapping.keys (combine_with_key f m1 m2) = Mapping.keys m1 \<union> Mapping.keys m2"
- by transfer (auto simp: dom_def combine_options_def split: option.splits)
+ by transfer (auto simp: dom_def combine_options_def split: option.splits)
lemma keys_combine [simp]: "Mapping.keys (combine f m1 m2) = Mapping.keys m1 \<union> Mapping.keys m2"
by (simp add: combine_altdef)
-lemma keys_tabulate [simp]:
- "keys (tabulate ks f) = set ks"
+lemma keys_tabulate [simp]: "keys (tabulate ks f) = set ks"
by transfer (simp add: map_of_map_restrict o_def)
lemma keys_of_alist [simp]: "keys (of_alist xs) = set (List.map fst xs)"
by transfer (simp_all add: dom_map_of_conv_image_fst)
-lemma keys_bulkload [simp]:
- "keys (bulkload xs) = {0..<length xs}"
+lemma keys_bulkload [simp]: "keys (bulkload xs) = {0..<length xs}"
by (simp add: bulkload_tabulate)
-lemma distinct_ordered_keys [simp]:
- "distinct (ordered_keys m)"
+lemma distinct_ordered_keys [simp]: "distinct (ordered_keys m)"
by (simp add: ordered_keys_def)
-lemma ordered_keys_infinite [simp]:
- "\<not> finite (keys m) \<Longrightarrow> ordered_keys m = []"
+lemma ordered_keys_infinite [simp]: "\<not> finite (keys m) \<Longrightarrow> ordered_keys m = []"
by (simp add: ordered_keys_def)
-lemma ordered_keys_empty [simp]:
- "ordered_keys empty = []"
+lemma ordered_keys_empty [simp]: "ordered_keys empty = []"
by (simp add: ordered_keys_def)
lemma ordered_keys_update [simp]:
"k \<in> keys m \<Longrightarrow> ordered_keys (update k v m) = ordered_keys m"
- "finite (keys m) \<Longrightarrow> k \<notin> keys m \<Longrightarrow> ordered_keys (update k v m) = insort k (ordered_keys m)"
- by (simp_all add: ordered_keys_def) (auto simp only: sorted_list_of_set_insert [symmetric] insert_absorb)
+ "finite (keys m) \<Longrightarrow> k \<notin> keys m \<Longrightarrow>
+ ordered_keys (update k v m) = insort k (ordered_keys m)"
+ by (simp_all add: ordered_keys_def)
+ (auto simp only: sorted_list_of_set_insert [symmetric] insert_absorb)
-lemma ordered_keys_delete [simp]:
- "ordered_keys (delete k m) = remove1 k (ordered_keys m)"
+lemma ordered_keys_delete [simp]: "ordered_keys (delete k m) = remove1 k (ordered_keys m)"
proof (cases "finite (keys m)")
- case False then show ?thesis by simp
+ case False
+ then show ?thesis by simp
next
- case True note fin = True
+ case fin: True
show ?thesis
proof (cases "k \<in> keys m")
- case False with fin have "k \<notin> set (sorted_list_of_set (keys m))" by simp
- with False show ?thesis by (simp add: ordered_keys_def remove1_idem)
+ case False
+ with fin have "k \<notin> set (sorted_list_of_set (keys m))"
+ by simp
+ with False show ?thesis
+ by (simp add: ordered_keys_def remove1_idem)
next
- case True with fin show ?thesis by (simp add: ordered_keys_def sorted_list_of_set_remove)
+ case True
+ with fin show ?thesis
+ by (simp add: ordered_keys_def sorted_list_of_set_remove)
qed
qed
-lemma ordered_keys_replace [simp]:
- "ordered_keys (replace k v m) = ordered_keys m"
+lemma ordered_keys_replace [simp]: "ordered_keys (replace k v m) = ordered_keys m"
by (simp add: replace_def)
lemma ordered_keys_default [simp]:
@@ -578,8 +548,7 @@
"finite (keys m) \<Longrightarrow> k \<notin> keys m \<Longrightarrow> ordered_keys (default k v m) = insort k (ordered_keys m)"
by (simp_all add: default_def)
-lemma ordered_keys_map_entry [simp]:
- "ordered_keys (map_entry k f m) = ordered_keys m"
+lemma ordered_keys_map_entry [simp]: "ordered_keys (map_entry k f m) = ordered_keys m"
by (simp add: ordered_keys_def)
lemma ordered_keys_map_default [simp]:
@@ -587,16 +556,13 @@
"finite (keys m) \<Longrightarrow> k \<notin> keys m \<Longrightarrow> ordered_keys (map_default k v f m) = insort k (ordered_keys m)"
by (simp_all add: map_default_def)
-lemma ordered_keys_tabulate [simp]:
- "ordered_keys (tabulate ks f) = sort (remdups ks)"
+lemma ordered_keys_tabulate [simp]: "ordered_keys (tabulate ks f) = sort (remdups ks)"
by (simp add: ordered_keys_def sorted_list_of_set_sort_remdups)
-lemma ordered_keys_bulkload [simp]:
- "ordered_keys (bulkload ks) = [0..<length ks]"
+lemma ordered_keys_bulkload [simp]: "ordered_keys (bulkload ks) = [0..<length ks]"
by (simp add: ordered_keys_def)
-lemma tabulate_fold:
- "tabulate xs f = fold (\<lambda>k m. update k (f k) m) xs empty"
+lemma tabulate_fold: "tabulate xs f = fold (\<lambda>k m. update k (f k) m) xs empty"
proof transfer
fix f :: "'a \<Rightarrow> 'b" and xs
have "map_of (List.map (\<lambda>k. (k, f k)) xs) = foldr (\<lambda>k m. m(k \<mapsto> f k)) xs Map.empty"
@@ -613,21 +579,23 @@
lemma All_mapping_empty [simp]: "All_mapping Mapping.empty P"
by (auto simp: All_mapping_def lookup_empty)
-
-lemma All_mapping_update_iff:
+
+lemma All_mapping_update_iff:
"All_mapping (Mapping.update k v m) P \<longleftrightarrow> P k v \<and> All_mapping m (\<lambda>k' v'. k = k' \<or> P k' v')"
- unfolding All_mapping_def
+ unfolding All_mapping_def
proof safe
assume "\<forall>x. case Mapping.lookup (Mapping.update k v m) x of None \<Rightarrow> True | Some y \<Rightarrow> P x y"
- hence A: "case Mapping.lookup (Mapping.update k v m) x of None \<Rightarrow> True | Some y \<Rightarrow> P x y" for x
+ then have *: "case Mapping.lookup (Mapping.update k v m) x of None \<Rightarrow> True | Some y \<Rightarrow> P x y" for x
by blast
- from A[of k] show "P k v" by (simp add: lookup_update)
+ from *[of k] show "P k v"
+ by (simp add: lookup_update)
show "case Mapping.lookup m x of None \<Rightarrow> True | Some v' \<Rightarrow> k = x \<or> P x v'" for x
- using A[of x] by (auto simp add: lookup_update' split: if_splits option.splits)
+ using *[of x] by (auto simp add: lookup_update' split: if_splits option.splits)
next
assume "P k v"
assume "\<forall>x. case Mapping.lookup m x of None \<Rightarrow> True | Some v' \<Rightarrow> k = x \<or> P x v'"
- hence A: "case Mapping.lookup m x of None \<Rightarrow> True | Some v' \<Rightarrow> k = x \<or> P x v'" for x by blast
+ then have A: "case Mapping.lookup m x of None \<Rightarrow> True | Some v' \<Rightarrow> k = x \<or> P x v'" for x
+ by blast
show "case Mapping.lookup (Mapping.update k v m) x of None \<Rightarrow> True | Some xa \<Rightarrow> P x xa" for x
using \<open>P k v\<close> A[of x] by (auto simp: lookup_update' split: option.splits)
qed
@@ -636,31 +604,28 @@
"P k v \<Longrightarrow> All_mapping m (\<lambda>k' v'. k = k' \<or> P k' v') \<Longrightarrow> All_mapping (Mapping.update k v m) P"
by (simp add: All_mapping_update_iff)
-lemma All_mapping_filter_iff:
- "All_mapping (filter P m) Q \<longleftrightarrow> All_mapping m (\<lambda>k v. P k v \<longrightarrow> Q k v)"
+lemma All_mapping_filter_iff: "All_mapping (filter P m) Q \<longleftrightarrow> All_mapping m (\<lambda>k v. P k v \<longrightarrow> Q k v)"
by (auto simp: All_mapping_def lookup_filter split: option.splits)
-lemma All_mapping_filter:
- "All_mapping m Q \<Longrightarrow> All_mapping (filter P m) Q"
+lemma All_mapping_filter: "All_mapping m Q \<Longrightarrow> All_mapping (filter P m) Q"
by (auto simp: All_mapping_filter_iff intro: All_mapping_mono)
-lemma All_mapping_map_values:
- "All_mapping (map_values f m) P \<longleftrightarrow> All_mapping m (\<lambda>k v. P k (f k v))"
+lemma All_mapping_map_values: "All_mapping (map_values f m) P \<longleftrightarrow> All_mapping m (\<lambda>k v. P k (f k v))"
by (auto simp: All_mapping_def lookup_map_values split: option.splits)
-lemma All_mapping_tabulate:
- "(\<forall>x\<in>set xs. P x (f x)) \<Longrightarrow> All_mapping (Mapping.tabulate xs f) P"
- unfolding All_mapping_def
- by (intro allI, transfer) (auto split: option.split dest!: map_of_SomeD)
+lemma All_mapping_tabulate: "(\<forall>x\<in>set xs. P x (f x)) \<Longrightarrow> All_mapping (Mapping.tabulate xs f) P"
+ unfolding All_mapping_def
+ apply (intro allI)
+ apply transfer
+ apply (auto split: option.split dest!: map_of_SomeD)
+ done
lemma All_mapping_alist:
"(\<And>k v. (k, v) \<in> set xs \<Longrightarrow> P k v) \<Longrightarrow> All_mapping (Mapping.of_alist xs) P"
by (auto simp: All_mapping_def lookup_of_alist dest!: map_of_SomeD split: option.splits)
-
-lemma combine_empty [simp]:
- "combine f Mapping.empty y = y" "combine f y Mapping.empty = y"
- by (transfer, force)+
+lemma combine_empty [simp]: "combine f Mapping.empty y = y" "combine f y Mapping.empty = y"
+ by (transfer; force)+
lemma (in abel_semigroup) comm_monoid_set_combine: "comm_monoid_set (combine f) Mapping.empty"
by standard (transfer fixing: f, simp add: combine_options_ac[of f] ac_simps)+
@@ -679,19 +644,16 @@
using that by (induction xs) simp_all
from this[of "remdups xs"] show ?thesis by simp
qed
-
-lemma keys_fold_combine:
- assumes "finite A"
- shows "Mapping.keys (combine.F g A) = (\<Union>x\<in>A. Mapping.keys (g x))"
- using assms by (induction A rule: finite_induct) simp_all
+
+lemma keys_fold_combine: "finite A \<Longrightarrow> Mapping.keys (combine.F g A) = (\<Union>x\<in>A. Mapping.keys (g x))"
+ by (induct A rule: finite_induct) simp_all
end
-
+
subsection \<open>Code generator setup\<close>
hide_const (open) empty is_empty rep lookup lookup_default filter update delete ordered_keys
keys size replace default map_entry map_default tabulate bulkload map map_values combine of_alist
end
-
--- a/src/Pure/Isar/outer_syntax.scala Tue Jul 12 13:55:35 2016 +0200
+++ b/src/Pure/Isar/outer_syntax.scala Tue Jul 12 16:04:19 2016 +0200
@@ -142,14 +142,28 @@
/* line-oriented structure */
+ private val close_structure =
+ Set(Keyword.NEXT_BLOCK, Keyword.QED_BLOCK, Keyword.PRF_CLOSE)
+
def line_structure(tokens: List[Token], structure: Outer_Syntax.Line_Structure)
: Outer_Syntax.Line_Structure =
{
val improper1 = tokens.forall(_.is_improper)
val command1 = tokens.exists(_.is_command)
+ val command_depth =
+ tokens.iterator.filter(_.is_proper).toStream.headOption match {
+ case Some(tok) =>
+ if (keywords.is_command(tok, close_structure))
+ Some(structure.after_span_depth - 1)
+ else None
+ case None => None
+ }
+
val depth1 =
- if (tokens.exists(keywords.is_command(_, Keyword.theory))) 0
+ if (tokens.exists(tok =>
+ keywords.is_before_command(tok) || keywords.is_command(tok, Keyword.theory))) 0
+ else if (command_depth.isDefined) command_depth.get
else if (command1) structure.after_span_depth
else structure.span_depth
@@ -161,7 +175,8 @@
else if (keywords.is_command(tok, Keyword.theory)) (1, 0)
else if (keywords.is_command(tok, Keyword.proof_open)) (y + 2, y + 1)
else if (keywords.is_command(tok, Keyword.PRF_BLOCK == _)) (y + 2, y + 1)
- else if (keywords.is_command(tok, Keyword.QED_BLOCK == _)) (y + 1, y - 2)
+ else if (keywords.is_command(tok, Keyword.QED_BLOCK == _)) (y - 1, y - 2)
+ else if (keywords.is_command(tok, Keyword.PRF_CLOSE == _)) (y, y - 1)
else if (keywords.is_command(tok, Keyword.proof_close)) (y + 1, y - 1)
else if (keywords.is_command(tok, Keyword.qed_global)) (1, 0)
else (x, y)
--- a/src/Tools/jEdit/src/fold_handling.scala Tue Jul 12 13:55:35 2016 +0200
+++ b/src/Tools/jEdit/src/fold_handling.scala Tue Jul 12 16:04:19 2016 +0200
@@ -34,12 +34,12 @@
override def getPrecedingFoldLevels(
buffer: JEditBuffer, line: Int, seg: Segment, level: Int): java.util.List[Integer] =
{
- val struct = Token_Markup.Line_Context.next(buffer, line).structure
+ val structure = Token_Markup.Line_Context.next(buffer, line).structure
val result =
- if (line > 0 && struct.command)
+ if (line > 0 && structure.command)
Range.inclusive(line - 1, 0, -1).iterator.
map(i => Token_Markup.Line_Context.next(buffer, i).structure).
- takeWhile(_.improper).map(_ => struct.depth max 0).toList
+ takeWhile(_.improper).map(_ => structure.depth max 0).toList
else Nil
if (result.isEmpty) null