src/HOL/ex/Interpretation_in_nested_targets.thy
author haftmann
Fri, 27 Jun 2025 08:09:26 +0200
changeset 82775 61c39a9e5415
parent 73846 9447668d1b77
permissions -rw-r--r--
typo

theory Interpretation_in_nested_targets
  imports Main
begin

locale injection =
  fixes f :: \<open>'a \<Rightarrow> 'b\<close>
  assumes eqI: \<open>f x = f y \<Longrightarrow> x = y\<close>
begin

lemma eq_iff:
  \<open>x = y \<longleftrightarrow> f x = f y\<close>
  by (auto intro: eqI)

lemma inv_apply:
  \<open>inv f (f x) = x\<close>
  by (rule inv_f_f) (simp add: eqI injI)

end

context
  fixes f :: \<open>'a::linorder \<Rightarrow> 'b::linorder\<close>
  assumes \<open>strict_mono f\<close>
begin

global_interpretation strict_mono: injection f
  by standard (simp add: \<open>strict_mono f\<close> strict_mono_eq)

thm strict_mono.eq_iff
thm strict_mono.inv_apply

end

thm strict_mono.eq_iff
thm strict_mono.inv_apply

end