merged
authorpaulson
Thu, 14 Jun 2018 14:23:48 +0100
changeset 68446 92ddca1edc43
parent 68441 3b11d48a711a (diff)
parent 68445 c183a6a69f2d (current diff)
child 68447 0beb927eed89
merged
src/HOL/Algebra/More_Finite_Product.thy
src/HOL/Algebra/More_Group.thy
src/HOL/Algebra/More_Ring.thy
--- a/src/HOL/Data_Structures/AA_Map.thy	Thu Jun 14 14:23:38 2018 +0100
+++ b/src/HOL/Data_Structures/AA_Map.thy	Thu Jun 14 14:23:48 2018 +0100
@@ -207,10 +207,10 @@
               post_split_max post_delete split_maxD split: prod.splits)
 
 interpretation I: Map_by_Ordered
-where empty = Leaf and lookup = lookup and update = update and delete = delete
+where empty = empty and lookup = lookup and update = update and delete = delete
 and inorder = inorder and inv = invar
 proof (standard, goal_cases)
-  case 1 show ?case by simp
+  case 1 show ?case by (simp add: empty_def)
 next
   case 2 thus ?case by(simp add: lookup_map_of)
 next
@@ -218,7 +218,7 @@
 next
   case 4 thus ?case by(simp add: inorder_delete)
 next
-  case 5 thus ?case by(simp)
+  case 5 thus ?case by(simp add: empty_def)
 next
   case 6 thus ?case by(simp add: invar_update)
 next
--- a/src/HOL/Data_Structures/AA_Set.thy	Thu Jun 14 14:23:38 2018 +0100
+++ b/src/HOL/Data_Structures/AA_Set.thy	Thu Jun 14 14:23:48 2018 +0100
@@ -12,6 +12,9 @@
 
 type_synonym 'a aa_tree = "('a,nat) tree"
 
+definition empty :: "'a aa_tree" where
+"empty = Leaf"
+
 fun lvl :: "'a aa_tree \<Rightarrow> nat" where
 "lvl Leaf = 0" |
 "lvl (Node _ _ lv _) = lv"
@@ -482,11 +485,11 @@
   (auto simp: del_list_simps inorder_adjust pre_adj_if_postL pre_adj_if_postR 
               post_split_max post_delete split_maxD split: prod.splits)
 
-interpretation I: Set_by_Ordered
-where empty = Leaf and isin = isin and insert = insert and delete = delete
+interpretation S: Set_by_Ordered
+where empty = empty and isin = isin and insert = insert and delete = delete
 and inorder = inorder and inv = invar
 proof (standard, goal_cases)
-  case 1 show ?case by simp
+  case 1 show ?case by (simp add: empty_def)
 next
   case 2 thus ?case by(simp add: isin_set_inorder)
 next
@@ -494,7 +497,7 @@
 next
   case 4 thus ?case by(simp add: inorder_delete)
 next
-  case 5 thus ?case by(simp)
+  case 5 thus ?case by(simp add: empty_def)
 next
   case 6 thus ?case by(simp add: invar_insert)
 next
--- a/src/HOL/Data_Structures/AVL_Map.thy	Thu Jun 14 14:23:38 2018 +0100
+++ b/src/HOL/Data_Structures/AVL_Map.thy	Thu Jun 14 14:23:48 2018 +0100
@@ -180,11 +180,11 @@
 qed simp_all
 
 
-interpretation Map_by_Ordered
-where empty = Leaf and lookup = lookup and update = update and delete = delete
+interpretation M: Map_by_Ordered
+where empty = empty and lookup = lookup and update = update and delete = delete
 and inorder = inorder and inv = avl
 proof (standard, goal_cases)
-  case 1 show ?case by simp
+  case 1 show ?case by (simp add: empty_def)
 next
   case 2 thus ?case by(simp add: lookup_map_of)
 next
@@ -192,7 +192,7 @@
 next
   case 4 thus ?case by(simp add: inorder_delete)
 next
-  case 5 show ?case by simp
+  case 5 show ?case by (simp add: empty_def)
 next
   case 6 thus ?case by(simp add: avl_update(1))
 next
--- a/src/HOL/Data_Structures/AVL_Set.thy	Thu Jun 14 14:23:38 2018 +0100
+++ b/src/HOL/Data_Structures/AVL_Set.thy	Thu Jun 14 14:23:48 2018 +0100
@@ -14,6 +14,9 @@
 
 type_synonym 'a avl_tree = "('a,nat) tree"
 
+definition empty :: "'a avl_tree" where
+"empty = Leaf"
+
 text \<open>Invariant:\<close>
 
 fun avl :: "'a avl_tree \<Rightarrow> bool" where
@@ -420,11 +423,11 @@
 
 subsection "Overall correctness"
 
-interpretation Set_by_Ordered
-where empty = Leaf and isin = isin and insert = insert and delete = delete
+interpretation S: Set_by_Ordered
+where empty = empty and isin = isin and insert = insert and delete = delete
 and inorder = inorder and inv = avl
 proof (standard, goal_cases)
-  case 1 show ?case by simp
+  case 1 show ?case by (simp add: empty_def)
 next
   case 2 thus ?case by(simp add: isin_set_inorder)
 next
@@ -432,7 +435,7 @@
 next
   case 4 thus ?case by(simp add: inorder_delete)
 next
-  case 5 thus ?case by simp
+  case 5 thus ?case by (simp add: empty_def)
 next
   case 6 thus ?case by (simp add: avl_insert(1))
 next
--- a/src/HOL/Data_Structures/Brother12_Map.thy	Thu Jun 14 14:23:38 2018 +0100
+++ b/src/HOL/Data_Structures/Brother12_Map.thy	Thu Jun 14 14:23:48 2018 +0100
@@ -192,7 +192,7 @@
 subsection "Overall correctness"
 
 interpretation Map_by_Ordered
-where empty = N0 and lookup = lookup and update = update.update
+where empty = empty and lookup = lookup and update = update.update
 and delete = delete.delete and inorder = inorder and inv = "\<lambda>t. \<exists>h. t \<in> B h"
 proof (standard, goal_cases)
   case 2 thus ?case by(auto intro!: lookup_map_of)
@@ -204,6 +204,6 @@
   case 6 thus ?case using update.update_type by (metis Un_iff)
 next
   case 7 thus ?case using delete.delete_type by blast
-qed auto
+qed (auto simp: empty_def)
 
 end
--- a/src/HOL/Data_Structures/Brother12_Set.thy	Thu Jun 14 14:23:38 2018 +0100
+++ b/src/HOL/Data_Structures/Brother12_Set.thy	Thu Jun 14 14:23:48 2018 +0100
@@ -19,6 +19,9 @@
   L2 'a |
   N3 "'a bro" 'a "'a bro" 'a "'a bro"
 
+definition empty :: "'a bro" where
+"empty = N0"
+
 fun inorder :: "'a bro \<Rightarrow> 'a list" where
 "inorder N0 = []" |
 "inorder (N1 t) = inorder t" |
@@ -465,7 +468,7 @@
 subsection "Overall correctness"
 
 interpretation Set_by_Ordered
-where empty = N0 and isin = isin and insert = insert.insert
+where empty = empty and isin = isin and insert = insert.insert
 and delete = delete.delete and inorder = inorder and inv = "\<lambda>t. \<exists>h. t \<in> B h"
 proof (standard, goal_cases)
   case 2 thus ?case by(auto intro!: isin_set)
@@ -477,7 +480,7 @@
   case 6 thus ?case using insert.insert_type by blast
 next
   case 7 thus ?case using delete.delete_type by blast
-qed auto
+qed (auto simp: empty_def)
 
 
 subsection \<open>Height-Size Relation\<close>
--- a/src/HOL/Data_Structures/Map_Specs.thy	Thu Jun 14 14:23:38 2018 +0100
+++ b/src/HOL/Data_Structures/Map_Specs.thy	Thu Jun 14 14:23:48 2018 +0100
@@ -31,36 +31,41 @@
 fixes lookup :: "'t \<Rightarrow> 'a \<Rightarrow> 'b option"
 fixes inorder :: "'t \<Rightarrow> ('a * 'b) list"
 fixes inv :: "'t \<Rightarrow> bool"
-assumes empty: "inorder empty = []"
-and lookup: "inv t \<and> sorted1 (inorder t) \<Longrightarrow>
+assumes inorder_empty: "inorder empty = []"
+and inorder_lookup: "inv t \<and> sorted1 (inorder t) \<Longrightarrow>
   lookup t a = map_of (inorder t) a"
-and update: "inv t \<and> sorted1 (inorder t) \<Longrightarrow>
+and inorder_update: "inv t \<and> sorted1 (inorder t) \<Longrightarrow>
   inorder(update a b t) = upd_list a b (inorder t)"
-and delete: "inv t \<and> sorted1 (inorder t) \<Longrightarrow>
+and inorder_delete: "inv t \<and> sorted1 (inorder t) \<Longrightarrow>
   inorder(delete a t) = del_list a (inorder t)"
-and inv_empty:  "inv empty"
-and inv_update: "inv t \<and> sorted1 (inorder t) \<Longrightarrow> inv(update a b t)"
-and inv_delete: "inv t \<and> sorted1 (inorder t) \<Longrightarrow> inv(delete a t)"
+and inorder_inv_empty:  "inv empty"
+and inorder_inv_update: "inv t \<and> sorted1 (inorder t) \<Longrightarrow> inv(update a b t)"
+and inorder_inv_delete: "inv t \<and> sorted1 (inorder t) \<Longrightarrow> inv(delete a t)"
 begin
 
 text \<open>It implements the traditional specification:\<close>
 
+definition invar :: "'t \<Rightarrow> bool" where
+"invar t == inv t \<and> sorted1 (inorder t)"
+
 sublocale Map
-  empty update delete lookup "\<lambda>t. inv t \<and> sorted1 (inorder t)"
+  empty update delete lookup invar
 proof(standard, goal_cases)
-  case 1 show ?case by (auto simp: lookup empty inv_empty)
+  case 1 show ?case by (auto simp: inorder_lookup inorder_empty inorder_inv_empty)
 next
   case 2 thus ?case
-    by(simp add: fun_eq_iff update inv_update map_of_ins_list lookup sorted_upd_list)
+    by(simp add: fun_eq_iff inorder_update inorder_inv_update map_of_ins_list inorder_lookup
+         sorted_upd_list invar_def)
 next
   case 3 thus ?case
-    by(simp add: fun_eq_iff delete inv_delete map_of_del_list lookup sorted_del_list)
-next
-  case 4 thus ?case by(simp add: empty inv_empty)
+    by(simp add: fun_eq_iff inorder_delete inorder_inv_delete map_of_del_list inorder_lookup
+         sorted_del_list invar_def)
 next
-  case 5 thus ?case by(simp add: update inv_update sorted_upd_list)
+  case 4 thus ?case by(simp add: inorder_empty inorder_inv_empty invar_def)
 next
-  case 6 thus ?case by (auto simp: delete inv_delete sorted_del_list)
+  case 5 thus ?case by(simp add: inorder_update inorder_inv_update sorted_upd_list invar_def)
+next
+  case 6 thus ?case by (auto simp: inorder_delete inorder_inv_delete sorted_del_list invar_def)
 qed
 
 end
--- a/src/HOL/Data_Structures/RBT_Map.thy	Thu Jun 14 14:23:38 2018 +0100
+++ b/src/HOL/Data_Structures/RBT_Map.thy	Thu Jun 14 14:23:48 2018 +0100
@@ -104,11 +104,11 @@
 theorem rbt_delete: "rbt t \<Longrightarrow> rbt (delete k t)"
 by (metis delete_def rbt_def color_paint_Black del_invc_invh invc_paint_Black invc2I invh_paint)
 
-interpretation Map_by_Ordered
-where empty = Leaf and lookup = lookup and update = update and delete = delete
+interpretation M: Map_by_Ordered
+where empty = empty and lookup = lookup and update = update and delete = delete
 and inorder = inorder and inv = rbt
 proof (standard, goal_cases)
-  case 1 show ?case by simp
+  case 1 show ?case by (simp add: empty_def)
 next
   case 2 thus ?case by(simp add: lookup_map_of)
 next
@@ -116,7 +116,7 @@
 next
   case 4 thus ?case by(simp add: inorder_delete)
 next
-  case 5 thus ?case by (simp add: rbt_Leaf) 
+  case 5 thus ?case by (simp add: rbt_def empty_def) 
 next
   case 6 thus ?case by (simp add: rbt_update) 
 next
--- a/src/HOL/Data_Structures/RBT_Set.thy	Thu Jun 14 14:23:38 2018 +0100
+++ b/src/HOL/Data_Structures/RBT_Set.thy	Thu Jun 14 14:23:48 2018 +0100
@@ -10,6 +10,9 @@
   Isin2
 begin
 
+definition empty :: "'a rbt" where
+"empty = Leaf"
+
 fun ins :: "'a::linorder \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt" where
 "ins x Leaf = R Leaf x Leaf" |
 "ins x (B l a r) =
@@ -121,9 +124,6 @@
 lemma color_paint_Black: "color (paint Black t) = Black"
 by (cases t) auto
 
-theorem rbt_Leaf: "rbt Leaf"
-by (simp add: rbt_def)
-
 lemma paint_invc2: "invc2 t \<Longrightarrow> invc2 (paint c t)"
 by (cases t) auto
 
@@ -256,11 +256,11 @@
 
 text \<open>Overall correctness:\<close>
 
-interpretation Set_by_Ordered
-where empty = Leaf and isin = isin and insert = insert and delete = delete
+interpretation S: Set_by_Ordered
+where empty = empty and isin = isin and insert = insert and delete = delete
 and inorder = inorder and inv = rbt
 proof (standard, goal_cases)
-  case 1 show ?case by simp
+  case 1 show ?case by (simp add: empty_def)
 next
   case 2 thus ?case by(simp add: isin_set_inorder)
 next
@@ -268,7 +268,7 @@
 next
   case 4 thus ?case by(simp add: inorder_delete)
 next
-  case 5 thus ?case by (simp add: rbt_Leaf) 
+  case 5 thus ?case by (simp add: rbt_def empty_def) 
 next
   case 6 thus ?case by (simp add: rbt_insert) 
 next
--- a/src/HOL/Data_Structures/Set_Specs.thy	Thu Jun 14 14:23:38 2018 +0100
+++ b/src/HOL/Data_Structures/Set_Specs.thy	Thu Jun 14 14:23:48 2018 +0100
@@ -33,37 +33,43 @@
 fixes isin :: "'t \<Rightarrow> 'a \<Rightarrow> bool"
 fixes inorder :: "'t \<Rightarrow> 'a list"
 fixes inv :: "'t \<Rightarrow> bool"
-assumes empty: "inorder empty = []"
+assumes inorder_empty: "inorder empty = []"
 assumes isin: "inv t \<and> sorted(inorder t) \<Longrightarrow>
   isin t x = (x \<in> set (inorder t))"
-assumes insert: "inv t \<and> sorted(inorder t) \<Longrightarrow>
+assumes inorder_insert: "inv t \<and> sorted(inorder t) \<Longrightarrow>
   inorder(insert x t) = ins_list x (inorder t)"
-assumes delete: "inv t \<and> sorted(inorder t) \<Longrightarrow>
+assumes inorder_delete: "inv t \<and> sorted(inorder t) \<Longrightarrow>
   inorder(delete x t) = del_list x (inorder t)"
-assumes inv_empty:  "inv empty"
-assumes inv_insert: "inv t \<and> sorted(inorder t) \<Longrightarrow> inv(insert x t)"
-assumes inv_delete: "inv t \<and> sorted(inorder t) \<Longrightarrow> inv(delete x t)"
+assumes inorder_inv_empty:  "inv empty"
+assumes inorder_inv_insert: "inv t \<and> sorted(inorder t) \<Longrightarrow> inv(insert x t)"
+assumes inorder_inv_delete: "inv t \<and> sorted(inorder t) \<Longrightarrow> inv(delete x t)"
 begin
 
 text \<open>It implements the traditional specification:\<close>
 
+definition set :: "'t \<Rightarrow> 'a set" where
+"set == List.set o inorder"
+
+definition invar :: "'t \<Rightarrow> bool" where
+"invar t == inv t \<and> sorted (inorder t)"
+
 sublocale Set
-  empty insert delete isin "set o inorder" "\<lambda>t. inv t \<and> sorted(inorder t)"
+  empty insert delete isin set invar
 proof(standard, goal_cases)
-  case 1 show ?case by (auto simp: empty)
+  case 1 show ?case by (auto simp: inorder_empty set_def)
 next
-  case 2 thus ?case by(simp add: isin)
+  case 2 thus ?case by(simp add: isin invar_def set_def)
 next
-  case 3 thus ?case by(simp add: insert set_ins_list)
+  case 3 thus ?case by(simp add: inorder_insert set_ins_list set_def invar_def)
 next
   case (4 s x) thus ?case
-    using delete[OF 4, of x] by (auto simp: distinct_if_sorted set_del_list_eq)
+    by (auto simp: inorder_delete distinct_if_sorted set_del_list_eq invar_def set_def)
 next
-  case 5 thus ?case by(simp add: empty inv_empty)
+  case 5 thus ?case by(simp add: inorder_empty inorder_inv_empty invar_def)
 next
-  case 6 thus ?case by(simp add: insert inv_insert sorted_ins_list)
+  case 6 thus ?case by(simp add: inorder_insert inorder_inv_insert sorted_ins_list invar_def)
 next
-  case 7 thus ?case by (auto simp: delete inv_delete sorted_del_list)
+  case 7 thus ?case by (auto simp: inorder_delete inorder_inv_delete sorted_del_list invar_def)
 qed
 
 end
--- a/src/HOL/Data_Structures/Tree234_Map.thy	Thu Jun 14 14:23:38 2018 +0100
+++ b/src/HOL/Data_Structures/Tree234_Map.thy	Thu Jun 14 14:23:48 2018 +0100
@@ -160,8 +160,8 @@
 
 subsection \<open>Overall Correctness\<close>
 
-interpretation Map_by_Ordered
-where empty = Leaf and lookup = lookup and update = update and delete = delete
+interpretation M: Map_by_Ordered
+where empty = empty and lookup = lookup and update = update and delete = delete
 and inorder = inorder and inv = bal
 proof (standard, goal_cases)
   case 2 thus ?case by(simp add: lookup_map_of)
@@ -173,6 +173,6 @@
   case 6 thus ?case by(simp add: bal_update)
 next
   case 7 thus ?case by(simp add: bal_delete)
-qed simp+
+qed (simp add: empty_def)+
 
 end
--- a/src/HOL/Data_Structures/Tree234_Set.thy	Thu Jun 14 14:23:38 2018 +0100
+++ b/src/HOL/Data_Structures/Tree234_Set.thy	Thu Jun 14 14:23:48 2018 +0100
@@ -13,6 +13,9 @@
 
 subsection \<open>Set operations on 2-3-4 trees\<close>
 
+definition empty :: "'a tree234" where
+"empty = Leaf"
+
 fun isin :: "'a::linorder tree234 \<Rightarrow> 'a \<Rightarrow> bool" where
 "isin Leaf x = False" |
 "isin (Node2 l a r) x =
@@ -501,8 +504,8 @@
 
 subsection \<open>Overall Correctness\<close>
 
-interpretation Set_by_Ordered
-where empty = Leaf and isin = isin and insert = insert and delete = delete
+interpretation S: Set_by_Ordered
+where empty = empty and isin = isin and insert = insert and delete = delete
 and inorder = inorder and inv = bal
 proof (standard, goal_cases)
   case 2 thus ?case by(simp add: isin_set)
@@ -514,6 +517,6 @@
   case 6 thus ?case by(simp add: bal_insert)
 next
   case 7 thus ?case by(simp add: bal_delete)
-qed simp+
+qed (simp add: empty_def)+
 
 end
--- a/src/HOL/Data_Structures/Tree23_Map.thy	Thu Jun 14 14:23:38 2018 +0100
+++ b/src/HOL/Data_Structures/Tree23_Map.thy	Thu Jun 14 14:23:48 2018 +0100
@@ -119,19 +119,23 @@
 
 subsection \<open>Overall Correctness\<close>
 
-interpretation Map_by_Ordered
-where empty = Leaf and lookup = lookup and update = update and delete = delete
+interpretation M: Map_by_Ordered
+where empty = empty and lookup = lookup and update = update and delete = delete
 and inorder = inorder and inv = bal
 proof (standard, goal_cases)
+  case 1 thus ?case by(simp add: empty_def)
+next
   case 2 thus ?case by(simp add: lookup_map_of)
 next
   case 3 thus ?case by(simp add: inorder_update)
 next
   case 4 thus ?case by(simp add: inorder_delete)
 next
+  case 5 thus ?case by(simp add: empty_def)
+next
   case 6 thus ?case by(simp add: bal_update)
 next
   case 7 thus ?case by(simp add: bal_delete)
-qed simp+
+qed
 
 end
--- a/src/HOL/Data_Structures/Tree23_Set.thy	Thu Jun 14 14:23:38 2018 +0100
+++ b/src/HOL/Data_Structures/Tree23_Set.thy	Thu Jun 14 14:23:48 2018 +0100
@@ -11,6 +11,9 @@
 
 declare sorted_wrt.simps(2)[simp del]
 
+definition empty :: "'a tree23" where
+"empty = Leaf"
+
 fun isin :: "'a::linorder tree23 \<Rightarrow> 'a \<Rightarrow> bool" where
 "isin Leaf x = False" |
 "isin (Node2 l a r) x =
@@ -376,8 +379,8 @@
 
 subsection \<open>Overall Correctness\<close>
 
-interpretation Set_by_Ordered
-where empty = Leaf and isin = isin and insert = insert and delete = delete
+interpretation S: Set_by_Ordered
+where empty = empty and isin = isin and insert = insert and delete = delete
 and inorder = inorder and inv = bal
 proof (standard, goal_cases)
   case 2 thus ?case by(simp add: isin_set)
@@ -389,6 +392,6 @@
   case 6 thus ?case by(simp add: bal_insert)
 next
   case 7 thus ?case by(simp add: bal_delete)
-qed simp+
+qed (simp add: empty_def)+
 
 end
--- a/src/HOL/Data_Structures/Tree_Map.thy	Thu Jun 14 14:23:38 2018 +0100
+++ b/src/HOL/Data_Structures/Tree_Map.thy	Thu Jun 14 14:23:48 2018 +0100
@@ -42,11 +42,11 @@
   "sorted1(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)"
 by(induction t) (auto simp: del_list_simps split_minD split: prod.splits)
 
-interpretation Map_by_Ordered
-where empty = Leaf and lookup = lookup and update = update and delete = delete
+interpretation M: Map_by_Ordered
+where empty = empty and lookup = lookup and update = update and delete = delete
 and inorder = inorder and inv = "\<lambda>_. True"
 proof (standard, goal_cases)
-  case 1 show ?case by simp
+  case 1 show ?case by (simp add: empty_def)
 next
   case 2 thus ?case by(simp add: lookup_map_of)
 next
--- a/src/HOL/Data_Structures/Tree_Set.thy	Thu Jun 14 14:23:38 2018 +0100
+++ b/src/HOL/Data_Structures/Tree_Set.thy	Thu Jun 14 14:23:48 2018 +0100
@@ -9,6 +9,9 @@
   Set_Specs
 begin
 
+definition empty :: "'a tree" where
+"empty == Leaf"
+
 fun isin :: "'a::linorder tree \<Rightarrow> 'a \<Rightarrow> bool" where
 "isin Leaf x = False" |
 "isin (Node l a r) x =
@@ -59,11 +62,11 @@
   "sorted(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)"
 by(induction t) (auto simp: del_list_simps split_minD split: prod.splits)
 
-interpretation Set_by_Ordered
-where empty = Leaf and isin = isin and insert = insert and delete = delete
+interpretation S: Set_by_Ordered
+where empty = empty and isin = isin and insert = insert and delete = delete
 and inorder = inorder and inv = "\<lambda>_. True"
 proof (standard, goal_cases)
-  case 1 show ?case by simp
+  case 1 show ?case by (simp add: empty_def)
 next
   case 2 thus ?case by(simp add: isin_set)
 next
--- a/src/HOL/Rat.thy	Thu Jun 14 14:23:38 2018 +0100
+++ b/src/HOL/Rat.thy	Thu Jun 14 14:23:48 2018 +0100
@@ -710,13 +710,13 @@
 lemma nonzero_of_rat_inverse: "a \<noteq> 0 \<Longrightarrow> of_rat (inverse a) = inverse (of_rat a)"
   by (rule inverse_unique [symmetric]) (simp add: of_rat_mult [symmetric])
 
-lemma of_rat_inverse: "(of_rat (inverse a) :: 'a::{field_char_0,field}) = inverse (of_rat a)"
+lemma of_rat_inverse: "(of_rat (inverse a) :: 'a::field_char_0) = inverse (of_rat a)"
   by (cases "a = 0") (simp_all add: nonzero_of_rat_inverse)
 
 lemma nonzero_of_rat_divide: "b \<noteq> 0 \<Longrightarrow> of_rat (a / b) = of_rat a / of_rat b"
   by (simp add: divide_inverse of_rat_mult nonzero_of_rat_inverse)
 
-lemma of_rat_divide: "(of_rat (a / b) :: 'a::{field_char_0,field}) = of_rat a / of_rat b"
+lemma of_rat_divide: "(of_rat (a / b) :: 'a::field_char_0) = of_rat a / of_rat b"
   by (cases "b = 0") (simp_all add: nonzero_of_rat_divide)
 
 lemma of_rat_power: "(of_rat (a ^ n) :: 'a::field_char_0) = of_rat a ^ n"
@@ -869,31 +869,17 @@
   apply (rule of_rat_mult [symmetric])
   done
 
-lemma nonzero_Rats_inverse: "a \<in> \<rat> \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> inverse a \<in> \<rat>"
+lemma Rats_inverse [simp]: "a \<in> \<rat> \<Longrightarrow> inverse a \<in> \<rat>"
   for a :: "'a::field_char_0"
   apply (auto simp add: Rats_def)
   apply (rule range_eqI)
-  apply (erule nonzero_of_rat_inverse [symmetric])
-  done
-
-lemma Rats_inverse [simp]: "a \<in> \<rat> \<Longrightarrow> inverse a \<in> \<rat>"
-  for a :: "'a::{field_char_0,field}"
-  apply (auto simp add: Rats_def)
-  apply (rule range_eqI)
   apply (rule of_rat_inverse [symmetric])
   done
 
-lemma nonzero_Rats_divide: "a \<in> \<rat> \<Longrightarrow> b \<in> \<rat> \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> a / b \<in> \<rat>"
+lemma Rats_divide [simp]: "a \<in> \<rat> \<Longrightarrow> b \<in> \<rat> \<Longrightarrow> a / b \<in> \<rat>"
   for a b :: "'a::field_char_0"
   apply (auto simp add: Rats_def)
   apply (rule range_eqI)
-  apply (erule nonzero_of_rat_divide [symmetric])
-  done
-
-lemma Rats_divide [simp]: "a \<in> \<rat> \<Longrightarrow> b \<in> \<rat> \<Longrightarrow> a / b \<in> \<rat>"
-  for a b :: "'a::{field_char_0, field}"
-  apply (auto simp add: Rats_def)
-  apply (rule range_eqI)
   apply (rule of_rat_divide [symmetric])
   done
 
--- a/src/HOL/Types_To_Sets/unoverload_type.ML	Thu Jun 14 14:23:38 2018 +0100
+++ b/src/HOL/Types_To_Sets/unoverload_type.ML	Thu Jun 14 14:23:48 2018 +0100
@@ -6,51 +6,58 @@
 
 signature UNOVERLOAD_TYPE =
 sig
-  val unoverload_type: Context.generic -> string -> thm -> thm
-  val unoverload_type_attr: string -> attribute
+  val unoverload_type: Context.generic -> indexname list -> thm -> thm
+  val unoverload_type_attr: indexname list -> attribute
 end;
 
 structure Unoverload_Type : UNOVERLOAD_TYPE =
 struct
 
-fun those [] = []
-  | those (NONE::xs) = those xs
-  | those (SOME x::xs) = x::those xs
-
-fun params_of_sort context sort =
+fun internalize_sort' ctvar thm =
   let
-    val algebra = Sign.classes_of (Context.theory_of context)
-    val params = List.concat (map (Sorts.super_classes algebra) sort) |>
-      map (try (Axclass.get_info (Context.theory_of context))) |>
-      those |>
-      map #params |>
-      List.concat
-  in params end
+    val (_, thm') = Internalize_Sort.internalize_sort ctvar thm
+    val class_premise = case Thm.prems_of thm' of t::_=> t | [] =>
+      raise THM ("internalize_sort': no premise?", 0, [thm'])
+    val class_vars = Term.add_tvars class_premise []
+    val tvar = case class_vars of [x] => TVar x | _ =>
+      raise TERM ("internalize_sort': not one type class variable.", [class_premise])
+  in
+    (tvar, thm')
+  end
 
-fun unoverload_type context s thm =
+fun params_of_class thy class = try (Axclass.get_info thy #> #params) class |> these
+
+fun params_of_super_classes thy class =
+  Sorts.super_classes (Sign.classes_of thy) class |> maps (params_of_class thy)
+
+fun params_of_sort thy sort = maps (params_of_super_classes thy) sort
+
+fun subst_TFree n' ty' ty = map_type_tfree (fn x as (n, _) => if n = n' then ty' else TFree x) ty
+
+fun unoverload_single_type context x thm =
   let
     val tvars = Term.add_tvars (Thm.prop_of thm) []
     val thy = Context.theory_of context
   in
-  case find_first (fn ((n,_), _) => n = s) tvars of NONE =>
-    raise THM ("unoverload_internalize_deps: type variable ("^s^") not in theorem", 0, [thm])
+  case find_first (fn (y, _) => x = y) tvars of NONE =>
+    raise TYPE ("unoverload_type: no such type variable in theorem", [TVar (x,[])], [])
   | SOME (x as (_, sort)) =>
     let
-      val (tvar, thm') = Internalize_Sort.internalize_sort (Thm.global_ctyp_of thy (TVar x)) thm
-      val consts = params_of_sort context sort |>
-        map (apsnd (map_type_tfree (fn ("'a", _) =>  tvar | x => TFree x)))
+      val (tvar, thm') = internalize_sort' (Thm.global_ctyp_of thy (TVar x)) thm
+      val consts = params_of_sort thy sort
+        |> map (apsnd (subst_TFree "'a" tvar) #> Const #> Thm.global_cterm_of thy)
     in
-      fold (fn c => Unoverloading.unoverload (Thm.global_cterm_of thy (Const c))) consts thm'
+      fold Unoverloading.unoverload consts thm'
+      |> Raw_Simplifier.norm_hhf (Context.proof_of context)
     end
   end
 
-val tyN = (Args.context -- Args.typ) >>
-  (fn (_, TFree (n, _)) => n
-  | (ctxt, t) => error ("Not a type variable: " ^ Syntax.string_of_typ ctxt t))
+fun unoverload_type context xs = fold (unoverload_single_type context) xs
 
-fun unoverload_type_attr n = Thm.rule_attribute [] (fn context => unoverload_type context n)
+fun unoverload_type_attr xs = Thm.rule_attribute [] (fn context => unoverload_type context xs)
 
 val _ = Context.>> (Context.map_theory (Attrib.setup @{binding unoverload_type}
-  (tyN >> unoverload_type_attr) "internalize a sort"))
+  (Scan.lift (Scan.repeat Args.var) >> unoverload_type_attr)
+    "internalize and unoverload type class parameters"))
 
 end
\ No newline at end of file