diff -r a80d8ec6c998 -r 3dda49e08b9d src/HOL/Nat.thy --- a/src/HOL/Nat.thy Fri Jan 04 21:49:06 2019 +0100 +++ b/src/HOL/Nat.thy Fri Jan 04 23:22:53 2019 +0100 @@ -142,8 +142,8 @@ ML \ val nat_basic_lfp_sugar = let - val ctr_sugar = the (Ctr_Sugar.ctr_sugar_of_global @{theory} @{type_name nat}); - val recx = Logic.varify_types_global @{term rec_nat}; + val ctr_sugar = the (Ctr_Sugar.ctr_sugar_of_global \<^theory> \<^type_name>\nat\); + val recx = Logic.varify_types_global \<^term>\rec_nat\; val C = body_type (fastype_of recx); in {T = HOLogic.natT, fp_res_index = 0, C = C, fun_arg_Tsss = [[], [[HOLogic.natT, C]]], @@ -153,7 +153,7 @@ setup \ let - fun basic_lfp_sugars_of _ [@{typ nat}] _ _ ctxt = + fun basic_lfp_sugars_of _ [\<^typ>\nat\] _ _ ctxt = ([], [0], [nat_basic_lfp_sugar], [], [], [], TrueI (*dummy*), [], false, ctxt) | basic_lfp_sugars_of bs arg_Ts callers callssss ctxt = BNF_LFP_Rec_Sugar.default_basic_lfp_sugars_of bs arg_Ts callers callssss ctxt; @@ -261,7 +261,7 @@ lemma Suc_n_not_n: "Suc n \ n" by (rule not_sym) (rule n_not_Suc_n) -text \A special form of induction for reasoning about @{term "m < n"} and @{term "m - n"}.\ +text \A special form of induction for reasoning about \<^term>\m < n\ and \<^term>\m - n\.\ lemma diff_induct: assumes "\x. P x 0" and "\y. P 0 (Suc y)" @@ -464,7 +464,7 @@ by (subst mult_cancel1) simp -subsection \Orders on @{typ nat}\ +subsection \Orders on \<^typ>\nat\\ subsubsection \Operation definition\ @@ -688,7 +688,7 @@ by simp (auto simp add: less_Suc_eq dest: Suc_lessD) qed -text \Can be used with \less_Suc_eq\ to get @{prop "n = m \ n < m"}.\ +text \Can be used with \less_Suc_eq\ to get \<^prop>\n = m \ n < m\.\ lemma not_less_eq: "\ m < n \ n < Suc m" by (simp only: not_less less_Suc_eq_le) @@ -883,7 +883,7 @@ qed text \Addition is the inverse of subtraction: - if @{term "n \ m"} then @{term "n + (m - n) = m"}.\ + if \<^term>\n \ m\ then \<^term>\n + (m - n) = m\.\ lemma add_diff_inverse_nat: "\ m < n \ n + (m - n) = m" for m n :: nat by (induct m n rule: diff_induct) simp_all @@ -921,7 +921,7 @@ instance nat :: ordered_cancel_comm_monoid_diff .. -subsubsection \@{term min} and @{term max}\ +subsubsection \\<^term>\min\ and \<^term>\max\\ lemma mono_Suc: "mono Suc" by (rule monoI) simp @@ -989,7 +989,7 @@ (auto dest: mult_left_le_imp_le mult_left_less_imp_less le_less_trans) -subsubsection \Additional theorems about @{term "(\)"}\ +subsubsection \Additional theorems about \<^term>\(\)\\ text \Complete induction, aka course-of-values induction\ @@ -1454,7 +1454,7 @@ with assms show "n * m \ n * q" by simp qed -text \The lattice order on @{typ nat}.\ +text \The lattice order on \<^typ>\nat\.\ instantiation nat :: distrib_lattice begin @@ -1702,7 +1702,7 @@ qed -subsection \Embedding of the naturals into any \semiring_1\: @{term of_nat}\ +subsection \Embedding of the naturals into any \semiring_1\: \<^term>\of_nat\\ context semiring_1 begin @@ -2345,7 +2345,7 @@ by (auto intro!: funpow_increasing simp: antimono_def) -subsection \The divides relation on @{typ nat}\ +subsection \The divides relation on \<^typ>\nat\\ lemma dvd_1_left [iff]: "Suc 0 dvd k" by (simp add: dvd_def)