Theory Interp
section ‹Interpretation Tools›
theory Interp
imports Main
begin
subsection ‹Interpretation Locale›
locale interp =
fixes f :: "'a ⇒ 'b"
assumes f_inj : "inj f"
begin
lemma meta_interp_law:
"(⋀P. PROP Q P) ≡ (⋀P. PROP Q (P o f))"
apply (rule equal_intr_rule)
apply (drule_tac x = "P o f" in meta_spec)
apply (assumption)
apply (drule_tac x = "P o inv f" in meta_spec)
apply (simp add: f_inj)
done
lemma all_interp_law:
"(∀P. Q P) = (∀P. Q (P o f))"
apply (safe)
apply (drule_tac x = "P o f" in spec)
apply (assumption)
apply (drule_tac x = "P o inv f" in spec)
apply (simp add: f_inj)
done
lemma exists_interp_law:
"(∃P. Q P) = (∃P. Q (P o f))"
apply (safe)
apply (rule_tac x = "P o inv f" in exI)
apply (simp add: f_inj)
apply (rule_tac x = "P o f" in exI)
apply (assumption)
done
end
end
Theory Two
section ‹ Types of Cardinality 2 or Greater ›
theory Two
imports HOL.Real
begin
text ‹ The two class states that a type's carrier is either infinite, or else it has a finite
cardinality of at least 2. It is needed when we depend on having at least two distinguishable
elements. ›
class two =
assumes card_two: "infinite (UNIV :: 'a set) ∨ card (UNIV :: 'a set) ≥ 2"
begin
lemma two_diff: "∃ x y :: 'a. x ≠ y"
proof -
obtain A where "finite A" "card A = 2" "A ⊆ (UNIV :: 'a set)"
proof (cases "infinite (UNIV :: 'a set)")
case True
with infinite_arbitrarily_large[of "UNIV :: 'a set" 2] that
show ?thesis by auto
next
case False
with card_two that
show ?thesis
by (metis UNIV_bool card_UNIV_bool card_image card_le_inj finite.intros(1) finite_insert finite_subset)
qed
thus ?thesis
by (metis (full_types) One_nat_def Suc_1 UNIV_eq_I card.empty card.insert finite.intros(1) insertCI nat.inject nat.simps(3))
qed
end
instance bool :: two
by (intro_classes, auto)
instance nat :: two
by (intro_classes, auto)
instance int :: two
by (intro_classes, auto simp add: infinite_UNIV_int)
instance rat :: two
by (intro_classes, auto simp add: infinite_UNIV_char_0)
instance real :: two
by (intro_classes, auto simp add: infinite_UNIV_char_0)
instance list :: (type) two
by (intro_classes, auto simp add: infinite_UNIV_listI)
end
Theory Lens_Laws
section ‹Core Lens Laws›
theory Lens_Laws
imports
Two Interp
begin
subsection ‹Lens Signature›
text ‹This theory introduces the signature of lenses and indentifies the core algebraic hierarchy of lens
classes, including laws for well-behaved, very well-behaved, and bijective lenses~\cite{Foster07,Fischer2015,Gibbons17}.›
record ('a, 'b) lens =
lens_get :: "'b ⇒ 'a" ("getı")
lens_put :: "'b ⇒ 'a ⇒ 'b" ("putı")
type_notation
lens (infixr "⟹" 0)
text ‹ Alternative parameters ordering, inspired by Back and von Wright's refinement
calculus~\cite{Back1998}, which similarly uses two functions to characterise updates to variables. ›
abbreviation (input) lens_set :: "('a ⟹ 'b) ⇒ 'a ⇒ 'b ⇒ 'b" ("lsetı") where
"lens_set ≡ (λ X v s. put⇘X⇙ s v)"
text ‹
\begin{figure}
\begin{center}
\includegraphics[width=6cm]{figures/Lens}
\end{center}
\vspace{-5ex}
\caption{Visualisation of a simple lens}
\label{fig:Lens}
\end{figure}
A lens $X : \view \lto \src$, for source type $\src$ and view type $\view$, identifies
$\view$ with a subregion of $\src$~\cite{Foster07,Foster09}, as illustrated in Figure~\ref{fig:Lens}. The arrow denotes
$X$ and the hatched area denotes the subregion $\view$ it characterises. Transformations on
$\view$ can be performed without affecting the parts of $\src$ outside the hatched area. The lens
signature consists of a pair of functions $\lget_X : \src \Rightarrow \view$ that extracts a view
from a source, and $\lput_X : \src \Rightarrow \view \Rightarrow \src$ that updates a view within
a given source. ›
named_theorems lens_defs
text ‹ @{text lens_source} gives the set of constructible sources; that is those that can be built
by putting a value into an arbitrary source. ›
definition lens_source :: "('a ⟹ 'b) ⇒ 'b set" ("𝒮ı") where
"lens_source X = {s. ∃ v s'. s = put⇘X⇙ s' v}"
abbreviation some_source :: "('a ⟹ 'b) ⇒ 'b" ("srcı") where
"some_source X ≡ (SOME s. s ∈ 𝒮⇘X⇙)"
definition lens_create :: "('a ⟹ 'b) ⇒ 'a ⇒ 'b" ("createı") where
[lens_defs]: "create⇘X⇙ v = put⇘X⇙ (src⇘X⇙) v"
text ‹ Function $\lcreate_X~v$ creates an instance of the source type of $X$ by injecting $v$
as the view, and leaving the remaining context arbitrary. ›
definition lens_update :: "('a ⟹ 'b) ⇒ ('a ⇒ 'a) ⇒ ('b ⇒ 'b)" ("updateı") where
[lens_defs]: "lens_update X f σ = put⇘X⇙ σ (f (get⇘X⇙ σ))"
text ‹ The update function is analogous to the record update function which lifts a function
on a view type to one on the source type. ›
definition lens_obs_eq :: "('b ⟹ 'a) ⇒ 'a ⇒ 'a ⇒ bool" (infix "≃ı" 50) where
[lens_defs]: "s⇩1 ≃⇘X⇙ s⇩2 = (s⇩1 = put⇘X⇙ s⇩2 (get⇘X⇙ s⇩1))"
text ‹ This relation states that two sources are equivalent outside of the region characterised
by lens $X$. ›
definition lens_override :: "('b ⟹ 'a) ⇒ 'a ⇒ 'a ⇒ 'a" (infixl "◃ı" 95) where
[lens_defs]: "S⇩1 ◃⇘X⇙ S⇩2 = put⇘X⇙ S⇩1 (get⇘X⇙ S⇩2)"
abbreviation (input) lens_override' :: "'a ⇒ 'a ⇒ ('b ⟹ 'a) ⇒ 'a" ("_ ⊕⇩L _ on _" [95,0,96] 95) where
"S⇩1 ⊕⇩L S⇩2 on X ≡ S⇩1 ◃⇘X⇙ S⇩2"
text ‹Lens override uses a lens to replace part of a source type with a given value for the
corresponding view.›
subsection ‹Weak Lenses›
text ‹ Weak lenses are the least constrained class of lenses in our algebraic hierarchy. They
simply require that the PutGet law~\cite{Foster09,Fischer2015} is satisfied, meaning that
$\lget$ is the inverse of $\lput$. ›
locale weak_lens =
fixes x :: "'a ⟹ 'b" (structure)
assumes put_get: "get (put σ v) = v"
begin
lemma source_nonempty: "∃ s. s ∈ 𝒮"
by (auto simp add: lens_source_def)
lemma put_closure: "put σ v ∈ 𝒮"
by (auto simp add: lens_source_def)
lemma create_closure: "create v ∈ 𝒮"
by (simp add: lens_create_def put_closure)
lemma src_source [simp]: "src ∈ 𝒮"
using some_in_eq source_nonempty by auto
lemma create_get: "get (create v) = v"
by (simp add: lens_create_def put_get)
lemma create_inj: "inj create"
by (metis create_get injI)
lemma get_update: "get (update f σ) = f (get σ)"
by (simp add: put_get lens_update_def)
lemma view_determination:
assumes "put σ u = put ρ v"
shows "u = v"
by (metis assms put_get)
lemma put_inj: "inj (put σ)"
by (simp add: injI view_determination)
end
declare weak_lens.put_get [simp]
declare weak_lens.create_get [simp]
subsection ‹Well-behaved Lenses›
text ‹ Well-behaved lenses add to weak lenses that requirement that the GetPut law~\cite{Foster09,Fischer2015}
is satisfied, meaning that $\lput$ is the inverse of $\lget$. ›
locale wb_lens = weak_lens +
assumes get_put: "put σ (get σ) = σ"
begin
lemma put_twice: "put (put σ v) v = put σ v"
by (metis get_put put_get)
lemma put_surjectivity: "∃ ρ v. put ρ v = σ"
using get_put by blast
lemma source_stability: "∃ v. put σ v = σ"
using get_put by auto
lemma source_UNIV [simp]: "𝒮 = UNIV"
by (metis UNIV_eq_I put_closure wb_lens.source_stability wb_lens_axioms)
end
declare wb_lens.get_put [simp]
lemma wb_lens_weak [simp]: "wb_lens x ⟹ weak_lens x"
by (simp add: wb_lens_def)
subsection ‹ Mainly Well-behaved Lenses ›
text ‹ Mainly well-behaved lenses extend weak lenses with the PutPut law that shows how one put
override a previous one. ›
locale mwb_lens = weak_lens +
assumes put_put: "put (put σ v) u = put σ u"
begin
lemma update_comp: "update f (update g σ) = update (f ∘ g) σ"
by (simp add: put_get put_put lens_update_def)
text ‹ Mainly well-behaved lenses give rise to a weakened version of the $get-put$ law,
where the source must be within the set of constructible sources. ›
lemma weak_get_put: "σ ∈ 𝒮 ⟹ put σ (get σ) = σ"
by (auto simp add: lens_source_def put_get put_put)
lemma weak_source_determination:
assumes "σ ∈ 𝒮" "ρ ∈ 𝒮" "get σ = get ρ" "put σ v = put ρ v"
shows "σ = ρ"
by (metis assms put_put weak_get_put)
lemma weak_put_eq:
assumes "σ ∈ 𝒮" "get σ = k" "put σ u = put ρ v"
shows "put ρ k = σ"
by (metis assms put_put weak_get_put)
text ‹ Provides $s$ is constructible, then @{term get} can be uniquely determined from @{term put} ›
lemma weak_get_via_put: "s ∈ 𝒮 ⟹ get s = (THE v. put s v = s)"
by (rule sym, auto intro!: the_equality weak_get_put, metis put_get)
end
abbreviation (input) "partial_lens ≡ mwb_lens"
declare mwb_lens.put_put [simp]
declare mwb_lens.weak_get_put [simp]
lemma mwb_lens_weak [simp]:
"mwb_lens x ⟹ weak_lens x"
by (simp add: mwb_lens.axioms(1))
subsection ‹Very Well-behaved Lenses›
text ‹Very well-behaved lenses combine all three laws, as in the literature~\cite{Foster09,Fischer2015}.
The same set of axioms can be found in Back and von Wright's refinement calculus~\cite{Back1998},
though with different names for the functions. ›
locale vwb_lens = wb_lens + mwb_lens
begin
lemma source_determination:
assumes "get σ = get ρ" "put σ v = put ρ v"
shows "σ = ρ"
by (metis assms get_put put_put)
lemma put_eq:
assumes "get σ = k" "put σ u = put ρ v"
shows "put ρ k = σ"
using assms weak_put_eq[of σ k u ρ v] by (simp)
text ‹ @{term get} can be uniquely determined from @{term put} ›
lemma get_via_put: "get s = (THE v. put s v = s)"
by (simp add: weak_get_via_put)
lemma get_surj: "surj get"
by (metis put_get surjI)
text ‹ Observation equivalence is an equivalence relation. ›
lemma lens_obs_equiv: "equivp (≃)"
proof (rule equivpI)
show "reflp (≃)"
by (rule reflpI, simp add: lens_obs_eq_def get_put)
show "symp (≃)"
by (rule sympI, simp add: lens_obs_eq_def, metis get_put put_put)
show "transp (≃)"
by (rule transpI, simp add: lens_obs_eq_def, metis put_put)
qed
end
abbreviation (input) "total_lens ≡ vwb_lens"
lemma vwb_lens_wb [simp]: "vwb_lens x ⟹ wb_lens x"
by (simp add: vwb_lens_def)
lemma vwb_lens_mwb [simp]: "vwb_lens x ⟹ mwb_lens x"
using vwb_lens_def by auto
lemma mwb_UNIV_src_is_vwb_lens:
"⟦ mwb_lens X; 𝒮⇘X⇙ = UNIV ⟧ ⟹ vwb_lens X"
using vwb_lens_def wb_lens_axioms_def wb_lens_def by fastforce
text ‹ Alternative characterisation: a very well-behaved (i.e. total) lens is a mainly well-behaved
(i.e. partial) lens whose source is the universe set. ›
lemma vwb_lens_iff_mwb_UNIV_src:
"vwb_lens X ⟷ (mwb_lens X ∧ 𝒮⇘X⇙ = UNIV)"
by (meson mwb_UNIV_src_is_vwb_lens vwb_lens_def wb_lens.source_UNIV)
subsection ‹ Ineffectual Lenses ›
text ‹Ineffectual lenses can have no effect on the view type -- application of the $\lput$ function
always yields the same source. They are thus, trivially, very well-behaved lenses.›
locale ief_lens = weak_lens +
assumes put_inef: "put σ v = σ"
begin
sublocale vwb_lens
proof
fix σ v u
show "put σ (get σ) = σ"
by (simp add: put_inef)
show "put (put σ v) u = put σ u"
by (simp add: put_inef)
qed
lemma ineffectual_const_get:
"∃ v. ∀ σ∈𝒮. get σ = v"
using put_get put_inef by auto
end
abbreviation "eff_lens X ≡ (weak_lens X ∧ (¬ ief_lens X))"
subsection ‹ Partially Bijective Lenses ›
locale pbij_lens = weak_lens +
assumes put_det: "put σ v = put ρ v"
begin
sublocale mwb_lens
proof
fix σ v u
show "put (put σ v) u = put σ u"
using put_det by blast
qed
lemma put_is_create: "put σ v = create v"
by (simp add: lens_create_def put_det)
lemma partial_get_put: "ρ ∈ 𝒮 ⟹ put σ (get ρ) = ρ"
by (metis put_det weak_get_put)
end
lemma pbij_lens_weak [simp]:
"pbij_lens x ⟹ weak_lens x"
by (simp_all add: pbij_lens_def)
lemma pbij_lens_mwb [simp]: "pbij_lens x ⟹ mwb_lens x"
by (simp add: mwb_lens_axioms.intro mwb_lens_def pbij_lens.put_is_create)
lemma pbij_alt_intro:
"⟦ weak_lens X; ⋀ s. s ∈ 𝒮⇘X⇙ ⟹ create⇘X⇙ (get⇘X⇙ s) = s ⟧ ⟹ pbij_lens X"
by (metis pbij_lens_axioms_def pbij_lens_def weak_lens.put_closure weak_lens.put_get)
subsection ‹ Bijective Lenses ›
text ‹Bijective lenses characterise the situation where the source and view type are equivalent:
in other words the view type full characterises the whole source type. It is often useful
when the view type and source type are syntactically different, but nevertheless correspond
precisely in terms of what they observe. Bijective lenses are formulates using
the strong GetPut law~\cite{Foster09,Fischer2015}.›
locale bij_lens = weak_lens +
assumes strong_get_put: "put σ (get ρ) = ρ"
begin
sublocale pbij_lens
proof
fix σ v ρ
show "put σ v = put ρ v"
by (metis put_get strong_get_put)
qed
sublocale vwb_lens
proof
fix σ v u
show "put σ (get σ) = σ"
by (simp add: strong_get_put)
qed
lemma put_bij: "bij_betw (put σ) UNIV UNIV"
by (metis bijI put_inj strong_get_put surj_def)
lemma get_create: "create (get σ) = σ"
by (simp add: lens_create_def strong_get_put)
end
declare bij_lens.strong_get_put [simp]
declare bij_lens.get_create [simp]
lemma bij_lens_weak [simp]:
"bij_lens x ⟹ weak_lens x"
by (simp_all add: bij_lens_def)
lemma bij_lens_pbij [simp]:
"bij_lens x ⟹ pbij_lens x"
by (metis bij_lens.get_create bij_lens_def pbij_lens_axioms.intro pbij_lens_def weak_lens.put_get)
lemma bij_lens_vwb [simp]: "bij_lens x ⟹ vwb_lens x"
by (metis bij_lens.strong_get_put bij_lens_weak mwb_lens.intro mwb_lens_axioms.intro vwb_lens_def wb_lens.intro wb_lens_axioms.intro weak_lens.put_get)
text ‹ Alternative characterisation: a bijective lens is a partial bijective lens that is also
very well-behaved (i.e. total). ›
lemma pbij_vwb_is_bij_lens:
"⟦ pbij_lens X; vwb_lens X ⟧ ⟹ bij_lens X"
by (unfold_locales, simp_all, meson pbij_lens.put_det vwb_lens.put_eq)
lemma bij_lens_iff_pbij_vwb:
"bij_lens X ⟷ (pbij_lens X ∧ vwb_lens X)"
using pbij_vwb_is_bij_lens by auto
subsection ‹Lens Independence›
text ‹
\begin{figure}
\begin{center}
\includegraphics[width=6cm]{figures/Independence}
\end{center}
\vspace{-5ex}
\caption{Lens Independence}
\label{fig:Indep}
\end{figure}
Lens independence shows when two lenses $X$ and $Y$ characterise disjoint regions of the
source type, as illustrated in Figure~\ref{fig:Indep}. We specify this by requiring that the $\lput$
functions of the two lenses commute, and that the $\lget$ function of each lens is unaffected by
application of $\lput$ from the corresponding lens. ›
locale lens_indep =
fixes X :: "'a ⟹ 'c" and Y :: "'b ⟹ 'c"
assumes lens_put_comm: "put⇘X⇙ (put⇘Y⇙ σ v) u = put⇘Y⇙ (put⇘X⇙ σ u) v"
and lens_put_irr1: "get⇘X⇙ (put⇘Y⇙ σ v) = get⇘X⇙ σ"
and lens_put_irr2: "get⇘Y⇙ (put⇘X⇙ σ u) = get⇘Y⇙ σ"
notation lens_indep (infix "⨝" 50)
lemma lens_indepI:
"⟦ ⋀ u v σ. put⇘x⇙ (put⇘y⇙ σ v) u = put⇘y⇙ (put⇘x⇙ σ u) v;
⋀ v σ. get⇘x⇙ (put⇘y⇙ σ v) = get⇘x⇙ σ;
⋀ u σ. get⇘y⇙ (put⇘x⇙ σ u) = get⇘y⇙ σ ⟧ ⟹ x ⨝ y"
by (simp add: lens_indep_def)
text ‹Lens independence is symmetric.›
lemma lens_indep_sym: "x ⨝ y ⟹ y ⨝ x"
by (simp add: lens_indep_def)
lemma lens_indep_comm:
"x ⨝ y ⟹ put⇘x⇙ (put⇘y⇙ σ v) u = put⇘y⇙ (put⇘x⇙ σ u) v"
by (simp add: lens_indep_def)
lemma lens_indep_get [simp]:
assumes "x ⨝ y"
shows "get⇘x⇙ (put⇘y⇙ σ v) = get⇘x⇙ σ"
using assms lens_indep_def by fastforce
text ‹ Characterisation of independence for two very well-behaved lenses ›
lemma lens_indep_vwb_iff:
assumes "vwb_lens x" "vwb_lens y"
shows "x ⨝ y ⟷ (∀ u v σ. put⇘x⇙ (put⇘y⇙ σ v) u = put⇘y⇙ (put⇘x⇙ σ u) v)"
proof
assume "x ⨝ y"
thus "∀ u v σ. put⇘x⇙ (put⇘y⇙ σ v) u = put⇘y⇙ (put⇘x⇙ σ u) v"
by (simp add: lens_indep_comm)
next
assume a: "∀ u v σ. put⇘x⇙ (put⇘y⇙ σ v) u = put⇘y⇙ (put⇘x⇙ σ u) v"
show "x ⨝ y"
proof (unfold_locales)
fix σ v u
from a show "put⇘x⇙ (put⇘y⇙ σ v) u = put⇘y⇙ (put⇘x⇙ σ u) v"
by auto
show "get⇘x⇙ (put⇘y⇙ σ v) = get⇘x⇙ σ"
by (metis a assms(1) vwb_lens.put_eq vwb_lens_wb wb_lens_def weak_lens.put_get)
show "get⇘y⇙ (put⇘x⇙ σ u) = get⇘y⇙ σ"
by (metis a assms(2) vwb_lens.put_eq vwb_lens_wb wb_lens_def weak_lens.put_get)
qed
qed
subsection ‹ Lens Compatibility ›
text ‹ Lens compatibility is a weaker notion than independence. It allows that two lenses can overlap
so long as they manipulate the source in the same way in that region. It is most easily defined
in terms of a function for copying a region from one source to another using a lens. ›
definition lens_compat (infix "##⇩L" 50) where
[lens_defs]: "lens_compat X Y = (∀s⇩1 s⇩2. s⇩1 ◃⇘X⇙ s⇩2 ◃⇘Y⇙ s⇩2 = s⇩1 ◃⇘Y⇙ s⇩2 ◃⇘X⇙ s⇩2)"
lemma lens_compat_idem [simp]: "x ##⇩L x"
by (simp add: lens_defs)
lemma lens_compat_sym: "x ##⇩L y ⟹ y ##⇩L x"
by (simp add: lens_defs)
lemma lens_indep_compat [simp]: "x ⨝ y ⟹ x ##⇩L y"
by (simp add: lens_override_def lens_compat_def lens_indep_comm)
end
Theory Lens_Algebra
section ‹Lens Algebraic Operators›
theory Lens_Algebra
imports Lens_Laws
begin
subsection ‹Lens Composition, Plus, Unit, and Identity›
text ‹
\begin{figure}
\begin{center}
\includegraphics[width=7cm]{figures/Composition}
\end{center}
\vspace{-5ex}
\caption{Lens Composition}
\label{fig:Comp}
\end{figure}
We introduce the algebraic lens operators; for more information please see our paper~\cite{Foster16a}.
Lens composition, illustrated in Figure~\ref{fig:Comp}, constructs a lens by composing the source
of one lens with the view of another.›
definition lens_comp :: "('a ⟹ 'b) ⇒ ('b ⟹ 'c) ⇒ ('a ⟹ 'c)" (infixl ";⇩L" 80) where
[lens_defs]: "lens_comp Y X = ⦇ lens_get = get⇘Y⇙ ∘ lens_get X
, lens_put = (λ σ v. lens_put X σ (lens_put Y (lens_get X σ) v)) ⦈"
text ‹
\begin{figure}
\begin{center}
\includegraphics[width=7cm]{figures/Sum}
\end{center}
\vspace{-5ex}
\caption{Lens Sum}
\label{fig:Sum}
\end{figure}
Lens plus, as illustrated in Figure~\ref{fig:Sum} parallel composes two independent lenses,
resulting in a lens whose view is the product of the two underlying lens views.›
definition lens_plus :: "('a ⟹ 'c) ⇒ ('b ⟹ 'c) ⇒ 'a × 'b ⟹ 'c" (infixr "+⇩L" 75) where
[lens_defs]: "X +⇩L Y = ⦇ lens_get = (λ σ. (lens_get X σ, lens_get Y σ))
, lens_put = (λ σ (u, v). lens_put X (lens_put Y σ v) u) ⦈"
text ‹The product functor lens similarly parallel composes two lenses, but in this case the lenses
have different sources and so the resulting source is also a product.›
definition lens_prod :: "('a ⟹ 'c) ⇒ ('b ⟹ 'd) ⇒ ('a × 'b ⟹ 'c × 'd)" (infixr "×⇩L" 85) where
[lens_defs]: "lens_prod X Y = ⦇ lens_get = map_prod get⇘X⇙ get⇘Y⇙
, lens_put = λ (u, v) (x, y). (put⇘X⇙ u x, put⇘Y⇙ v y) ⦈"
text ‹The $\lfst$ and $\lsnd$ lenses project the first and second elements, respectively, of a
product source type.›
definition fst_lens :: "'a ⟹ 'a × 'b" ("fst⇩L") where
[lens_defs]: "fst⇩L = ⦇ lens_get = fst, lens_put = (λ (σ, ρ) u. (u, ρ)) ⦈"
definition snd_lens :: "'b ⟹ 'a × 'b" ("snd⇩L") where
[lens_defs]: "snd⇩L = ⦇ lens_get = snd, lens_put = (λ (σ, ρ) u. (σ, u)) ⦈"
lemma get_fst_lens [simp]: "get⇘fst⇩L⇙ (x, y) = x"
by (simp add: fst_lens_def)
lemma get_snd_lens [simp]: "get⇘snd⇩L⇙ (x, y) = y"
by (simp add: snd_lens_def)
text ‹The swap lens is a bijective lens which swaps over the elements of the product source type.›
abbreviation swap_lens :: "'a × 'b ⟹ 'b × 'a" ("swap⇩L") where
"swap⇩L ≡ snd⇩L +⇩L fst⇩L"
text ‹The zero lens is an ineffectual lens whose view is a unit type. This means the zero lens
cannot distinguish or change the source type.›
definition zero_lens :: "unit ⟹ 'a" ("0⇩L") where
[lens_defs]: "0⇩L = ⦇ lens_get = (λ _. ()), lens_put = (λ σ x. σ) ⦈"
text ‹The identity lens is a bijective lens where the source and view type are the same.›
definition id_lens :: "'a ⟹ 'a" ("1⇩L") where
[lens_defs]: "1⇩L = ⦇ lens_get = id, lens_put = (λ _. id) ⦈"
text ‹The quotient operator $X \lquot Y$ shortens lens $X$ by cutting off $Y$ from the end. It is
thus the dual of the composition operator.›
definition lens_quotient :: "('a ⟹ 'c) ⇒ ('b ⟹ 'c) ⇒ 'a ⟹ 'b" (infixr "'/⇩L" 90) where
[lens_defs]: "X /⇩L Y = ⦇ lens_get = λ σ. get⇘X⇙ (create⇘Y⇙ σ)
, lens_put = λ σ v. get⇘Y⇙ (put⇘X⇙ (create⇘Y⇙ σ) v) ⦈"
text ‹Lens inverse take a bijective lens and swaps the source and view types.›
definition lens_inv :: "('a ⟹ 'b) ⇒ ('b ⟹ 'a)" ("inv⇩L") where
[lens_defs]: "lens_inv x = ⦇ lens_get = create⇘x⇙, lens_put = λ σ. get⇘x⇙ ⦈"
subsection ‹Closure Poperties›
text ‹We show that the core lenses combinators defined above are closed under the key lens classes.›
lemma id_wb_lens: "wb_lens 1⇩L"
by (unfold_locales, simp_all add: id_lens_def)
lemma source_id_lens: "𝒮⇘1⇩L⇙ = UNIV"
by (simp add: id_lens_def lens_source_def)
lemma unit_wb_lens: "wb_lens 0⇩L"
by (unfold_locales, simp_all add: zero_lens_def)
lemma source_zero_lens: "𝒮⇘0⇩L⇙ = UNIV"
by (simp_all add: zero_lens_def lens_source_def)
lemma comp_weak_lens: "⟦ weak_lens x; weak_lens y ⟧ ⟹ weak_lens (x ;⇩L y)"
by (unfold_locales, simp_all add: lens_comp_def)
lemma comp_wb_lens: "⟦ wb_lens x; wb_lens y ⟧ ⟹ wb_lens (x ;⇩L y)"
by (unfold_locales, auto simp add: lens_comp_def wb_lens_def weak_lens.put_closure)
lemma comp_mwb_lens: "⟦ mwb_lens x; mwb_lens y ⟧ ⟹ mwb_lens (x ;⇩L y)"
by (unfold_locales, auto simp add: lens_comp_def mwb_lens_def weak_lens.put_closure)
lemma source_lens_comp: "⟦ mwb_lens x; mwb_lens y ⟧ ⟹ 𝒮⇘x ;⇩L y⇙ = {s ∈ 𝒮⇘y⇙. get⇘y⇙ s ∈ 𝒮⇘x⇙}"
by (auto simp add: lens_comp_def lens_source_def, blast, metis mwb_lens.put_put mwb_lens_def weak_lens.put_get)
lemma id_vwb_lens [simp]: "vwb_lens 1⇩L"
by (unfold_locales, simp_all add: id_lens_def)
lemma unit_vwb_lens [simp]: "vwb_lens 0⇩L"
by (unfold_locales, simp_all add: zero_lens_def)
lemma comp_vwb_lens: "⟦ vwb_lens x; vwb_lens y ⟧ ⟹ vwb_lens (x ;⇩L y)"
by (unfold_locales, simp_all add: lens_comp_def weak_lens.put_closure)
lemma unit_ief_lens: "ief_lens 0⇩L"
by (unfold_locales, simp_all add: zero_lens_def)
text ‹Lens plus requires that the lenses be independent to show closure.›
lemma plus_mwb_lens:
assumes "mwb_lens x" "mwb_lens y" "x ⨝ y"
shows "mwb_lens (x +⇩L y)"
using assms
apply (unfold_locales)
apply (simp_all add: lens_plus_def prod.case_eq_if lens_indep_sym)
apply (simp add: lens_indep_comm)
done
lemma plus_wb_lens:
assumes "wb_lens x" "wb_lens y" "x ⨝ y"
shows "wb_lens (x +⇩L y)"
using assms
apply (unfold_locales, simp_all add: lens_plus_def)
apply (simp add: lens_indep_sym prod.case_eq_if)
done
lemma plus_vwb_lens [simp]:
assumes "vwb_lens x" "vwb_lens y" "x ⨝ y"
shows "vwb_lens (x +⇩L y)"
using assms
apply (unfold_locales, simp_all add: lens_plus_def)
apply (simp add: lens_indep_sym prod.case_eq_if)
apply (simp add: lens_indep_comm prod.case_eq_if)
done
lemma source_plus_lens:
assumes "mwb_lens x" "mwb_lens y" "x ⨝ y"
shows "𝒮⇘x +⇩L y⇙ = 𝒮⇘x⇙ ∩ 𝒮⇘y⇙"
apply (auto simp add: lens_source_def lens_plus_def)
apply (meson assms(3) lens_indep_comm)
apply (metis assms(1) mwb_lens.weak_get_put mwb_lens_weak weak_lens.put_closure)
done
lemma prod_mwb_lens:
"⟦ mwb_lens X; mwb_lens Y ⟧ ⟹ mwb_lens (X ×⇩L Y)"
by (unfold_locales, simp_all add: lens_prod_def prod.case_eq_if)
lemma prod_wb_lens:
"⟦ wb_lens X; wb_lens Y ⟧ ⟹ wb_lens (X ×⇩L Y)"
by (unfold_locales, simp_all add: lens_prod_def prod.case_eq_if)
lemma prod_vwb_lens:
"⟦ vwb_lens X; vwb_lens Y ⟧ ⟹ vwb_lens (X ×⇩L Y)"
by (unfold_locales, simp_all add: lens_prod_def prod.case_eq_if)
lemma prod_bij_lens:
"⟦ bij_lens X; bij_lens Y ⟧ ⟹ bij_lens (X ×⇩L Y)"
by (unfold_locales, simp_all add: lens_prod_def prod.case_eq_if)
lemma fst_vwb_lens: "vwb_lens fst⇩L"
by (unfold_locales, simp_all add: fst_lens_def prod.case_eq_if)
lemma snd_vwb_lens: "vwb_lens snd⇩L"
by (unfold_locales, simp_all add: snd_lens_def prod.case_eq_if)
lemma id_bij_lens: "bij_lens 1⇩L"
by (unfold_locales, simp_all add: id_lens_def)
lemma inv_id_lens: "inv⇩L 1⇩L = 1⇩L"
by (auto simp add: lens_inv_def id_lens_def lens_create_def)
lemma inv_inv_lens: "bij_lens X ⟹ inv⇩L (inv⇩L X) = X"
apply (cases X)
apply (auto simp add: lens_defs fun_eq_iff)
apply (metis (no_types) bij_lens.strong_get_put bij_lens_def select_convs(2) weak_lens.put_get)
done
lemma lens_inv_bij: "bij_lens X ⟹ bij_lens (inv⇩L X)"
by (unfold_locales, simp_all add: lens_inv_def lens_create_def)
lemma swap_bij_lens: "bij_lens swap⇩L"
by (unfold_locales, simp_all add: lens_plus_def prod.case_eq_if fst_lens_def snd_lens_def)
subsection ‹Composition Laws›
text ‹Lens composition is monoidal, with unit @{term "1⇩L"}, as the following theorems demonstrate.
It also has @{term "0⇩L"} as a right annihilator. ›
lemma lens_comp_assoc: "X ;⇩L (Y ;⇩L Z) = (X ;⇩L Y) ;⇩L Z"
by (auto simp add: lens_comp_def)
lemma lens_comp_left_id [simp]: "1⇩L ;⇩L X = X"
by (simp add: id_lens_def lens_comp_def)
lemma lens_comp_right_id [simp]: "X ;⇩L 1⇩L = X"
by (simp add: id_lens_def lens_comp_def)
lemma lens_comp_anhil [simp]: "wb_lens X ⟹ 0⇩L ;⇩L X = 0⇩L"
by (simp add: zero_lens_def lens_comp_def comp_def)
lemma lens_comp_anhil_right [simp]: "wb_lens X ⟹ X ;⇩L 0⇩L = 0⇩L"
by (simp add: zero_lens_def lens_comp_def comp_def)
subsection ‹Independence Laws›
text ‹The zero lens @{term "0⇩L"} is independent of any lens. This is because nothing can be observed
or changed using @{term "0⇩L"}. ›
lemma zero_lens_indep [simp]: "0⇩L ⨝ X"
by (auto simp add: zero_lens_def lens_indep_def)
lemma zero_lens_indep' [simp]: "X ⨝ 0⇩L"
by (auto simp add: zero_lens_def lens_indep_def)
text ‹Lens independence is irreflexive, but only for effectual lenses as otherwise nothing can
be observed.›
lemma lens_indep_quasi_irrefl: "⟦ wb_lens x; eff_lens x ⟧ ⟹ ¬ (x ⨝ x)"
by (auto simp add: lens_indep_def ief_lens_def ief_lens_axioms_def, metis (full_types) wb_lens.get_put)
text ‹Lens independence is a congruence with respect to composition, as the following properties demonstrate.›
lemma lens_indep_left_comp [simp]:
"⟦ mwb_lens z; x ⨝ y ⟧ ⟹ (x ;⇩L z) ⨝ (y ;⇩L z)"
apply (rule lens_indepI)
apply (auto simp add: lens_comp_def)
apply (simp add: lens_indep_comm)
apply (simp add: lens_indep_sym)
done
lemma lens_indep_right_comp:
"y ⨝ z ⟹ (x ;⇩L y) ⨝ (x ;⇩L z)"
apply (auto intro!: lens_indepI simp add: lens_comp_def)
using lens_indep_comm lens_indep_sym apply fastforce
apply (simp add: lens_indep_sym)
done
lemma lens_indep_left_ext [intro]:
"y ⨝ z ⟹ (x ;⇩L y) ⨝ z"
apply (auto intro!: lens_indepI simp add: lens_comp_def)
apply (simp add: lens_indep_comm)
apply (simp add: lens_indep_sym)
done
lemma lens_indep_right_ext [intro]:
"x ⨝ z ⟹ x ⨝ (y ;⇩L z)"
by (simp add: lens_indep_left_ext lens_indep_sym)
lemma lens_comp_indep_cong_left:
"⟦ mwb_lens Z; X ;⇩L Z ⨝ Y ;⇩L Z ⟧ ⟹ X ⨝ Y"
apply (rule lens_indepI)
apply (rename_tac u v σ)
apply (drule_tac u=u and v=v and σ="create⇘Z⇙ σ" in lens_indep_comm)
apply (simp add: lens_comp_def)
apply (meson mwb_lens_weak weak_lens.view_determination)
apply (rename_tac v σ)
apply (drule_tac v=v and σ="create⇘Z⇙ σ" in lens_indep_get)
apply (simp add: lens_comp_def)
apply (drule lens_indep_sym)
apply (rename_tac u σ)
apply (drule_tac v=u and σ="create⇘Z⇙ σ" in lens_indep_get)
apply (simp add: lens_comp_def)
done
lemma lens_comp_indep_cong:
"mwb_lens Z ⟹ (X ;⇩L Z) ⨝ (Y ;⇩L Z) ⟷ X ⨝ Y"
using lens_comp_indep_cong_left lens_indep_left_comp by blast
text ‹The first and second lenses are independent since the view different parts of a product source.›
lemma fst_snd_lens_indep [simp]:
"fst⇩L ⨝ snd⇩L"
by (simp add: lens_indep_def fst_lens_def snd_lens_def)
lemma snd_fst_lens_indep [simp]:
"snd⇩L ⨝ fst⇩L"
by (simp add: lens_indep_def fst_lens_def snd_lens_def)
lemma split_prod_lens_indep:
assumes "mwb_lens X"
shows "(fst⇩L ;⇩L X) ⨝ (snd⇩L ;⇩L X)"
using assms fst_snd_lens_indep lens_indep_left_comp vwb_lens_mwb by blast
text ‹Lens independence is preserved by summation.›
lemma plus_pres_lens_indep [simp]: "⟦ X ⨝ Z; Y ⨝ Z ⟧ ⟹ (X +⇩L Y) ⨝ Z"
apply (rule lens_indepI)
apply (simp_all add: lens_plus_def prod.case_eq_if)
apply (simp add: lens_indep_comm)
apply (simp add: lens_indep_sym)
done
lemma plus_pres_lens_indep' [simp]:
"⟦ X ⨝ Y; X ⨝ Z ⟧ ⟹ X ⨝ Y +⇩L Z"
by (auto intro: lens_indep_sym plus_pres_lens_indep)
text ‹Lens independence is preserved by product.›
lemma lens_indep_prod:
"⟦ X⇩1 ⨝ X⇩2; Y⇩1 ⨝ Y⇩2 ⟧ ⟹ X⇩1 ×⇩L Y⇩1 ⨝ X⇩2 ×⇩L Y⇩2"
apply (rule lens_indepI)
apply (auto simp add: lens_prod_def prod.case_eq_if lens_indep_comm map_prod_def)
apply (simp_all add: lens_indep_sym)
done
subsection ‹ Compatibility Laws ›
lemma zero_lens_compat [simp]: "0⇩L ##⇩L X"
by (auto simp add: zero_lens_def lens_override_def lens_compat_def)
lemma id_lens_compat [simp]: "vwb_lens X ⟹ 1⇩L ##⇩L X"
by (auto simp add: id_lens_def lens_override_def lens_compat_def)
subsection ‹Algebraic Laws›
text ‹Lens plus distributes to the right through composition.›
lemma plus_lens_distr: "mwb_lens Z ⟹ (X +⇩L Y) ;⇩L Z = (X ;⇩L Z) +⇩L (Y ;⇩L Z)"
by (auto simp add: lens_comp_def lens_plus_def comp_def)
text ‹The first lens projects the first part of a summation.›
lemma fst_lens_plus:
"wb_lens y ⟹ fst⇩L ;⇩L (x +⇩L y) = x"
by (simp add: fst_lens_def lens_plus_def lens_comp_def comp_def)
text ‹The second law requires independence as we have to apply x first, before y›
lemma snd_lens_plus:
"⟦ wb_lens x; x ⨝ y ⟧ ⟹ snd⇩L ;⇩L (x +⇩L y) = y"
apply (simp add: snd_lens_def lens_plus_def lens_comp_def comp_def)
apply (subst lens_indep_comm)
apply (simp_all)
done
text ‹The swap lens switches over a summation.›
lemma lens_plus_swap:
"X ⨝ Y ⟹ swap⇩L ;⇩L (X +⇩L Y) = (Y +⇩L X)"
by (auto simp add: lens_plus_def fst_lens_def snd_lens_def id_lens_def lens_comp_def lens_indep_comm)
text ‹The first, second, and swap lenses are all closely related.›
lemma fst_snd_id_lens: "fst⇩L +⇩L snd⇩L = 1⇩L"
by (auto simp add: lens_plus_def fst_lens_def snd_lens_def id_lens_def)
lemma swap_lens_idem: "swap⇩L ;⇩L swap⇩L = 1⇩L"
by (simp add: fst_snd_id_lens lens_indep_sym lens_plus_swap)
lemma swap_lens_fst: "fst⇩L ;⇩L swap⇩L = snd⇩L"
by (simp add: fst_lens_plus fst_vwb_lens)
lemma swap_lens_snd: "snd⇩L ;⇩L swap⇩L = fst⇩L"
by (simp add: lens_indep_sym snd_lens_plus snd_vwb_lens)
text ‹The product lens can be rewritten as a sum lens.›
lemma prod_as_plus: "X ×⇩L Y = X ;⇩L fst⇩L +⇩L Y ;⇩L snd⇩L"
by (auto simp add: lens_prod_def fst_lens_def snd_lens_def lens_comp_def lens_plus_def)
lemma prod_lens_id_equiv:
"1⇩L ×⇩L 1⇩L = 1⇩L"
by (auto simp add: lens_prod_def id_lens_def)
lemma prod_lens_comp_plus:
"X⇩2 ⨝ Y⇩2 ⟹ ((X⇩1 ×⇩L Y⇩1) ;⇩L (X⇩2 +⇩L Y⇩2)) = (X⇩1 ;⇩L X⇩2) +⇩L (Y⇩1 ;⇩L Y⇩2)"
by (auto simp add: lens_comp_def lens_plus_def lens_prod_def prod.case_eq_if fun_eq_iff)
text ‹The following laws about quotient are similar to their arithmetic analogues. Lens quotient
reverse the effect of a composition.›
lemma lens_comp_quotient:
"weak_lens Y ⟹ (X ;⇩L Y) /⇩L Y = X"
by (simp add: lens_quotient_def lens_comp_def)
lemma lens_quotient_id [simp]: "weak_lens X ⟹ (X /⇩L X) = 1⇩L"
by (force simp add: lens_quotient_def id_lens_def)
lemma lens_quotient_id_denom: "X /⇩L 1⇩L = X"
by (simp add: lens_quotient_def id_lens_def lens_create_def)
lemma lens_quotient_unit: "weak_lens X ⟹ (0⇩L /⇩L X) = 0⇩L"
by (simp add: lens_quotient_def zero_lens_def)
lemma lens_obs_eq_zero: "s⇩1 ≃⇘0⇩L⇙ s⇩2 = (s⇩1 = s⇩2)"
by (simp add: lens_defs)
lemma lens_obs_eq_one: "s⇩1 ≃⇘1⇩L⇙ s⇩2"
by (simp add: lens_defs)
lemma lens_obs_eq_as_override: "vwb_lens X ⟹ s⇩1 ≃⇘X⇙ s⇩2 ⟷ (s⇩2 = s⇩1 ⊕⇩L s⇩2 on X)"
by (auto simp add: lens_defs; metis vwb_lens.put_eq)
end
Theory Lens_Order
section ‹Order and Equivalence on Lenses›
theory Lens_Order
imports Lens_Algebra
begin
subsection ‹Sub-lens Relation›
text ‹A lens $X$ is a sub-lens of $Y$ if there is a well-behaved lens $Z$ such that $X = Z \lcomp Y$,
or in other words if $X$ can be expressed purely in terms of $Y$.›
definition sublens :: "('a ⟹ 'c) ⇒ ('b ⟹ 'c) ⇒ bool" (infix "⊆⇩L" 55) where
[lens_defs]: "sublens X Y = (∃ Z :: ('a, 'b) lens. vwb_lens Z ∧ X = Z ;⇩L Y)"
text ‹Various lens classes are downward closed under the sublens relation.›
lemma sublens_pres_mwb:
"⟦ mwb_lens Y; X ⊆⇩L Y ⟧ ⟹ mwb_lens X"
by (unfold_locales, auto simp add: sublens_def lens_comp_def)
lemma sublens_pres_wb:
"⟦ wb_lens Y; X ⊆⇩L Y ⟧ ⟹ wb_lens X"
by (unfold_locales, auto simp add: sublens_def lens_comp_def)
lemma sublens_pres_vwb:
"⟦ vwb_lens Y; X ⊆⇩L Y ⟧ ⟹ vwb_lens X"
by (unfold_locales, auto simp add: sublens_def lens_comp_def)
text ‹Sublens is a preorder as the following two theorems show.›
lemma sublens_refl [simp]:
"X ⊆⇩L X"
using id_vwb_lens sublens_def by fastforce
lemma sublens_trans [trans]:
"⟦ X ⊆⇩L Y; Y ⊆⇩L Z ⟧ ⟹ X ⊆⇩L Z"
apply (auto simp add: sublens_def lens_comp_assoc)
apply (rename_tac Z⇩1 Z⇩2)
apply (rule_tac x="Z⇩1 ;⇩L Z⇩2" in exI)
apply (simp add: lens_comp_assoc)
using comp_vwb_lens apply blast
done
text ‹Sublens has a least element -- @{text "0⇩L"} -- and a greatest element -- @{text "1⇩L"}.
Intuitively, this shows that sublens orders how large a portion of the source type a particular
lens views, with @{text "0⇩L"} observing the least, and @{text "1⇩L"} observing the most.›
lemma sublens_least: "wb_lens X ⟹ 0⇩L ⊆⇩L X"
using sublens_def unit_vwb_lens by fastforce
lemma sublens_greatest: "vwb_lens X ⟹ X ⊆⇩L 1⇩L"
by (simp add: sublens_def)
text ‹If $Y$ is a sublens of $X$ then any put using $X$ will necessarily erase any put using $Y$.
Similarly, any two source types are observationally equivalent by $Y$ when performed
following a put using $X$.›
lemma sublens_put_put:
"⟦ mwb_lens X; Y ⊆⇩L X ⟧ ⟹ put⇘X⇙ (put⇘Y⇙ σ v) u = put⇘X⇙ σ u"
by (auto simp add: sublens_def lens_comp_def)
lemma sublens_obs_get:
"⟦ mwb_lens X; Y ⊆⇩L X ⟧ ⟹ get⇘Y⇙ (put⇘X⇙ σ v) = get⇘Y⇙ (put⇘X⇙ ρ v)"
by (auto simp add: sublens_def lens_comp_def)
text ‹Sublens preserves independence; in other words if $Y$ is independent of $Z$, then also
any $X$ smaller than $Y$ is independent of $Z$.›
lemma sublens_pres_indep:
"⟦ X ⊆⇩L Y; Y ⨝ Z ⟧ ⟹ X ⨝ Z"
apply (auto intro!:lens_indepI simp add: sublens_def lens_comp_def lens_indep_comm)
apply (simp add: lens_indep_sym)
done
lemma sublens_pres_indep':
"⟦ X ⊆⇩L Y; Z ⨝ Y ⟧ ⟹ Z ⨝ X"
by (meson lens_indep_sym sublens_pres_indep)
lemma sublens_compat: "⟦ vwb_lens X; vwb_lens Y; X ⊆⇩L Y ⟧ ⟹ X ##⇩L Y"
unfolding lens_compat_def lens_override_def
by (metis (no_types, hide_lams) sublens_obs_get sublens_put_put vwb_lens_mwb vwb_lens_wb wb_lens.get_put)
text ‹Well-behavedness of lens quotient has sublens as a proviso. This is because we can only
remove $X$ from $Y$ if $X$ is smaller than $Y$. ›
lemma lens_quotient_mwb:
"⟦ mwb_lens Y; X ⊆⇩L Y ⟧ ⟹ mwb_lens (X /⇩L Y)"
by (unfold_locales, auto simp add: lens_quotient_def lens_create_def sublens_def lens_comp_def comp_def)
subsection ‹Lens Equivalence›
text ‹Using our preorder, we can also derive an equivalence on lenses as follows. It should be
noted that this equality, like sublens, is heterogeneously typed -- it can compare lenses whose
view types are different, so long as the source types are the same. We show that it is reflexive,
symmetric, and transitive. ›
definition lens_equiv :: "('a ⟹ 'c) ⇒ ('b ⟹ 'c) ⇒ bool" (infix "≈⇩L" 51) where
[lens_defs]: "lens_equiv X Y = (X ⊆⇩L Y ∧ Y ⊆⇩L X)"
lemma lens_equivI [intro]:
"⟦ X ⊆⇩L Y; Y ⊆⇩L X ⟧ ⟹ X ≈⇩L Y"
by (simp add: lens_equiv_def)
lemma lens_equiv_refl:
"X ≈⇩L X"
by (simp add: lens_equiv_def)
lemma lens_equiv_sym:
"X ≈⇩L Y ⟹ Y ≈⇩L X"
by (simp add: lens_equiv_def)
lemma lens_equiv_trans [trans]:
"⟦ X ≈⇩L Y; Y ≈⇩L Z ⟧ ⟹ X ≈⇩L Z"
by (auto intro: sublens_trans simp add: lens_equiv_def)
lemma lens_equiv_pres_indep:
"⟦ X ≈⇩L Y; Y ⨝ Z ⟧ ⟹ X ⨝ Z"
using lens_equiv_def sublens_pres_indep by blast
lemma lens_equiv_pres_indep':
"⟦ X ≈⇩L Y; Z ⨝ Y ⟧ ⟹ Z ⨝ X"
using lens_equiv_def sublens_pres_indep' by blast
lemma lens_comp_cong_1: "X ≈⇩L Y ⟹ X ;⇩L Z ≈⇩L Y ;⇩L Z"
unfolding lens_equiv_def
by (metis (no_types, lifting) lens_comp_assoc sublens_def)
subsection ‹Further Algebraic Laws›
text ‹This law explains the behaviour of lens quotient.›
lemma lens_quotient_comp:
"⟦ weak_lens Y; X ⊆⇩L Y ⟧ ⟹ (X /⇩L Y) ;⇩L Y = X"
by (auto simp add: lens_quotient_def lens_comp_def comp_def sublens_def)
text ‹Plus distributes through quotient.›
lemma lens_quotient_plus:
"⟦ mwb_lens Z; X ⊆⇩L Z; Y ⊆⇩L Z ⟧ ⟹ (X +⇩L Y) /⇩L Z = (X /⇩L Z) +⇩L (Y /⇩L Z)"
apply (auto simp add: lens_quotient_def lens_plus_def sublens_def lens_comp_def comp_def)
apply (rule ext)
apply (rule ext)
apply (simp add: prod.case_eq_if)
done
text ‹Laws for for lens plus on the denominator. These laws allow us to extract compositions
of @{term "fst⇩L"} and @{term "snd⇩L"} terms. ›
lemma lens_quotient_plus_den1:
"⟦ weak_lens x; weak_lens y; x ⨝ y ⟧ ⟹ x /⇩L (x +⇩L y) = fst⇩L"
by (auto simp add: lens_defs prod.case_eq_if fun_eq_iff, metis (lifting) lens_indep_def weak_lens.put_get)
lemma lens_quotient_plus_den2: "⟦ weak_lens x; weak_lens z; x ⨝ z; y ⊆⇩L z ⟧ ⟹ y /⇩L (x +⇩L z) = (y /⇩L z) ;⇩L snd⇩L "
by (auto simp add: lens_defs prod.case_eq_if fun_eq_iff lens_indep.lens_put_irr2)
text ‹There follows a number of laws relating sublens and summation. Firstly, sublens is preserved
by summation. ›
lemma plus_pred_sublens: "⟦ mwb_lens Z; X ⊆⇩L Z; Y ⊆⇩L Z; X ⨝ Y ⟧ ⟹ (X +⇩L Y) ⊆⇩L Z"
apply (auto simp add: sublens_def)
apply (rename_tac Z⇩1 Z⇩2)
apply (rule_tac x="Z⇩1 +⇩L Z⇩2" in exI)
apply (auto intro!: plus_wb_lens)
apply (simp add: lens_comp_indep_cong_left plus_vwb_lens)
apply (simp add: plus_lens_distr)
done
text ‹Intuitively, lens plus is associative. However we cannot prove this using HOL equality due
to monomorphic typing of this operator. But since sublens and lens equivalence are both heterogeneous
we can now prove this in the following three lemmas.›
lemma lens_plus_sub_assoc_1:
"X +⇩L Y +⇩L Z ⊆⇩L (X +⇩L Y) +⇩L Z"
apply (simp add: sublens_def)
apply (rule_tac x="(fst⇩L ;⇩L fst⇩L) +⇩L (snd⇩L ;⇩L fst⇩L) +⇩L snd⇩L" in exI)
apply (auto)
apply (rule plus_vwb_lens)
apply (simp add: comp_vwb_lens fst_vwb_lens)
apply (rule plus_vwb_lens)
apply (simp add: comp_vwb_lens fst_vwb_lens snd_vwb_lens)
apply (simp add: snd_vwb_lens)
apply (simp add: lens_indep_left_ext)
apply (rule lens_indep_sym)
apply (rule plus_pres_lens_indep)
using fst_snd_lens_indep fst_vwb_lens lens_indep_left_comp lens_indep_sym vwb_lens_mwb apply blast
using fst_snd_lens_indep lens_indep_left_ext lens_indep_sym apply blast
apply (auto simp add: lens_plus_def lens_comp_def fst_lens_def snd_lens_def prod.case_eq_if split_beta')[1]
done
lemma lens_plus_sub_assoc_2:
"(X +⇩L Y) +⇩L Z ⊆⇩L X +⇩L Y +⇩L Z"
apply (simp add: sublens_def)
apply (rule_tac x="(fst⇩L +⇩L (fst⇩L ;⇩L snd⇩L)) +⇩L (snd⇩L ;⇩L snd⇩L)" in exI)
apply (auto)
apply (rule plus_vwb_lens)
apply (rule plus_vwb_lens)
apply (simp add: fst_vwb_lens)
apply (simp add: comp_vwb_lens fst_vwb_lens snd_vwb_lens)
apply (rule lens_indep_sym)
apply (rule lens_indep_left_ext)
using fst_snd_lens_indep lens_indep_sym apply blast
apply (auto intro: comp_vwb_lens simp add: snd_vwb_lens)
apply (rule plus_pres_lens_indep)
apply (simp add: lens_indep_left_ext lens_indep_sym)
apply (simp add: snd_vwb_lens)
apply (auto simp add: lens_plus_def lens_comp_def fst_lens_def snd_lens_def prod.case_eq_if split_beta')[1]
done
lemma lens_plus_assoc:
"(X +⇩L Y) +⇩L Z ≈⇩L X +⇩L Y +⇩L Z"
by (simp add: lens_equivI lens_plus_sub_assoc_1 lens_plus_sub_assoc_2)
text ‹We can similarly show that it is commutative.›
lemma lens_plus_sub_comm: "X ⨝ Y ⟹ X +⇩L Y ⊆⇩L Y +⇩L X"
apply (simp add: sublens_def)
apply (rule_tac x="snd⇩L +⇩L fst⇩L" in exI)
apply (auto)
apply (simp add: fst_vwb_lens lens_indep_sym snd_vwb_lens)
apply (simp add: lens_indep_sym lens_plus_swap)
done
lemma lens_plus_comm: "X ⨝ Y ⟹ X +⇩L Y ≈⇩L Y +⇩L X"
by (simp add: lens_equivI lens_indep_sym lens_plus_sub_comm)
text ‹Any composite lens is larger than an element of the lens, as demonstrated by the following
four laws.›
lemma lens_plus_ub [simp]: "wb_lens Y ⟹ X ⊆⇩L X +⇩L Y"
by (metis fst_lens_plus fst_vwb_lens sublens_def)
lemma lens_plus_right_sublens:
"⟦ vwb_lens Y; Y ⨝ Z; X ⊆⇩L Z ⟧ ⟹ X ⊆⇩L Y +⇩L Z"
apply (auto simp add: sublens_def)
apply (rename_tac Z')
apply (rule_tac x="Z' ;⇩L snd⇩L" in exI)
apply (auto)
using comp_vwb_lens snd_vwb_lens apply blast
apply (metis lens_comp_assoc snd_lens_plus vwb_lens_def)
done
lemma lens_plus_mono_left:
"⟦ Y ⨝ Z; X ⊆⇩L Y ⟧ ⟹ X +⇩L Z ⊆⇩L Y +⇩L Z"
apply (auto simp add: sublens_def)
apply (rename_tac Z')
apply (rule_tac x="Z' ×⇩L 1⇩L" in exI)
apply (subst prod_lens_comp_plus)
apply (simp_all)
using id_vwb_lens prod_vwb_lens apply blast
done
lemma lens_plus_mono_right:
"⟦ X ⨝ Z; Y ⊆⇩L Z ⟧ ⟹ X +⇩L Y ⊆⇩L X +⇩L Z"
by (metis prod_lens_comp_plus prod_vwb_lens sublens_def sublens_refl)
text ‹If we compose a lens $X$ with lens $Y$ then naturally the resulting lens must be smaller
than $Y$, as $X$ views a part of $Y$. ›
lemma lens_comp_lb [simp]: "vwb_lens X ⟹ X ;⇩L Y ⊆⇩L Y"
by (auto simp add: sublens_def)
lemma sublens_comp [simp]:
assumes "vwb_lens b" "c ⊆⇩L a"
shows "(b ;⇩L c) ⊆⇩L a"
by (metis assms sublens_def sublens_trans)
text ‹We can now also show that @{text "0⇩L"} is the unit of lens plus›
lemma lens_unit_plus_sublens_1: "X ⊆⇩L 0⇩L +⇩L X"
by (metis lens_comp_lb snd_lens_plus snd_vwb_lens zero_lens_indep unit_wb_lens)
lemma lens_unit_prod_sublens_2: "0⇩L +⇩L X ⊆⇩L X"
apply (auto simp add: sublens_def)
apply (rule_tac x="0⇩L +⇩L 1⇩L" in exI)
apply (auto)
apply (auto simp add: lens_plus_def zero_lens_def lens_comp_def id_lens_def prod.case_eq_if comp_def)
apply (rule ext)
apply (rule ext)
apply (auto)
done
lemma lens_plus_left_unit: "0⇩L +⇩L X ≈⇩L X"
by (simp add: lens_equivI lens_unit_plus_sublens_1 lens_unit_prod_sublens_2)
lemma lens_plus_right_unit: "X +⇩L 0⇩L ≈⇩L X"
using lens_equiv_trans lens_indep_sym lens_plus_comm lens_plus_left_unit zero_lens_indep by blast
text ‹We can also show that both sublens and equivalence are congruences with respect to lens plus
and lens product.›
lemma lens_plus_subcong: "⟦ Y⇩1 ⨝ Y⇩2; X⇩1 ⊆⇩L Y⇩1; X⇩2 ⊆⇩L Y⇩2 ⟧ ⟹ X⇩1 +⇩L X⇩2 ⊆⇩L Y⇩1 +⇩L Y⇩2"
by (metis prod_lens_comp_plus prod_vwb_lens sublens_def)
lemma lens_plus_eq_left: "⟦ X ⨝ Z; X ≈⇩L Y ⟧ ⟹ X +⇩L Z ≈⇩L Y +⇩L Z"
by (meson lens_equiv_def lens_plus_mono_left sublens_pres_indep)
lemma lens_plus_eq_right: "⟦ X ⨝ Y; Y ≈⇩L Z ⟧ ⟹ X +⇩L Y ≈⇩L X +⇩L Z"
by (meson lens_equiv_def lens_indep_sym lens_plus_mono_right sublens_pres_indep)
lemma lens_plus_cong:
assumes "X⇩1 ⨝ X⇩2" "X⇩1 ≈⇩L Y⇩1" "X⇩2 ≈⇩L Y⇩2"
shows "X⇩1 +⇩L X⇩2 ≈⇩L Y⇩1 +⇩L Y⇩2"
proof -
have "X⇩1 +⇩L X⇩2 ≈⇩L Y⇩1 +⇩L X⇩2"
by (simp add: assms(1) assms(2) lens_plus_eq_left)
moreover have "... ≈⇩L Y⇩1 +⇩L Y⇩2"
using assms(1) assms(2) assms(3) lens_equiv_def lens_plus_eq_right sublens_pres_indep by blast
ultimately show ?thesis
using lens_equiv_trans by blast
qed
lemma prod_lens_sublens_cong:
"⟦ X⇩1 ⊆⇩L X⇩2; Y⇩1 ⊆⇩L Y⇩2 ⟧ ⟹ (X⇩1 ×⇩L Y⇩1) ⊆⇩L (X⇩2 ×⇩L Y⇩2)"
apply (auto simp add: sublens_def)
apply (rename_tac Z⇩1 Z⇩2)
apply (rule_tac x="Z⇩1 ×⇩L Z⇩2" in exI)
apply (auto)
using prod_vwb_lens apply blast
apply (auto simp add: lens_prod_def lens_comp_def prod.case_eq_if)
apply (rule ext, rule ext)
apply (auto simp add: lens_prod_def lens_comp_def prod.case_eq_if)
done
lemma prod_lens_equiv_cong:
"⟦ X⇩1 ≈⇩L X⇩2; Y⇩1 ≈⇩L Y⇩2 ⟧ ⟹ (X⇩1 ×⇩L Y⇩1) ≈⇩L (X⇩2 ×⇩L Y⇩2)"
by (simp add: lens_equiv_def prod_lens_sublens_cong)
text ‹We also have the following "exchange" law that allows us to switch over a lens product and plus.›
lemma lens_plus_prod_exchange:
"(X⇩1 +⇩L X⇩2) ×⇩L (Y⇩1 +⇩L Y⇩2) ≈⇩L (X⇩1 ×⇩L Y⇩1) +⇩L (X⇩2 ×⇩L Y⇩2)"
proof (rule lens_equivI)
show "(X⇩1 +⇩L X⇩2) ×⇩L (Y⇩1 +⇩L Y⇩2) ⊆⇩L (X⇩1 ×⇩L Y⇩1) +⇩L (X⇩2 ×⇩L Y⇩2)"
apply (simp add: sublens_def)
apply (rule_tac x="((fst⇩L ;⇩L fst⇩L) +⇩L (fst⇩L ;⇩L snd⇩L)) +⇩L ((snd⇩L ;⇩L fst⇩L) +⇩L (snd⇩L ;⇩L snd⇩L))" in exI)
apply (auto)
apply (auto intro!: plus_vwb_lens comp_vwb_lens fst_vwb_lens snd_vwb_lens lens_indep_right_comp)
apply (auto intro!: lens_indepI simp add: lens_comp_def lens_plus_def fst_lens_def snd_lens_def)
apply (auto simp add: lens_prod_def lens_plus_def lens_comp_def fst_lens_def snd_lens_def prod.case_eq_if comp_def)[1]
apply (rule ext, rule ext, auto simp add: prod.case_eq_if)
done
show "(X⇩1 ×⇩L Y⇩1) +⇩L (X⇩2 ×⇩L Y⇩2) ⊆⇩L (X⇩1 +⇩L X⇩2) ×⇩L (Y⇩1 +⇩L Y⇩2)"
apply (simp add: sublens_def)
apply (rule_tac x="((fst⇩L ;⇩L fst⇩L) +⇩L (fst⇩L ;⇩L snd⇩L)) +⇩L ((snd⇩L ;⇩L fst⇩L) +⇩L (snd⇩L ;⇩L snd⇩L))" in exI)
apply (auto)
apply (auto intro!: plus_vwb_lens comp_vwb_lens fst_vwb_lens snd_vwb_lens lens_indep_right_comp)
apply (auto intro!: lens_indepI simp add: lens_comp_def lens_plus_def fst_lens_def snd_lens_def)
apply (auto simp add: lens_prod_def lens_plus_def lens_comp_def fst_lens_def snd_lens_def prod.case_eq_if comp_def)[1]
apply (rule ext, rule ext, auto simp add: lens_prod_def prod.case_eq_if)
done
qed
lemma lens_get_put_quasi_commute:
"⟦ vwb_lens Y; X ⊆⇩L Y ⟧ ⟹ get⇘Y⇙ (put⇘X⇙ s v) = put⇘X /⇩L Y⇙ (get⇘Y⇙ s) v"
proof -
assume a1: "vwb_lens Y"
assume a2: "X ⊆⇩L Y"
have "⋀l la. put⇘l ;⇩L la⇙ = (λb c. put⇘la⇙ (b::'b) (put⇘l⇙ (get⇘la⇙ b::'a) (c::'c)))"
by (simp add: lens_comp_def)
then have "⋀l la b c. get⇘l⇙ (put⇘la ;⇩L l⇙ (b::'b) (c::'c)) = put⇘la⇙ (get⇘l⇙ b::'a) c ∨ ¬ weak_lens l"
by force
then show ?thesis
using a2 a1 by (metis lens_quotient_comp vwb_lens_wb wb_lens_def)
qed
lemma lens_put_of_quotient:
"⟦ vwb_lens Y; X ⊆⇩L Y ⟧ ⟹ put⇘Y⇙ s (put⇘X /⇩L Y⇙ v⇩2 v⇩1) = put⇘X⇙ (put⇘Y⇙ s v⇩2) v⇩1"
proof -
assume a1: "vwb_lens Y"
assume a2: "X ⊆⇩L Y"
have f3: "⋀l b. put⇘l⇙ (b::'b) (get⇘l⇙ b::'a) = b ∨ ¬ vwb_lens l"
by force
have f4: "⋀b c. put⇘X /⇩L Y⇙ (get⇘Y⇙ b) c = get⇘Y⇙ (put⇘X⇙ b c)"
using a2 a1 by (simp add: lens_get_put_quasi_commute)
have "⋀b c a. put⇘Y⇙ (put⇘X⇙ b c) a = put⇘Y⇙ b a"
using a2 a1 by (simp add: sublens_put_put)
then show ?thesis
using f4 f3 a1 by (metis mwb_lens.put_put mwb_lens_def vwb_lens_mwb weak_lens.put_get)
qed
subsection ‹Bijective Lens Equivalences›
text ‹A bijective lens, like a bijective function, is its own inverse. Thus, if we compose its inverse
with itself we get @{term "1⇩L"}.›
lemma bij_lens_inv_left:
"bij_lens X ⟹ inv⇩L X ;⇩L X = 1⇩L"
by (auto simp add: lens_inv_def lens_comp_def lens_create_def comp_def id_lens_def, rule ext, auto)
lemma bij_lens_inv_right:
"bij_lens X ⟹ X ;⇩L inv⇩L X = 1⇩L"
by (auto simp add: lens_inv_def lens_comp_def comp_def id_lens_def, rule ext, auto)
text ‹The following important results shows that bijective lenses are precisely those that are
equivalent to identity. In other words, a bijective lens views all of the source type.›
lemma bij_lens_equiv_id:
"bij_lens X ⟷ X ≈⇩L 1⇩L"
apply (auto)
apply (rule lens_equivI)
apply (simp_all add: sublens_def)
apply (rule_tac x="lens_inv X" in exI)
apply (simp add: bij_lens_inv_left lens_inv_bij)
apply (auto simp add: lens_equiv_def sublens_def id_lens_def lens_comp_def comp_def)
apply (unfold_locales)
apply (simp)
apply (simp)
apply (metis (no_types, lifting) vwb_lens_wb wb_lens_weak weak_lens.put_get)
done
text ‹For this reason, by transitivity, any two bijective lenses with the same source type must be equivalent.›
lemma bij_lens_equiv:
"⟦ bij_lens X; X ≈⇩L Y ⟧ ⟹ bij_lens Y"
by (meson bij_lens_equiv_id lens_equiv_def sublens_trans)
lemma bij_lens_cong:
"X ≈⇩L Y ⟹ bij_lens X = bij_lens Y"
by (meson bij_lens_equiv lens_equiv_sym)
text ‹We can also show that the identity lens @{term "1⇩L"} is unique. That is to say it is the only
lens which when compose with $Y$ will yield $Y$.›
lemma lens_id_unique:
"weak_lens Y ⟹ Y = X ;⇩L Y ⟹ X = 1⇩L"
apply (cases Y)
apply (cases X)
apply (auto simp add: lens_comp_def comp_def id_lens_def fun_eq_iff)
apply (metis select_convs(1) weak_lens.create_get)
apply (metis select_convs(1) select_convs(2) weak_lens.put_get)
done
text ‹Consequently, if composition of two lenses $X$ and $Y$ yields @{text "1⇩L"}, then both
of the composed lenses must be bijective.›
lemma bij_lens_via_comp_id_left:
"⟦ wb_lens X; wb_lens Y; X ;⇩L Y = 1⇩L ⟧ ⟹ bij_lens X"
apply (cases Y)
apply (cases X)
apply (auto simp add: lens_comp_def comp_def id_lens_def fun_eq_iff)
apply (unfold_locales)
apply (simp_all)
using vwb_lens_wb wb_lens_weak weak_lens.put_get apply fastforce
apply (metis select_convs(1) select_convs(2) wb_lens_weak weak_lens.put_get)
done
lemma bij_lens_via_comp_id_right:
"⟦ wb_lens X; wb_lens Y; X ;⇩L Y = 1⇩L ⟧ ⟹ bij_lens Y"
apply (cases Y)
apply (cases X)
apply (auto simp add: lens_comp_def comp_def id_lens_def fun_eq_iff)
apply (unfold_locales)
apply (simp_all)
using vwb_lens_wb wb_lens_weak weak_lens.put_get apply fastforce
apply (metis select_convs(1) select_convs(2) wb_lens_weak weak_lens.put_get)
done
text ‹Importantly, an equivalence between two lenses can be demonstrated by showing that one lens
can be converted to the other by application of a suitable bijective lens $Z$. This $Z$ lens
converts the view type of one to the view type of the other.›
lemma lens_equiv_via_bij:
assumes "bij_lens Z" "X = Z ;⇩L Y"
shows "X ≈⇩L Y"
using assms
apply (auto simp add: lens_equiv_def sublens_def)
using bij_lens_vwb apply blast
apply (rule_tac x="lens_inv Z" in exI)
apply (auto simp add: lens_comp_assoc bij_lens_inv_left)
using bij_lens_vwb lens_inv_bij apply blast
done
text ‹Indeed, we actually have a stronger result than this -- the equivalent lenses are precisely
those than can be converted to one another through a suitable bijective lens. Bijective lenses
can thus be seen as a special class of "adapter" lens.›
lemma lens_equiv_iff_bij:
assumes "weak_lens Y"
shows "X ≈⇩L Y ⟷ (∃ Z. bij_lens Z ∧ X = Z ;⇩L Y)"
apply (rule iffI)
apply (auto simp add: lens_equiv_def sublens_def lens_id_unique)[1]
apply (rename_tac Z⇩1 Z⇩2)
apply (rule_tac x="Z⇩1" in exI)
apply (simp)
apply (subgoal_tac "Z⇩2 ;⇩L Z⇩1 = 1⇩L")
apply (meson bij_lens_via_comp_id_right vwb_lens_wb)
apply (metis assms lens_comp_assoc lens_id_unique)
using lens_equiv_via_bij apply blast
done
lemma pbij_plus_commute:
"⟦ a ⨝ b; mwb_lens a; mwb_lens b; pbij_lens (b +⇩L a) ⟧ ⟹ pbij_lens (a +⇩L b)"
apply (unfold_locales, simp_all add:lens_defs lens_indep_sym prod.case_eq_if)
using lens_indep.lens_put_comm pbij_lens.put_det apply fastforce
done
subsection ‹Lens Override Laws›
text ‹The following laws are analogus to the equivalent laws for functions.›
lemma lens_override_id [simp]:
"S⇩1 ⊕⇩L S⇩2 on 1⇩L = S⇩2"
by (simp add: lens_override_def id_lens_def)
lemma lens_override_unit [simp]:
"S⇩1 ⊕⇩L S⇩2 on 0⇩L = S⇩1"
by (simp add: lens_override_def zero_lens_def)
lemma lens_override_overshadow:
assumes "mwb_lens Y" "X ⊆⇩L Y"
shows "(S⇩1 ⊕⇩L S⇩2 on X) ⊕⇩L S⇩3 on Y = S⇩1 ⊕⇩L S⇩3 on Y"
using assms by (simp add: lens_override_def sublens_put_put)
lemma lens_override_irr:
assumes "X ⨝ Y"
shows "S⇩1 ⊕⇩L (S⇩2 ⊕⇩L S⇩3 on Y) on X = S⇩1 ⊕⇩L S⇩2 on X"
using assms by (simp add: lens_override_def)
lemma lens_override_overshadow_left:
assumes "mwb_lens X"
shows "(S⇩1 ⊕⇩L S⇩2 on X) ⊕⇩L S⇩3 on X = S⇩1 ⊕⇩L S⇩3 on X"
by (simp add: assms lens_override_def)
lemma lens_override_overshadow_right:
assumes "mwb_lens X"
shows "S⇩1 ⊕⇩L (S⇩2 ⊕⇩L S⇩3 on X) on X = S⇩1 ⊕⇩L S⇩3 on X"
by (simp add: assms lens_override_def)
lemma lens_override_plus:
"X ⨝ Y ⟹ S⇩1 ⊕⇩L S⇩2 on (X +⇩L Y) = (S⇩1 ⊕⇩L S⇩2 on X) ⊕⇩L S⇩2 on Y"
by (simp add: lens_indep_comm lens_override_def lens_plus_def)
lemma lens_override_idem [simp]:
"vwb_lens X ⟹ S ⊕⇩L S on X = S"
by (simp add: lens_override_def)
lemma lens_override_mwb_idem [simp]:
"⟦ mwb_lens X; S ∈ 𝒮⇘X⇙ ⟧ ⟹ S ⊕⇩L S on X = S"
by (simp add: lens_override_def)
lemma lens_override_put_right_in:
"⟦ vwb_lens A; X ⊆⇩L A ⟧ ⟹ S⇩1 ⊕⇩L (put⇘X⇙ S⇩2 v) on A = put⇘X⇙ (S⇩1 ⊕⇩L S⇩2 on A) v"
by (simp add: lens_override_def lens_get_put_quasi_commute lens_put_of_quotient)
lemma lens_override_put_right_out:
"⟦ vwb_lens A; X ⨝ A ⟧ ⟹ S⇩1 ⊕⇩L (put⇘X⇙ S⇩2 v) on A = (S⇩1 ⊕⇩L S⇩2 on A)"
by (simp add: lens_override_def lens_indep.lens_put_irr2)
lemma lens_indep_overrideI:
assumes "vwb_lens X" "vwb_lens Y" "(⋀ s⇩1 s⇩2 s⇩3. s⇩1 ⊕⇩L s⇩2 on X ⊕⇩L s⇩3 on Y = s⇩1 ⊕⇩L s⇩3 on Y ⊕⇩L s⇩2 on X)"
shows "X ⨝ Y"
using assms
apply (unfold_locales)
apply (simp_all add: lens_override_def)
apply (metis mwb_lens_def vwb_lens_mwb weak_lens.put_get)
apply (metis lens_override_def lens_override_idem mwb_lens_def vwb_lens_mwb weak_lens.put_get)
apply (metis mwb_lens_weak vwb_lens_mwb vwb_lens_wb wb_lens.get_put weak_lens.put_get)
done
lemma lens_indep_override_def:
assumes "vwb_lens X" "vwb_lens Y"
shows "X ⨝ Y ⟷ (∀ s⇩1 s⇩2 s⇩3. s⇩1 ⊕⇩L s⇩2 on X ⊕⇩L s⇩3 on Y = s⇩1 ⊕⇩L s⇩3 on Y ⊕⇩L s⇩2 on X)"
by (metis assms(1) assms(2) lens_indep_comm lens_indep_overrideI lens_override_def)
text ‹ Alternative characterisation of very-well behaved lenses: override is idempotent. ›
lemma override_idem_implies_vwb:
"⟦ mwb_lens X; ⋀ s. s ⊕⇩L s on X = s ⟧ ⟹ vwb_lens X"
by (unfold_locales, simp_all add: lens_defs)
subsection ‹ Alternative Sublens Characterisation ›
text ‹ The following definition is equivalent to the above when the two lenses are very well behaved. ›
definition sublens' :: "('a ⟹ 'c) ⇒ ('b ⟹ 'c) ⇒ bool" (infix "⊆⇩L''" 55) where
[lens_defs]: "sublens' X Y = (∀ s⇩1 s⇩2 s⇩3. s⇩1 ⊕⇩L s⇩2 on Y ⊕⇩L s⇩3 on X = s⇩1 ⊕⇩L s⇩2 ⊕⇩L s⇩3 on X on Y)"
text ‹ We next prove some characteristic properties of our alternative definition of sublens. ›
lemma sublens'_prop1:
assumes "vwb_lens X" "X ⊆⇩L' Y"
shows "put⇘X⇙ (put⇘Y⇙ s⇩1 (get⇘Y⇙ s⇩2)) s⇩3 = put⇘Y⇙ s⇩1 (get⇘Y⇙ (put⇘X⇙ s⇩2 s⇩3))"
using assms
by (simp add: sublens'_def, metis lens_override_def mwb_lens_def vwb_lens_mwb weak_lens.put_get)
lemma sublens'_prop2:
assumes "vwb_lens X" "X ⊆⇩L' Y"
shows "get⇘X⇙ (put⇘Y⇙ s⇩1 (get⇘Y⇙ s⇩2)) = get⇘X⇙ s⇩2"
using assms unfolding sublens'_def
by (metis lens_override_def vwb_lens_wb wb_lens_axioms_def wb_lens_def weak_lens.put_get)
lemma sublens'_prop3:
assumes "vwb_lens X" "vwb_lens Y" "X ⊆⇩L' Y"
shows "put⇘Y⇙ σ (get⇘Y⇙ (put⇘X⇙ (put⇘Y⇙ ρ (get⇘Y⇙ σ)) v)) = put⇘X⇙ σ v"
by (metis assms(1) assms(2) assms(3) mwb_lens_def sublens'_prop1 vwb_lens.put_eq vwb_lens_mwb weak_lens.put_get)
text ‹ Finally we show our two definitions of sublens are equivalent, assuming very well behaved lenses. ›
lemma sublens'_implies_sublens:
assumes "vwb_lens X" "vwb_lens Y" "X ⊆⇩L' Y"
shows "X ⊆⇩L Y"
proof -
have "vwb_lens (X /⇩L Y)"
by (unfold_locales
,auto simp add: assms lens_quotient_def lens_comp_def lens_create_def sublens'_prop1 sublens'_prop2)
moreover have "X = X /⇩L Y ;⇩L Y"
proof -
have "get⇘X⇙ = (λσ. get⇘X⇙ (create⇘Y⇙ σ)) ∘ get⇘Y⇙"
by (rule ext, simp add: assms(1) assms(3) lens_create_def sublens'_prop2)
moreover have "put⇘X⇙ = (λσ v. put⇘Y⇙ σ (get⇘Y⇙ (put⇘X⇙ (create⇘Y⇙ (get⇘Y⇙ σ)) v)))"
by (rule ext, rule ext, simp add: assms(1) assms(2) assms(3) lens_create_def sublens'_prop3)
ultimately show ?thesis
by (simp add: lens_quotient_def lens_comp_def)
qed
ultimately show ?thesis
using sublens_def by blast
qed
lemma sublens_implies_sublens':
assumes "vwb_lens Y" "X ⊆⇩L Y"
shows "X ⊆⇩L' Y"
by (metis assms lens_override_def lens_override_put_right_in sublens'_def)
lemma sublens_iff_sublens':
assumes "vwb_lens X" "vwb_lens Y"
shows "X ⊆⇩L Y ⟷ X ⊆⇩L' Y"
using assms sublens'_implies_sublens sublens_implies_sublens' by blast
subsection ‹ Alternative Equivalence Characterisation ›
definition lens_equiv' :: "('a ⟹ 'c) ⇒ ('b ⟹ 'c) ⇒ bool" (infix "≈⇩L''" 51) where
[lens_defs]: "lens_equiv' X Y = (∀ s⇩1 s⇩2. (s⇩1 ⊕⇩L s⇩2 on X = s⇩1 ⊕⇩L s⇩2 on Y))"
lemma lens_equiv_iff_lens_equiv':
assumes "vwb_lens X" "vwb_lens Y"
shows "X ≈⇩L Y ⟷ X ≈⇩L' Y"
apply (simp add: lens_equiv_def sublens_iff_sublens' assms)
apply (auto simp add: lens_defs assms)
apply (metis assms(2) mwb_lens.put_put vwb_lens_mwb vwb_lens_wb wb_lens.get_put)
done
end
Theory Lens_Symmetric
section ‹ Symmetric Lenses ›
theory Lens_Symmetric
imports Lens_Order
begin
text ‹ A characterisation of Hofmann's ``Symmetric Lenses''~\cite{Hofmann2011}, where
a lens is accompanied by its complement. ›
record ('a, 'b, 's) slens =
view :: "'a ⟹ 's" ("𝒱ı")
coview :: "'b ⟹ 's" ("𝒞ı")
type_notation
slens ("<_, _> ⟺ _" [0, 0, 0] 0)
declare slens.defs [lens_defs]
definition slens_compl :: "(<'a, 'c> ⟺ 'b) ⇒ <'c, 'a> ⟺ 'b" ("-⇩L _" [81] 80) where
[lens_defs]: "slens_compl a = ⦇ view = coview a, coview = view a ⦈"
lemma view_slens_compl [simp]: "𝒱⇘-⇩L a⇙ = 𝒞⇘a⇙"
by (simp add: slens_compl_def)
lemma coview_slens_compl [simp]: "𝒞⇘-⇩L a⇙ = 𝒱⇘a⇙"
by (simp add: slens_compl_def)
subsection ‹ Partial Symmetric Lenses ›
locale psym_lens =
fixes S :: "<'a, 'b> ⟺ 's" (structure)
assumes
mwb_region [simp]: "mwb_lens 𝒱" and
mwb_coregion [simp]: "mwb_lens 𝒞" and
indep_region_coregion [simp]: "𝒱 ⨝ 𝒞" and
pbij_region_coregion [simp]: "pbij_lens (𝒱 +⇩L 𝒞)"
declare psym_lens.mwb_region [simp]
declare psym_lens.mwb_coregion [simp]
declare psym_lens.indep_region_coregion [simp]
lemma psym_lens_compl [simp]: "psym_lens a ⟹ psym_lens (-⇩L a)"
apply (simp add: slens_compl_def)
apply (rule psym_lens.intro)
apply (simp_all)
using lens_indep_sym psym_lens.indep_region_coregion apply blast
using lens_indep_sym pbij_plus_commute psym_lens_def apply blast
done
subsection ‹ Symmetric Lenses ›
locale sym_lens =
fixes S :: "<'a, 'b> ⟺ 's" (structure)
assumes
vwb_region: "vwb_lens 𝒱" and
vwb_coregion: "vwb_lens 𝒞" and
indep_region_coregion: "𝒱 ⨝ 𝒞" and
bij_region_coregion: "bij_lens (𝒱 +⇩L 𝒞)"
begin
sublocale psym_lens
proof (rule psym_lens.intro)
show "mwb_lens 𝒱"
by (simp add: vwb_region)
show "mwb_lens 𝒞"
by (simp add: vwb_coregion)
show "𝒱 ⨝ 𝒞"
using indep_region_coregion by blast
show "pbij_lens (𝒱 +⇩L 𝒞)"
by (simp add: bij_region_coregion)
qed
lemma put_region_coregion_cover:
"put⇘𝒱⇙ (put⇘𝒞⇙ s⇩1 (get⇘𝒞⇙ s⇩2)) (get⇘𝒱⇙ s⇩2) = s⇩2"
proof -
have "put⇘𝒱⇙ (put⇘𝒞⇙ s⇩1 (get⇘𝒞⇙ s⇩2)) (get⇘𝒱⇙ s⇩2) = put⇘𝒱 +⇩L 𝒞⇙ s⇩1 (get⇘𝒱 +⇩L 𝒞⇙ s⇩2)"
by (simp add: lens_defs)
also have "... = s⇩2"
by (simp add: bij_region_coregion)
finally show ?thesis .
qed
end
declare sym_lens.vwb_region [simp]
declare sym_lens.vwb_coregion [simp]
declare sym_lens.indep_region_coregion [simp]
lemma sym_lens_psym [simp]: "sym_lens x ⟹ psym_lens x"
by (simp add: psym_lens_def sym_lens.bij_region_coregion)
lemma sym_lens_compl [simp]: "sym_lens a ⟹ sym_lens (-⇩L a)"
apply (simp add: slens_compl_def)
apply (rule sym_lens.intro, simp_all)
using lens_indep_sym sym_lens.indep_region_coregion apply blast
using bij_lens_equiv lens_plus_comm sym_lens_def apply blast
done
end
Theory Lens_Instances
section ‹Lens Instances›
theory Lens_Instances
imports Lens_Order Lens_Symmetric "HOL-Eisbach.Eisbach"
keywords "alphabet" "statespace" :: "thy_defn"
begin
text ‹In this section we define a number of concrete instantiations of the lens locales, including
functions lenses, list lenses, and record lenses.›
subsection ‹Function Lens›
text ‹A function lens views the valuation associated with a particular domain element @{typ "'a"}.
We require that range type of a lens function has cardinality of at least 2; this ensures
that properties of independence are provable.›
definition fun_lens :: "'a ⇒ ('b::two ⟹ ('a ⇒ 'b))" where
[lens_defs]: "fun_lens x = ⦇ lens_get = (λ f. f x), lens_put = (λ f u. f(x := u)) ⦈"
lemma fun_vwb_lens: "vwb_lens (fun_lens x)"
by (unfold_locales, simp_all add: fun_lens_def)
text ‹Two function lenses are independent if and only if the domain elements are different.›
lemma fun_lens_indep:
"fun_lens x ⨝ fun_lens y ⟷ x ≠ y"
proof -
obtain u v :: "'a::two" where "u ≠ v"
using two_diff by auto
thus ?thesis
by (auto simp add: fun_lens_def lens_indep_def)
qed
subsection ‹Function Range Lens›
text ‹The function range lens allows us to focus on a particular region of a function's range.›
definition fun_ran_lens :: "('c ⟹ 'b) ⇒ (('a ⇒ 'b) ⟹ 'α) ⇒ (('a ⇒ 'c) ⟹ 'α)" where
[lens_defs]: "fun_ran_lens X Y = ⦇ lens_get = λ s. get⇘X⇙ ∘ get⇘Y⇙ s
, lens_put = λ s v. put⇘Y⇙ s (λ x::'a. put⇘X⇙ (get⇘Y⇙ s x) (v x)) ⦈"
lemma fun_ran_mwb_lens: "⟦ mwb_lens X; mwb_lens Y ⟧ ⟹ mwb_lens (fun_ran_lens X Y)"
by (unfold_locales, auto simp add: fun_ran_lens_def)
lemma fun_ran_wb_lens: "⟦ wb_lens X; wb_lens Y ⟧ ⟹ wb_lens (fun_ran_lens X Y)"
by (unfold_locales, auto simp add: fun_ran_lens_def)
lemma fun_ran_vwb_lens: "⟦ vwb_lens X; vwb_lens Y ⟧ ⟹ vwb_lens (fun_ran_lens X Y)"
by (unfold_locales, auto simp add: fun_ran_lens_def)
subsection ‹Map Lens›
text ‹The map lens allows us to focus on a particular region of a partial function's range. It
is only a mainly well-behaved lens because it does not satisfy the PutGet law when the view
is not in the domain.›
definition map_lens :: "'a ⇒ ('b ⟹ ('a ⇀ 'b))" where
[lens_defs]: "map_lens x = ⦇ lens_get = (λ f. the (f x)), lens_put = (λ f u. f(x ↦ u)) ⦈"
lemma map_mwb_lens: "mwb_lens (map_lens x)"
by (unfold_locales, simp_all add: map_lens_def)
lemma source_map_lens: "𝒮⇘map_lens x⇙ = {f. x ∈ dom(f)}"
by (force simp add: map_lens_def lens_source_def)
subsection ‹List Lens›
text ‹The list lens allows us to view a particular element of a list. In order to show it is mainly
well-behaved we need to define to additional list functions. The following function adds
a number undefined elements to the end of a list.›
definition list_pad_out :: "'a list ⇒ nat ⇒ 'a list" where
"list_pad_out xs k = xs @ replicate (k + 1 - length xs) undefined"
text ‹The following function is like @{term "list_update"} but it adds additional elements to the
list if the list isn't long enough first.›
definition list_augment :: "'a list ⇒ nat ⇒ 'a ⇒ 'a list" where
"list_augment xs k v = (list_pad_out xs k)[k := v]"
text ‹The following function is like @{term nth} but it expressly returns @{term undefined} when
the list isn't long enough.›
definition nth' :: "'a list ⇒ nat ⇒ 'a" where
"nth' xs i = (if (length xs > i) then xs ! i else undefined)"
text ‹We can prove some additional laws about list update and append.›
lemma list_update_append_lemma1: "i < length xs ⟹ xs[i := v] @ ys = (xs @ ys)[i := v]"
by (simp add: list_update_append)
lemma list_update_append_lemma2: "i < length ys ⟹ xs @ ys[i := v] = (xs @ ys)[i + length xs := v]"
by (simp add: list_update_append)
text ‹We can also prove some laws about our new operators.›
lemma nth'_0 [simp]: "nth' (x # xs) 0 = x"
by (simp add: nth'_def)
lemma nth'_Suc [simp]: "nth' (x # xs) (Suc n) = nth' xs n"
by (simp add: nth'_def)
lemma list_augment_0 [simp]:
"list_augment (x # xs) 0 y = y # xs"
by (simp add: list_augment_def list_pad_out_def)
lemma list_augment_Suc [simp]:
"list_augment (x # xs) (Suc n) y = x # list_augment xs n y"
by (simp add: list_augment_def list_pad_out_def)
lemma list_augment_twice:
"list_augment (list_augment xs i u) j v = (list_pad_out xs (max i j))[i:=u, j:=v]"
apply (auto simp add: list_augment_def list_pad_out_def list_update_append_lemma1 replicate_add[THEN sym] max_def)
apply (metis Suc_le_mono add.commute diff_diff_add diff_le_mono le_add_diff_inverse2)
done
lemma list_augment_last [simp]:
"list_augment (xs @ [y]) (length xs) z = xs @ [z]"
by (induct xs, simp_all)
lemma list_augment_idem [simp]:
"i < length xs ⟹ list_augment xs i (xs ! i) = xs"
by (simp add: list_augment_def list_pad_out_def)
text ‹We can now prove that @{term list_augment} is commutative for different (arbitrary) indices.›
lemma list_augment_commute:
"i ≠ j ⟹ list_augment (list_augment σ j v) i u = list_augment (list_augment σ i u) j v"
by (simp add: list_augment_twice list_update_swap max.commute)
text ‹We can also prove that we can always retrieve an element we have added to the list, since
@{term list_augment} extends the list when necessary. This isn't true of @{term list_update}. ›
lemma nth_list_augment: "list_augment xs k v ! k = v"
by (simp add: list_augment_def list_pad_out_def)
lemma nth'_list_augment: "nth' (list_augment xs k v) k = v"
by (auto simp add: nth'_def nth_list_augment list_augment_def list_pad_out_def)
text ‹ The length is expanded if not already long enough, or otherwise left as it is. ›
lemma length_list_augment_1: "k ≥ length xs ⟹ length (list_augment xs k v) = Suc k"
by (simp add: list_augment_def list_pad_out_def)
lemma length_list_augment_2: "k < length xs ⟹ length (list_augment xs k v) = length xs"
by (simp add: list_augment_def list_pad_out_def)
text ‹We also have it that @{term list_augment} cancels itself.›
lemma list_augment_same_twice: "list_augment (list_augment xs k u) k v = list_augment xs k v"
by (simp add: list_augment_def list_pad_out_def)
lemma nth'_list_augment_diff: "i ≠ j ⟹ nth' (list_augment σ i v) j = nth' σ j"
by (auto simp add: list_augment_def list_pad_out_def nth_append nth'_def)
text ‹Finally we can create the list lenses, of which there are three varieties. One that allows
us to view an index, one that allows us to view the head, and one that allows us to view the tail.
They are all mainly well-behaved lenses.›
definition list_lens :: "nat ⇒ ('a::two ⟹ 'a list)" where
[lens_defs]: "list_lens i = ⦇ lens_get = (λ xs. nth' xs i)
, lens_put = (λ xs x. list_augment xs i x) ⦈"
abbreviation hd_lens ("hd⇩L") where "hd_lens ≡ list_lens 0"
definition tl_lens :: "'a list ⟹ 'a list" ("tl⇩L") where
[lens_defs]: "tl_lens = ⦇ lens_get = (λ xs. tl xs)
, lens_put = (λ xs xs'. hd xs # xs') ⦈"
lemma list_mwb_lens: "mwb_lens (list_lens x)"
by (unfold_locales, simp_all add: list_lens_def nth'_list_augment list_augment_same_twice)
text ‹ The set of constructible sources is precisely those where the length is greater than the
given index. ›
lemma source_list_lens: "𝒮⇘list_lens i⇙ = {xs. length xs > i}"
apply (auto simp add: lens_source_def list_lens_def)
apply (metis length_list_augment_1 length_list_augment_2 lessI not_less)
apply (metis list_augment_idem)
done
lemma tail_lens_mwb:
"mwb_lens tl⇩L"
by (unfold_locales, simp_all add: tl_lens_def)
lemma source_tail_lens: "𝒮⇘tl⇩L⇙ = {xs. xs ≠ []}"
using list.exhaust_sel by (auto simp add: tl_lens_def lens_source_def)
text ‹Independence of list lenses follows when the two indices are different.›
lemma list_lens_indep:
"i ≠ j ⟹ list_lens i ⨝ list_lens j"
by (simp add: list_lens_def lens_indep_def list_augment_commute nth'_list_augment_diff)
lemma hd_tl_lens_indep [simp]:
"hd⇩L ⨝ tl⇩L"
apply (rule lens_indepI)
apply (simp_all add: list_lens_def tl_lens_def)
apply (metis hd_conv_nth hd_def length_greater_0_conv list.case(1) nth'_def nth'_list_augment)
apply (metis (full_types) hd_conv_nth hd_def length_greater_0_conv list.case(1) nth'_def)
apply (metis One_nat_def diff_Suc_Suc diff_zero length_0_conv length_list_augment_1 length_tl linorder_not_less list.exhaust list.sel(2) list.sel(3) list_augment_0 not_less_zero)
done
lemma hd_tl_lens_pbij: "pbij_lens (hd⇩L +⇩L tl⇩L)"
by (unfold_locales, auto simp add: lens_defs)
subsection ‹Record Field Lenses›
text ‹We also add support for record lenses. Every record created can yield a lens for each field.
These cannot be created generically and thus must be defined case by case as new records are
created. We thus create a new Isabelle outer syntax command \textbf{alphabet} which enables this.
We first create syntax that allows us to obtain a lens from a given field using the internal
record syntax translations.›
abbreviation (input) "fld_put f ≡ (λ σ u. f (λ_. u) σ)"
syntax
"_FLDLENS" :: "id ⇒ logic" ("FLDLENS _")
translations
"FLDLENS x" => "⦇ lens_get = x, lens_put = CONST fld_put (_update_name x) ⦈"
text ‹ We also allow the extraction of the "base lens", which characterises all the fields added
by a record without the extension. ›
syntax
"_BASELENS" :: "id ⇒ logic" ("BASELENS _")
abbreviation (input) "base_lens t e m ≡ ⦇ lens_get = t, lens_put = λ s v. e v (m s) ⦈"
ML ‹
fun baselens_tr [Free (name, _)] =
let
val extend = Free (name ^ ".extend", dummyT);
val truncate = Free (name ^ ".truncate", dummyT);
val more = Free (name ^ ".more", dummyT);
in
Const (@{const_syntax "base_lens"}, dummyT) $ truncate $ extend $ more
end
| baselens_tr _ = raise Match;
›
parse_translation ‹[(@{syntax_const "_BASELENS"}, K baselens_tr)]›
text ‹We also introduce the \textbf{alphabet} command that creates a record with lenses for each field.
For each field a lens is created together with a proof that it is very well-behaved, and for each
pair of lenses an independence theorem is generated. Alphabets can also be extended which yields
sublens proofs between the extension field lens and record extension lenses. ›
ML_file ‹Lens_Lib.ML›
ML_file ‹Lens_Record.ML›
text ‹The following theorem attribute stores splitting theorems for alphabet types which which is useful
for proof automation.›
named_theorems alpha_splits
subsection ‹Locale State Spaces ›
text ‹ Alternative to the alphabet command, we also introduce the statespace command, which
implements Schirmer and Wenzel's locale-based approach to state space modelling~\cite{Schirmer2009}.
It has the advantage of allowing multiple inheritance of state spaces, and also variable names are
fully internalised with the locales. The approach is also far simpler than record-based state
spaces.
It has the disadvantage that variables may not be fully polymorphic, unlike for the
alphabet command. This makes it in general unsuitable for UTP theory alphabets. ›
ML_file ‹Lens_Statespace.ML›
subsection ‹Type Definition Lens›
text ‹ Every type defined by a ❙‹typedef› command induces a partial bijective lens constructed
using the abstraction and representation functions. ›
context type_definition
begin
definition typedef_lens :: "'b ⟹ 'a" ("typedef⇩L") where
[lens_defs]: "typedef⇩L = ⦇ lens_get = Abs, lens_put = (λ s. Rep) ⦈"
lemma pbij_typedef_lens [simp]: "pbij_lens typedef⇩L"
by (unfold_locales, simp_all add: lens_defs Rep_inverse)
lemma source_typedef_lens: "𝒮⇘typedef⇩L⇙ = A"
using Rep_cases by (auto simp add: lens_source_def lens_defs Rep)
lemma bij_typedef_lens_UNIV: "A = UNIV ⟹ bij_lens typedef⇩L"
by (auto intro: pbij_vwb_is_bij_lens simp add: mwb_UNIV_src_is_vwb_lens source_typedef_lens)
end
subsection ‹Mapper Lenses›
definition lmap_lens ::
"(('α ⇒ 'β) ⇒ ('γ ⇒ 'δ)) ⇒
(('β ⇒ 'α) ⇒ 'δ ⇒ 'γ) ⇒
('γ ⇒ 'α) ⇒
('β ⟹ 'α) ⇒
('δ ⟹ 'γ)" where
[lens_defs]:
"lmap_lens f g h l = ⦇
lens_get = f (get⇘l⇙),
lens_put = g o (put⇘l⇙) o h ⦈"
text ‹
The parse translation below yields a heterogeneous mapping lens for any
record type. This achieved through the utility function above that
constructs a functorial lens. This takes as input a heterogeneous mapping
function that lifts a function on a record's extension type to an update
on the entire record, and also the record's ``more'' function. The first input
is given twice as it has different polymorphic types, being effectively
a type functor construction which are not explicitly supported by HOL. We note
that the ‹more_update› function does something similar to the extension lifting,
but is not precisely suitable here since it only considers homogeneous functions,
namely of type ‹'a ⇒ 'a› rather than ‹'a ⇒ 'b›.
›
syntax
"_lmap" :: "id ⇒ logic" ("lmap[_]")
ML ‹
fun lmap_tr [Free (name, _)] =
let
val extend = Free (name ^ ".extend", dummyT);
val truncate = Free (name ^ ".truncate", dummyT);
val more = Free (name ^ ".more", dummyT);
val map_ext = Abs ("f", dummyT,
Abs ("r", dummyT,
extend $ (truncate $ Bound 0) $ (Bound 1 $ (more $ (Bound 0)))))
in
Const (@{const_syntax "lmap_lens"}, dummyT) $ map_ext $ map_ext $ more
end
| lmap_tr _ = raise Match;
›
parse_translation ‹[(@{syntax_const "_lmap"}, K lmap_tr)]›
subsection ‹Lens Interpretation›
named_theorems lens_interp_laws
locale lens_interp = interp
begin
declare meta_interp_law [lens_interp_laws]
declare all_interp_law [lens_interp_laws]
declare exists_interp_law [lens_interp_laws]
end
subsection ‹ Tactic ›
text ‹ A simple tactic for simplifying lens expressions ›
declare split_paired_all [alpha_splits]
method lens_simp = (simp add: alpha_splits lens_defs prod.case_eq_if)
end
File ‹Lens_Lib.ML›
signature LENS_LIB =
sig
val bij_lensN: string
val vwb_lensN: string
val sym_lensN: string
val lens_indepN: string
val lens_plusN: string
val lens_quotientN: string
val lens_compN: string
val id_lensN: string
val sublensN: string
val lens_equivN: string
val lens_defsN: string
val indepsN: string
val sublensesN: string
val quotientsN: string
val compositionsN: string
val lensT: typ -> typ -> typ
val isLensT: typ -> bool
val astateT: typ
val pairsWith: 'a list -> 'a list -> ('a * 'a) list
val pairings: 'a list -> ('a * 'a) list
val mk_vwb_lens: term -> term
val mk_indep: term -> term -> term
end
structure Lens_Lib : LENS_LIB =
struct
open Syntax
open HOLogic
val bij_lensN = @{const_name bij_lens}
val vwb_lensN = @{const_name vwb_lens}
val sym_lensN = @{const_name sym_lens}
val lens_indepN = @{const_name lens_indep}
val lens_plusN = @{const_name lens_plus}
val lens_quotientN = @{const_name lens_quotient}
val lens_compN = @{const_name lens_comp}
val id_lensN = @{const_name id_lens}
val sublensN = @{const_name sublens}
val lens_equivN = @{const_name lens_equiv}
val lens_defsN = "lens_defs"
val indepsN = "indeps"
val sublensesN = "sublenses"
val quotientsN = "quotients"
val compositionsN = "compositions"
fun lensT a b = Type (@{type_name lens_ext}, [a, b, unitT])
fun isLensT (Type (n, _)) = (n = @{type_name lens_ext}) |
isLensT _ = false
val astateT = (TFree ("'st", ["HOL.type"]))
fun pairWith _ [] = []
| pairWith x (y :: ys) = [(x, y), (y, x)] @ pairWith x ys;
fun pairsWith [] _ = []
| pairsWith (x :: xs) ys = pairWith x ys @ pairsWith xs ys;
fun pairings [] = []
| pairings (x :: xs) = pairWith x xs @ pairings xs;
fun mk_vwb_lens t = mk_Trueprop (const vwb_lensN $ t)
fun mk_indep x y = mk_Trueprop (const lens_indepN $ x $ y)
end
File ‹Lens_Record.ML›
signature LENS_UTILS =
sig
val add_alphabet_cmd : {overloaded: bool} ->
(string * string option) list * binding ->
string option -> (binding * string * mixfix) list -> theory -> theory
end;
structure Lens_Utils : LENS_UTILS =
struct
open Syntax;
open Lens_Lib;
val FLDLENS = "FLDLENS"
val BASELENS = "BASELENS"
val base_lensN = "base⇩L"
val child_lensN = "more⇩L"
val all_lensN = "all⇩L"
val base_moreN = "base_more"
val bij_lens_suffix = "_bij_lens"
val vwb_lens_suffix = "_vwb_lens"
val sym_lens_suffix = "_sym_lens"
val Trueprop = @{const_name Trueprop}
val HOLeq = @{const_name HOL.eq}
val lens_suffix = "⇩v"
val lens_defsN = "lens_defs"
val lens_defs = (Binding.empty, [Token.make_src (lens_defsN, Position.none) []])
val alpha_splitsN = "alpha_splits"
val alpha_splits = [Token.make_src (alpha_splitsN, Position.none) []]
val equivN = "equivs"
val splits_suffix = ".splits"
val defs_suffix = ".defs"
val prod_case_eq_ifN = "prod.case_eq_if"
val slens_view = "view"
val slens_coview = "coview"
fun read_parent NONE ctxt = (NONE, ctxt)
| read_parent (SOME raw_T) ctxt =
(case Proof_Context.read_typ_abbrev ctxt raw_T of
Type (name, Ts) => (SOME (Ts, name), fold Variable.declare_typ Ts ctxt)
| T => error ("Bad parent record specification: " ^ Syntax.string_of_typ ctxt T));
fun read_fields raw_fields ctxt =
let
val Ts = Syntax.read_typs ctxt (map (fn (_, raw_T, _) => raw_T) raw_fields);
val fields = map2 (fn (x, _, mx) => fn T => (x, T, mx)) raw_fields Ts;
val ctxt' = fold Variable.declare_typ Ts ctxt;
in (fields, ctxt') end;
fun add_record_cmd overloaded (raw_params, binding) raw_parent raw_fields thy =
let
val ctxt = Proof_Context.init_global thy;
val params = map (apsnd (Typedecl.read_constraint ctxt)) raw_params;
val ctxt1 = fold (Variable.declare_typ o TFree) params ctxt;
val (parent, ctxt2) = read_parent raw_parent ctxt1;
val (fields, ctxt3) = read_fields raw_fields ctxt2;
val params' = map (Proof_Context.check_tfree ctxt3) params;
in thy |> Record.add_record overloaded (params', binding) parent fields end;
fun lens_proof tname x thy =
let open Simplifier; open Global_Theory in
Goal.prove_global thy [] []
(hd (Type_Infer_Context.infer_types
(Proof_Context.init_global thy)
[mk_vwb_lens (const (Context.theory_name thy ^ "." ^ tname ^ "." ^ x))]))
(fn {context = context, prems = _}
=> EVERY [ Locale.intro_locales_tac {strict = true, eager = true} context []
, PARALLEL_ALLGOALS (asm_simp_tac
(fold add_simp (get_thm thy (x ^ "_def") :: get_thms thy (tname ^ ".defs"))
context))])
end
fun lens_sym_proof tname thy =
let open Simplifier; open Global_Theory in
Goal.prove_global thy [] []
(hd (Type_Infer_Context.infer_types
(Proof_Context.init_global thy)
[ Const (Trueprop, dummyT)
$ (const sym_lensN $ const (Context.theory_name thy ^ "." ^ tname ^ "." ^ all_lensN))]))
(fn {context = context, prems = _}
=> EVERY [ Classical.rule_tac context [@{thm sym_lens.intro}] [] 1
, PARALLEL_ALLGOALS (asm_simp_tac
(fold add_simp (@{thms slens.defs} @ get_thms thy (tname ^ ".defs"))
context))])
end
fun prove_lens_goal tname thy ctx =
let open Simplifier; open Global_Theory in
auto_tac (fold add_simp (get_thms thy lens_defsN @
get_thms thy (tname ^ splits_suffix) @
[get_thm thy prod_case_eq_ifN]) ctx)
end
fun prove_indep tname thy =
let open Simplifier; open Global_Theory in
(fn {context, prems = _} =>
EVERY [auto_tac (add_simp @{thm lens_indep_vwb_iff} context)
,prove_lens_goal tname thy context])
end
fun prove_sublens tname thy =
let open Simplifier; open Global_Theory in
(fn {context, prems = _} =>
EVERY [auto_tac (add_simp @{thm sublens_iff_sublens'} context)
,prove_lens_goal tname thy context])
end
fun prove_quotient tname thy =
let open Simplifier; open Global_Theory in
(fn {context, prems = _} =>
EVERY [prove_lens_goal tname thy context])
end
fun prove_equiv tname thy =
let open Simplifier; open Global_Theory in
(fn {context, prems = _} =>
EVERY [auto_tac (add_simp @{thm lens_equiv_iff_lens_equiv'} context)
,prove_lens_goal tname thy context])
end
fun lens_bij_proof tname thy =
let open Simplifier; open Global_Theory in
Goal.prove_global thy [] []
(hd (Type_Infer_Context.infer_types
(Proof_Context.init_global thy)
[ Const (Trueprop, dummyT)
$ (const (bij_lensN) $
(const (lens_plusN) $ const (Context.theory_name thy ^ "." ^ tname ^ "." ^ base_lensN)
$ const (Context.theory_name thy ^ "." ^ tname ^ "." ^ child_lensN)))]))
(fn {context = context, prems = _}
=> EVERY [ Locale.intro_locales_tac {strict = true, eager = true} context []
, auto_tac (fold add_simp (get_thms thy lens_defsN @ [get_thm thy prod_case_eq_ifN])
context)])
end
fun indep_proof tname thy (x, y) =
let open Simplifier; open Global_Theory in
Goal.prove_global thy [] []
(hd (Type_Infer_Context.infer_types
(Proof_Context.init_global thy)
[ mk_indep
(const (Context.theory_name thy ^ "." ^ tname ^ "." ^ x))
(const (Context.theory_name thy ^ "." ^ tname ^ "." ^ y))
]))
(prove_indep tname thy)
end
fun equiv_more_proof tname pname thy fs =
let open Simplifier; open Global_Theory; open Context; open Term in
Goal.prove_global thy [] []
(hd (Type_Infer_Context.infer_types
(Proof_Context.init_global thy)
[ Const (Trueprop, dummyT)
$ ( Const (lens_equivN, dummyT)
$ Const (pname ^ "." ^ child_lensN, dummyT)
$ foldr1 (fn (x, y) => Const (lens_plusN, dummyT) $ x $ y)
(map (fn n => Const (theory_name thy ^ "." ^ tname ^ "." ^ n, dummyT)) (fs @ [child_lensN]))
)]))
(prove_equiv tname thy)
end
fun equiv_base_proof tname parent thy fs =
let open Simplifier; open Global_Theory; open Context; open Term in
Goal.prove_global thy [] []
(hd (Type_Infer_Context.infer_types
(Proof_Context.init_global thy)
[ Const (Trueprop, dummyT)
$ ( Const (lens_equivN, dummyT)
$ Const (theory_name thy ^ "." ^ tname ^ "." ^ base_lensN, dummyT)
$ foldr1 (fn (x, y) => Const (lens_plusN, dummyT) $ x $ y)
((case parent of NONE => [] | SOME (_, pname) => [Const (pname ^ "." ^ base_lensN, dummyT)]) @
map (fn n => Const (theory_name thy ^ "." ^ tname ^ "." ^ n, dummyT)) fs)
)]))
(prove_equiv tname thy)
end
fun equiv_partition_proof tname thy =
let open Simplifier; open Global_Theory; open Context; open Term in
Goal.prove_global thy [] []
(hd (Type_Infer_Context.infer_types
(Proof_Context.init_global thy)
[ Const (Trueprop, dummyT)
$ ( Const (lens_equivN, dummyT)
$ ( Const (lens_plusN, dummyT)
$ Const (theory_name thy ^ "." ^ tname ^ "." ^ base_lensN, dummyT)
$ Const (theory_name thy ^ "." ^ tname ^ "." ^ child_lensN, dummyT))
$ Const (id_lensN, dummyT)
)]))
(prove_equiv tname thy)
end
fun sublens_proof tname pname thy y x =
let open Simplifier; open Global_Theory in
Goal.prove_global thy [] []
(hd (Type_Infer_Context.infer_types
(Proof_Context.init_global thy)
[ Const (Trueprop, dummyT)
$ ( Const (sublensN, dummyT)
$ Const (Context.theory_name thy ^ "." ^ tname ^ "." ^ x, dummyT)
$ Const (pname ^ "." ^ y, dummyT)
)]))
(prove_sublens tname thy)
end
fun quotient_proof tname thy x =
let open Simplifier; open Global_Theory in
Goal.prove_global thy [] []
(hd (Type_Infer_Context.infer_types
(Proof_Context.init_global thy)
[ Const (Trueprop, dummyT)
$ ( Const (HOLeq, dummyT)
$ (Const (lens_quotientN, dummyT)
$ Const (Context.theory_name thy ^ "." ^ tname ^ "." ^ x, dummyT)
$ Const (Context.theory_name thy ^ "." ^ tname ^ "." ^ base_lensN, dummyT)
)
$ Const (Context.theory_name thy ^ "." ^ tname ^ "." ^ x, dummyT)
)]))
(prove_quotient tname thy)
end
fun composition_proof tname thy x =
let open Simplifier; open Global_Theory in
Goal.prove_global thy [] []
(hd (Type_Infer_Context.infer_types
(Proof_Context.init_global thy)
[ Const (Trueprop, dummyT)
$ ( Const (HOLeq, dummyT)
$ (Const (lens_compN, dummyT)
$ Const (Context.theory_name thy ^ "." ^ tname ^ "." ^ x, dummyT)
$ Const (Context.theory_name thy ^ "." ^ tname ^ "." ^ base_lensN, dummyT)
)
$ Const (Context.theory_name thy ^ "." ^ tname ^ "." ^ x, dummyT)
)]))
(prove_quotient tname thy)
end
fun add_alphabet_cmd _ (raw_params, binding) raw_parent raw_fields thy =
let
open Simplifier; open Global_Theory
val tname = Binding.name_of binding
val fields = map (fn (x, y, z) => (Binding.suffix_name lens_suffix x, y, z)) raw_fields
val lnames = map (fn (x, _, _) => Binding.name_of x) raw_fields
val (parent, _) = read_parent raw_parent (Proof_Context.init_global thy);
fun ldef x = (x, x ^ " = " ^ FLDLENS ^ " " ^ x ^ lens_suffix)
val pname = case parent of SOME (_,r) => r | NONE => "";
val plchild =
case parent of
SOME (_, _) => child_lensN |
NONE => ""
val bldef = (base_lensN, base_lensN ^ " = " ^ BASELENS ^ " " ^ tname);
val mldef = (child_lensN, child_lensN ^ " = " ^ FLDLENS ^ " more");
val sldef = (all_lensN, all_lensN ^ " ≡ ⦇ " ^ slens_view ^ " = " ^ base_lensN ^ ", " ^ slens_coview ^ " = " ^ child_lensN ^ " ⦈");
val plnames = if (raw_parent = NONE) then [] else lnames @ [child_lensN];
fun pindeps thy = map (fn thm => @{thm sublens_pres_indep} OF [thm]) (get_thms thy sublensesN)
@ map (fn thm => @{thm sublens_pres_indep'} OF [thm]) (get_thms thy sublensesN)
in thy
|> add_record_cmd {overloaded = false} (raw_params, binding) raw_parent fields
|> Named_Target.theory_map (snd o Specification.theorems_cmd "" [((Binding.empty, []), [(Facts.named (tname ^ defs_suffix), snd lens_defs)])] [] false)
|> Named_Target.theory_map (snd o Specification.theorems_cmd "" [((Binding.empty, []), [(Facts.named (tname ^ splits_suffix), alpha_splits)])] [] false)
|> Sign.qualified_path false binding
|> Named_Target.theory_map
(fold (fn (n, d) => snd o Specification.definition_cmd (SOME (Binding.make (n, Position.none), NONE, NoSyn)) [] [] (lens_defs, d) true) (map ldef lnames @ [bldef, mldef]))
|> Named_Target.theory_map
(fold (fn (n, d) => Specification.abbreviation_cmd Syntax.mode_default (SOME (Binding.make (n, Position.none), NONE, NoSyn)) [] d true) [sldef])
|> fold (fn x => fn thy => snd (add_thm ((Binding.make (x ^ vwb_lens_suffix, Position.none), lens_proof tname x thy), [Simplifier.simp_add]) thy)) (lnames @ [base_lensN, child_lensN])
|> (fn thy => snd (add_thm ((Binding.make (base_moreN ^ bij_lens_suffix, Position.none), lens_bij_proof tname thy), [Simplifier.simp_add]) thy))
|> (fn thy => snd (add_thmss [((Binding.make (sublensesN, Position.none), map (sublens_proof tname pname thy plchild) plnames @ map (sublens_proof tname (Context.theory_name thy ^ "." ^ tname) thy base_lensN) lnames), [Simplifier.simp_add])] thy))
|> (fn thy => snd (add_thmss [((Binding.make (quotientsN, Position.none), map (quotient_proof tname thy) lnames), [Simplifier.simp_add])] thy))
|> (fn thy => snd (add_thmss [((Binding.make (compositionsN, Position.none), map (composition_proof tname thy) lnames), [Simplifier.simp_add])] thy))
|> (fn thy => snd (add_thmss
[((Binding.make (indepsN, Position.none), map (indep_proof tname thy) (pairings (lnames @ [child_lensN]) @ pairings [base_lensN, child_lensN]) @ pindeps thy), [Simplifier.simp_add])] thy))
|> (fn thy => snd (add_thmss
[((Binding.make (equivN, Position.none), (if (raw_parent = NONE) then [] else [equiv_more_proof tname pname thy lnames]) @ [equiv_base_proof tname parent thy lnames, equiv_partition_proof tname thy]), [Simplifier.simp_add])] thy))
|> (fn thy => snd (add_thm ((Binding.make (all_lensN ^ sym_lens_suffix, Position.none), lens_sym_proof tname thy), [Simplifier.simp_add]) thy))
|> Sign.parent_path
end;
val _ =
Outer_Syntax.command @{command_keyword alphabet} "define record with lenses"
(Parse_Spec.overloaded -- (Parse.type_args_constrained -- Parse.binding) --
(@{keyword "="} |-- Scan.option (Parse.typ --| @{keyword "+"}) --
Scan.repeat1 Parse.const_binding)
>> (fn ((overloaded, x), (y, z)) =>
Toplevel.theory (add_alphabet_cmd {overloaded = overloaded} x y z)));
end
File ‹Lens_Statespace.ML›
signature LENS_STATESPACE =
sig
val mk_statespace_elements: bstring -> (bstring * typ) list -> xstring list -> theory -> Element.context_i list
val compile_statespace: bstring -> (bstring * typ) list -> xstring list -> theory -> theory
val statespace_cmd: bstring -> xstring list -> (bstring * string) list -> theory -> theory
end
structure Lens_Statespace =
struct
fun mk_statespace_elements vds exts thy =
let
open Lens_Lib;
open Syntax;
val ctx = Named_Target.theory_init thy
val vns = map fst vds
val vfixes = map (fn (n, t) => (Binding.name n, SOME (lensT t astateT), NoSyn)) vds
val vwbs = map (mk_vwb_lens o free) vns;
val simp = Attrib.check_src ctx (Token.make_src ("simp", Position.none) [])
val imps = map (fn e => Locale.check thy (e, Position.none)) exts
val ilnsls = map (filter (isLensT o snd) o map fst o Locale.params_of thy) imps
val impind = List.concat (map (pairsWith vns) (map (map fst) ilnsls))
val indeps = map (fn (x, y) => mk_indep (free x) (free y))
(pairings (vns) @ impind)
in
[ Element.Fixes vfixes
, Element.Assumes [((Binding.name "vwbs", [simp])
, map (fn vwb => (vwb, [])) vwbs),
((Binding.name "indeps", [simp])
, map (fn vwb => (vwb, [])) indeps)]
]
end
fun compile_statespace ssn vds exts thy =
let
open Lens_Lib;
open Syntax;
val imps = map (fn e => Locale.check thy (e, Position.none)) exts
val locexs = map (fn i => (i, (("", false), (Expression.Positional [], [])))) imps
in
(Local_Theory.exit_global o snd o
Expression.add_locale (Binding.name ssn) Binding.empty [] (locexs,[])
(mk_statespace_elements vds exts thy)) thy
end
fun statespace_cmd n exts vds thy =
let
open Syntax
val ctx = Named_Target.theory_init thy in
compile_statespace n (map (fn (n, t) => (n, read_typ ctx t)) vds) exts thy
end
end;
let open Parse in
Outer_Syntax.command @{command_keyword statespace} "define locale-based statespace with lenses"
((Parse.name --
(@{keyword "="} |-- Scan.repeat (Parse.name --| @{keyword "+"}) --
Scan.repeat (Parse.name -- ($$$ "::" |-- !!! typ))))
>> (fn (n, (exts, vds)) => Toplevel.theory (Lens_Statespace.statespace_cmd n exts vds)))
end
Theory Prisms
section ‹Prisms›
theory Prisms
imports Lenses
begin
subsection ‹ Signature and Axioms ›
text ‹Prisms are like lenses, but they act on sum types rather than product types~\cite{Gibbons17}.
See \url{https://hackage.haskell.org/package/lens-4.15.2/docs/Control-Lens-Prism.html}
for more information.›
record ('v, 's) prism =
prism_match :: "'s ⇒ 'v option" ("matchı")
prism_build :: "'v ⇒ 's" ("buildı")
type_notation
prism (infixr "⟹⇩△" 0)
locale wb_prism =
fixes x :: "'v ⟹⇩△ 's" (structure)
assumes match_build: "match (build v) = Some v"
and build_match: "match s = Some v ⟹ s = build v"
begin
lemma build_match_iff: "match s = Some v ⟷ s = build v"
using build_match match_build by blast
lemma range_build: "range build = dom match"
using build_match match_build by fastforce
end
declare wb_prism.match_build [simp]
declare wb_prism.build_match [simp]
subsection ‹ Co-dependence ›
text ‹ The relation states that two prisms construct disjoint elements of the source. This
can occur, for example, when the two prisms characterise different constructors of an
algebraic datatype. ›
definition prism_diff :: "('a ⟹⇩△ 's) ⇒ ('b ⟹⇩△ 's) ⇒ bool" (infix "∇" 50) where
"prism_diff X Y = (range build⇘X⇙ ∩ range build⇘Y⇙ = {})"
lemma prism_diff_intro:
"(⋀ s⇩1 s⇩2. build⇘X⇙ s⇩1 = build⇘Y⇙ s⇩2 ⟹ False) ⟹ X ∇ Y"
by (auto simp add: prism_diff_def)
lemma prism_diff_irrefl: "¬ X ∇ X"
by (simp add: prism_diff_def)
lemma prism_diff_sym: "X ∇ Y ⟹ Y ∇ X"
by (auto simp add: prism_diff_def)
lemma prism_diff_build: "X ∇ Y ⟹ build⇘X⇙ u ≠ build⇘Y⇙ v"
by (simp add: disjoint_iff_not_equal prism_diff_def)
subsection ‹ Summation ›
definition prism_plus :: "('a ⟹⇩△ 's) ⇒ ('b ⟹⇩△ 's) ⇒ 'a + 'b ⟹⇩△ 's" (infixl "+⇩△" 85)
where
"X +⇩△ Y = ⦇ prism_match = (λ s. case (match⇘X⇙ s, match⇘Y⇙ s) of
(Some u, _) ⇒ Some (Inl u) |
(None, Some v) ⇒ Some (Inr v) |
(None, None) ⇒ None),
prism_build = (λ v. case v of Inl x ⇒ build⇘X⇙ x | Inr y ⇒ build⇘Y⇙ y) ⦈"
subsection ‹ Instances ›
definition prism_suml :: "('a, 'a + 'b) prism" ("Inl⇩△") where
[lens_defs]: "prism_suml = ⦇ prism_match = (λ v. case v of Inl x ⇒ Some x | _ ⇒ None), prism_build = Inl ⦈"
definition prism_sumr :: "('b, 'a + 'b) prism" ("Inr⇩△") where
[lens_defs]: "prism_sumr = ⦇ prism_match = (λ v. case v of Inr x ⇒ Some x | _ ⇒ None), prism_build = Inr ⦈"
lemma wb_prim_suml: "wb_prism Inl⇩△"
apply (unfold_locales)
apply (simp_all add: prism_suml_def sum.case_eq_if)
apply (metis option.inject option.simps(3) sum.collapse(1))
done
lemma wb_prim_sumr: "wb_prism Inr⇩△"
apply (unfold_locales)
apply (simp_all add: prism_sumr_def sum.case_eq_if)
apply (metis option.distinct(1) option.inject sum.collapse(2))
done
lemma prism_suml_indep_sumr [simp]: "Inl⇩△ ∇ Inr⇩△"
by (auto simp add: prism_diff_def lens_defs)
subsection ‹ Lens correspondence ›
text ‹ Every well-behaved prism can be represented by a partial bijective lens. We prove
this by exhibiting conversion functions and showing they are (almost) inverses. ›
definition prism_lens :: "('a, 's) prism ⇒ ('a ⟹ 's)" where
"prism_lens X = ⦇ lens_get = (λ s. the (match⇘X⇙ s)), lens_put = (λ s v. build⇘X⇙ v) ⦈"
definition lens_prism :: "('a ⟹ 's) ⇒ ('a, 's) prism" where
"lens_prism X = ⦇ prism_match = (λ s. if (s ∈ 𝒮⇘X⇙) then Some (get⇘X⇙ s) else None)
, prism_build = create⇘X⇙ ⦈"
lemma get_prism_lens: "get⇘prism_lens X⇙ = the ∘ match⇘X⇙"
by (simp add: prism_lens_def fun_eq_iff)
lemma src_prism_lens: "𝒮⇘prism_lens X⇙ = range (build⇘X⇙)"
by (auto simp add: prism_lens_def lens_source_def)
lemma create_prism_lens: "create⇘prism_lens X⇙ = build⇘X⇙"
by (simp add: prism_lens_def lens_create_def fun_eq_iff)
lemma prism_lens_inverse:
"wb_prism X ⟹ lens_prism (prism_lens X) = X"
unfolding lens_prism_def src_prism_lens create_prism_lens get_prism_lens
by (auto intro: prism.equality simp add: fun_eq_iff domIff wb_prism.range_build)
text ‹ Function @{const lens_prism} is almost inverted by @{const prism_lens}. The $put$
functions are identical, but the $get$ functions differ when applied to a source where
the prism @{term X} is undefined. ›
lemma lens_prism_put_inverse:
"pbij_lens X ⟹ put⇘prism_lens (lens_prism X)⇙ = put⇘X⇙"
unfolding prism_lens_def lens_prism_def
by (auto simp add: fun_eq_iff pbij_lens.put_is_create)
lemma wb_prism_implies_pbij_lens:
"wb_prism X ⟹ pbij_lens (prism_lens X)"
by (unfold_locales, simp_all add: prism_lens_def)
lemma pbij_lens_implies_wb_prism:
assumes "pbij_lens X"
shows "wb_prism (lens_prism X)"
proof (unfold_locales)
fix s v
show "match⇘lens_prism X⇙ (build⇘lens_prism X⇙ v) = Some v"
by (simp add: lens_prism_def weak_lens.create_closure assms)
assume a: "match⇘lens_prism X⇙ s = Some v"
show "s = build⇘lens_prism X⇙ v"
proof (cases "s ∈ 𝒮⇘X⇙")
case True
with a assms show ?thesis
by (simp add: lens_prism_def lens_create_def,
metis mwb_lens.weak_get_put pbij_lens.put_det pbij_lens_mwb)
next
case False
with a assms show ?thesis by (simp add: lens_prism_def)
qed
qed
ML_file ‹Prism_Lib.ML›
end
File ‹Prism_Lib.ML›
signature PRISM_LIB =
sig
val wb_prismN: string
val prism_diffN: string
val codepsN: string
val prismT: typ -> typ -> typ
val isPrismT: typ -> bool
val mk_wb_prism: term -> term
val mk_codep: term -> term -> term
end
structure Prism_Lib : PRISM_LIB =
struct
open Syntax
open HOLogic
val wb_prismN = @{const_name wb_prism}
val prism_diffN = @{const_name prism_diff}
val codepsN = "codeps"
fun prismT a b = Type (@{type_name prism_ext}, [a, b, unitT])
fun isPrismT (Type (n, _)) = (n = @{type_name prism_ext}) |
isPrismT _ = false
fun mk_wb_prism t = mk_Trueprop (const wb_prismN $ t)
fun mk_codep x y = mk_Trueprop (const prism_diffN $ x $ y)
end
Theory Channel_Type
section ‹ Channel Types ›
theory Channel_Type
imports Prisms
keywords "chantype" :: "thy_defn"
begin
text ‹ A channel type is a simplified algebraic datatype where each constructor has exactly
one parameter, and it is wrapped up as a prism. It a dual of an alphabet type. ›
definition ctor_prism :: "('a ⇒ 'd) ⇒ ('d ⇒ bool) ⇒ ('d ⇒ 'a) ⇒ ('a ⟹⇩△ 'd)" where
[lens_defs]:
"ctor_prism ctor disc sel = ⦇ prism_match = (λ d. if (disc d) then Some (sel d) else None)
, prism_build = ctor ⦈"
lemma wb_ctor_prism_intro:
assumes
"⋀ v. disc (ctor v)"
"⋀ v. sel (ctor v) = v"
"⋀ s. disc s ⟹ ctor (sel s) = s"
shows "wb_prism (ctor_prism ctor disc sel)"
by (unfold_locales, simp_all add: lens_defs assms)
(metis assms(3) option.distinct(1) option.sel)
lemma ctor_codep_intro:
assumes "⋀ x y. ctor1 x ≠ ctor2 y"
shows "ctor_prism ctor1 disc1 sel1 ∇ ctor_prism ctor2 disc2 sel2"
by (rule prism_diff_intro, simp add: lens_defs assms)
ML_file "Channel_Type.ML"
end
File ‹Channel_Type.ML›
structure Channel_Type =
struct
fun prove_prism_goal thy =
let
open Simplifier; open Global_Theory; open Lens_Lib
val ctx = Named_Target.theory_init thy
in
auto_tac (fold add_simp (get_thms thy lens_defsN) ctx)
end
val wb_prism_suffix = "_wb_prism"
val codep_suffix = "_codeps"
val ctor_suffix = "_C"
fun wb_prism_proof x thms ctx =
let open Simplifier; open Prism_Lib; open Syntax
in
Goal.prove ctx [] []
(Syntax.check_term ctx (mk_wb_prism (free x)))
(fn {context = context, prems = _}
=> EVERY [ PARALLEL_ALLGOALS (full_simp_tac (fold add_simp thms context))
, Classical.rule_tac context [@{thm wb_ctor_prism_intro}] [] 1
, auto_tac (context delsimprocs [@{simproc unit_eq}])
])
end
fun codep_proof thms ctx (x, y) =
let open Simplifier; open Prism_Lib; open Syntax in
Goal.prove ctx [] []
(Syntax.check_term ctx (mk_codep (free x) (free y)))
(fn {context = context, prems = _}
=> EVERY [ PARALLEL_ALLGOALS (full_simp_tac (fold add_simp thms context))
, Classical.rule_tac context [@{thm ctor_codep_intro}] [] 1
, auto_tac ctx
])
end
fun mk_def ty x v = Const ("Pure.eq", ty --> ty --> Term.propT) $ Free (x, ty) $ v;
val is_prefix = "is_";
val un_prefix = "un_";
fun def qual (x, tm) ctx =
let open Specification; open Syntax
val ((_, (_, thm)), d) = definition (SOME (Binding.qualify false qual (Binding.name x), NONE, NoSyn)) [] [] ((Binding.empty, [Attrib.check_src @{context} (Token.make_src ("lens_defs", Position.none) [])]), mk_def dummyT x tm) ctx
in (thm, d)
end
fun defs qual inds f ctx =
fold (fn i => fn (thms, ctx) =>
let val (thm, ctx') = def qual (i, f i) ctx
in (thms @ [thm], ctx') end) inds ([], ctx)
fun compile_chantype (name, chans) ctx =
let
open BNF_FP_Def_Sugar; open BNF_FP_Rec_Sugar_Util; open BNF_LFP; open Ctr_Sugar
open Prism_Lib; open Lens_Lib; open Local_Theory; open Specification; open Syntax
val ctrs = map (fn (n, t) => (((Binding.empty, Binding.name (n ^ ctor_suffix)), [(Binding.empty, t)]), Mixfix.NoSyn)) chans
val pnames = map fst chans
val thypfx =
case (Named_Target.locale_of ctx) of
SOME loc => loc ^ "." |
NONE => Context.theory_name (Local_Theory.exit_global ctx) ^ "."
val prefix = thypfx ^ name ^ "."
val simp = Attrib.check_src @{context} (Token.make_src ("simp", Position.none) [])
val dummy_disc = absdummy dummyT @{term True}
in
(co_datatype_cmd Least_FP construct_lfp
((K Plugin_Name.default_filter, true),
[((((([],Binding.name name), Mixfix.NoSyn), ctrs), (Binding.empty, Binding.empty, Binding.empty)),[])])
#> (if (length pnames = 2)
then abbreviation
Syntax.mode_default (SOME (Binding.qualify false name (Binding.name (is_prefix ^ nth pnames 1 ^ ctor_suffix)), NONE, NoSyn)) []
(mk_def dummyT
(is_prefix ^ nth pnames 1 ^ ctor_suffix)
(const @{const_name comp} $ @{term Not} $ const (prefix ^ is_prefix ^ nth pnames 0 ^ ctor_suffix))) false
else I)
#> defs name pnames (fn x => (const @{const_name ctor_prism}
$ const (prefix ^ x ^ ctor_suffix)
$ (if (length pnames = 1) then dummy_disc else const (prefix ^ is_prefix ^ x ^ ctor_suffix))
$ const (prefix ^ un_prefix ^ x ^ ctor_suffix)))
#> (fn (thms, ctx) =>
(fold (fn x => fn thy => snd (note ((Binding.qualify false name (Binding.name (x ^ wb_prism_suffix)), [simp]), [wb_prism_proof x thms thy]) thy)) pnames
#> (snd o note ((Binding.qualify false name (Binding.name codepsN), [simp]), map (codep_proof thms ctx) (pairings pnames)))
) ctx))
ctx
end;
end;
let open Parse; open Parse_Spec; open Scan in
Outer_Syntax.command @{command_keyword chantype} "define a channel datatype"
((name --
(@{keyword "="} |-- repeat1 (name -- ($$$ "::" |-- !!! typ))))
>> (fn x => Toplevel.local_theory NONE NONE (Channel_Type.compile_chantype x)))
end;
Theory Dataspace
section ‹ Data spaces ›
theory Dataspace
imports Lenses Prisms
keywords "dataspace" :: "thy_defn" and "constants" "variables" "channels"
begin
text ‹ A data space is like a more sophisticated version of a locale-based state space. It
allows us to introduce both variables, modelled by lenses, and channels, modelled by
prisms. It also allows local constants, and assumptions over them. ›
ML_file "Dataspace.ML"
end
File ‹Dataspace.ML›
structure Dataspace =
struct
val achanT = TFree ("'ch", ["HOL.type"])
fun mk_constant_elements cds =
let
open Syntax;
val cfixes = map (fn (n, t) => (Binding.name n, SOME t, NoSyn)) cds
in [ Element.Fixes cfixes ] end;
fun mk_channel_elements exts chds thy =
let
open Lens_Lib;
open Prism_Lib;
open Syntax;
val ctx = Named_Target.theory_init thy
val chns = map fst chds
val chfixes = map (fn (n, t) => (Binding.name n, SOME (prismT t achanT), NoSyn)) chds
val wbs = map (mk_wb_prism o free) chns;
val simp = Attrib.check_src ctx (Token.make_src ("simp", Position.none) [])
val imps = map (fn e => Locale.check thy (e, Position.none)) exts
val ilnsls = map (filter (isPrismT o snd) o map fst o Locale.params_of thy) imps
val impind = List.concat (map (pairsWith chns) (map (map fst) ilnsls))
val codeps = map (fn (x, y) => mk_codep (free x) (free y))
(pairings (chns) @ impind)
in
[ Element.Fixes chfixes
, Element.Assumes [((Binding.name "prisms", [simp])
, map (fn wb => (wb, [])) wbs),
((Binding.name "codeps", [simp])
, map (fn dp => (dp, [])) codeps)]
]
end
val STATE = "STATE_TYPE"
val CHAN = "CHAN_TYPE"
fun compile_dataspace nm exts cds assms vds chds thy =
let
open Lens_Lib;
open Syntax;
val imps = map (fn e => Locale.check thy (e, Position.none)) exts
val pinsts = Expression.Named
[(STATE, Logic.mk_type astateT)
,(CHAN, Logic.mk_type achanT)]
val locexs = map (fn i => (i, (("", false), (pinsts, [])))) imps
val st_par = (Binding.name STATE, SOME (Term.itselfT astateT), NoSyn)
val ch_par = (Binding.name CHAN, SOME (Term.itselfT achanT), NoSyn)
in
(Local_Theory.exit_global o snd o
Expression.add_locale (Binding.name nm) Binding.empty [] (locexs, [])
(mk_constant_elements cds
@ [Element.Assumes assms]
@ Lens_Statespace.mk_statespace_elements vds exts thy
@ mk_channel_elements exts chds thy
@ [Element.Fixes [st_par, ch_par]])) thy
end
fun dataspace_cmd nm exts cds rassms vds chds thy =
let
open Syntax; open HOLogic;
val ctx = Named_Target.theory_init thy
val assms = map (fn (b, td) => (b, map (fn (t, ts) => (mk_Trueprop (parse_term ctx t), map (mk_Trueprop o parse_term ctx) ts)) td)) rassms in
compile_dataspace nm exts (map (fn (n, t) => (n, read_typ ctx t)) cds) assms (map (fn (n, t) => (n, read_typ ctx t)) vds) (map (fn (n, t) => (n, read_typ ctx t)) chds) thy
end
end;
let open Parse; open Parse_Spec; open Scan in
Outer_Syntax.command @{command_keyword dataspace} "define reactive contexts"
((name --
((@{keyword "="} |-- repeat (Parse.name --| @{keyword "+"}))
-- optional (@{keyword "constants"} |-- repeat1 (name -- ($$$ "::" |-- !!! typ))) []
-- optional (@{keyword "assumes"} |-- !!! (and_list1 (opt_thm_name ":" -- repeat1 propp))) []
-- optional (@{keyword "variables"} |-- repeat1 (name -- ($$$ "::" |-- !!! typ))) []
-- optional (@{keyword "channels"} |-- repeat1 (name -- ($$$ "::" |-- !!! typ))) [])
>> (fn (nm, ((((exts, cds), assms), vds), chds)) => Toplevel.theory (Dataspace.dataspace_cmd nm exts cds assms vds chds))))
end;
Theory Scenes
section ‹ Scenes ›
theory Scenes
imports Lens_Instances
begin
text ‹ Like lenses, scenes characterise a region of a source type. However, unlike lenses, scenes
do not explicitly assign a view type to this region, and consequently they have just one type
parameter. This means they can be more flexibly composed, and in particular it is possible to
show they form nice algebraic structures in Isabelle/HOL. They are mainly of use in characterising
sets of variables, where, of course, we do not care about the types of those variables and
therefore representing them as lenses is inconvenient. ›
subsection ‹ Overriding Functions ›
text ‹ Overriding functions provide an abstract way of replacing a region of an existing source
with the corresponding region of another source. ›
locale overrider =
fixes F :: "'s ⇒ 's ⇒ 's" (infixl "▹" 65)
assumes
ovr_overshadow_left: "x ▹ y ▹ z = x ▹ z" and
ovr_overshadow_right: "x ▹ (y ▹ z) = x ▹ z"
begin
lemma ovr_assoc: "x ▹ (y ▹ z) = x ▹ y ▹ z"
by (simp add: ovr_overshadow_left ovr_overshadow_right)
end
locale idem_overrider = overrider +
assumes ovr_idem: "x ▹ x = x"
declare overrider.ovr_overshadow_left [simp]
declare overrider.ovr_overshadow_right [simp]
declare idem_overrider.ovr_idem [simp]
subsection ‹ Scene Type ›
typedef 's scene = "{F :: 's ⇒ 's ⇒ 's. overrider F}"
by (rule_tac x="λ x y. x" in exI, simp, unfold_locales, simp_all)
setup_lifting type_definition_scene
lift_definition idem_scene :: "'s scene ⇒ bool" is idem_overrider .
lift_definition region :: "'s scene ⇒ 's rel"
is "λ F. {(s⇩1, s⇩2). (∀ s. F s s⇩1 = F s s⇩2)}" .
lift_definition coregion :: "'s scene ⇒ 's rel"
is "λ F. {(s⇩1, s⇩2). (∀ s. F s⇩1 s = F s⇩2 s)}" .
lemma equiv_region: "equiv UNIV (region X)"
apply (transfer)
apply (rule equivI)
apply (rule refl_onI)
apply (auto)
apply (rule symI)
apply (auto)
apply (rule transI)
apply (auto)
done
lemma equiv_coregion: "equiv UNIV (coregion X)"
apply (transfer)
apply (rule equivI)
apply (rule refl_onI)
apply (auto)
apply (rule symI)
apply (auto)
apply (rule transI)
apply (auto)
done
lemma region_coregion_Id:
"idem_scene X ⟹ region X ∩ coregion X = Id"
by (transfer, auto, metis idem_overrider.ovr_idem)
lemma state_eq_iff: "idem_scene S ⟹ x = y ⟷ (x, y) ∈ region S ∧ (x, y) ∈ coregion S"
by (metis IntE IntI pair_in_Id_conv region_coregion_Id)
lift_definition scene_override :: "'a ⇒ 'a ⇒ ('a scene) ⇒ 'a" ("_ ⊕⇩S _ on _" [95,0,96] 95)
is "λ s⇩1 s⇩2 F. F s⇩1 s⇩2" .
abbreviation (input) scene_copy :: "'a scene ⇒ 'a ⇒ ('a ⇒ 'a)" ("cp⇘_⇙") where
"cp⇘A⇙ s ≡ (λ s'. s' ⊕⇩S s on A)"
lemma scene_override_idem [simp]: "idem_scene X ⟹ s ⊕⇩S s on X = s"
by (transfer, simp)
lemma scene_override_overshadow_left [simp]:
"S⇩1 ⊕⇩S S⇩2 on X ⊕⇩S S⇩3 on X = S⇩1 ⊕⇩S S⇩3 on X"
by (transfer, simp)
lemma scene_override_overshadow_right [simp]:
"S⇩1 ⊕⇩S (S⇩2 ⊕⇩S S⇩3 on X) on X = S⇩1 ⊕⇩S S⇩3 on X"
by (transfer, simp)
definition scene_equiv :: "'a ⇒ 'a ⇒ ('a scene) ⇒ bool" ("_ ≈⇩S _ on _" [65,0,66] 65) where
[lens_defs]: "S⇩1 ≈⇩S S⇩2 on X = (S⇩1 ⊕⇩S S⇩2 on X = S⇩1)"
lemma scene_equiv_region: "idem_scene X ⟹ region X = {(S⇩1, S⇩2). S⇩1 ≈⇩S S⇩2 on X}"
by (simp add: lens_defs, transfer, auto)
(metis idem_overrider.ovr_idem, metis overrider.ovr_overshadow_right)
lift_definition scene_indep :: "'a scene ⇒ 'a scene ⇒ bool" (infix "⨝⇩S" 50)
is "λ F G. (∀ s⇩1 s⇩2 s⇩3. G (F s⇩1 s⇩2) s⇩3 = F (G s⇩1 s⇩3) s⇩2)" .
lemma scene_indep_override:
"X ⨝⇩S Y = (∀ s⇩1 s⇩2 s⇩3. s⇩1 ⊕⇩S s⇩2 on X ⊕⇩S s⇩3 on Y = s⇩1 ⊕⇩S s⇩3 on Y ⊕⇩S s⇩2 on X)"
by (transfer, auto)
lemma scene_indep_copy:
"X ⨝⇩S Y = (∀ s⇩1 s⇩2. cp⇘X⇙ s⇩1 ∘ cp⇘Y⇙ s⇩2 = cp⇘Y⇙ s⇩2 ∘ cp⇘X⇙ s⇩1)"
by (auto simp add: scene_indep_override comp_def fun_eq_iff)
lemma scene_indep_sym:
"X ⨝⇩S Y ⟹ Y ⨝⇩S X"
by (transfer, auto)
text ‹ Compatibility is a weaker notion than independence; the scenes can overlap but they must
agree when they do. ›
lift_definition scene_compat :: "'a scene ⇒ 'a scene ⇒ bool" (infix "##⇩S" 50)
is "λ F G. (∀ s⇩1 s⇩2. G (F s⇩1 s⇩2) s⇩2 = F (G s⇩1 s⇩2) s⇩2)" .
lemma scene_compat_copy:
"X ##⇩S Y = (∀ s. cp⇘X⇙ s ∘ cp⇘Y⇙ s = cp⇘Y⇙ s ∘ cp⇘X⇙ s)"
by (transfer, auto simp add: fun_eq_iff)
lemma scene_indep_compat [simp]: "X ⨝⇩S Y ⟹ X ##⇩S Y"
by (transfer, auto)
lemma scene_compat_refl: "X ##⇩S X"
by (transfer, simp)
lemma scene_compat_sym: "X ##⇩S Y ⟹ Y ##⇩S X"
by (transfer, simp)
lemma scene_override_commute_indep:
assumes "X ⨝⇩S Y"
shows "S⇩1 ⊕⇩S S⇩2 on X ⊕⇩S S⇩3 on Y = S⇩1 ⊕⇩S S⇩3 on Y ⊕⇩S S⇩2 on X"
using assms
by (transfer, auto)
instantiation scene :: (type) "{bot, top, uminus, sup, inf}"
begin
lift_definition bot_scene :: "'s scene" is "λ x y. x" by (unfold_locales, simp_all)
lift_definition top_scene :: "'s scene" is "λ x y. y" by (unfold_locales, simp_all)
lift_definition uminus_scene :: "'s scene ⇒ 's scene" is "λ F x y. F y x"
by (unfold_locales, simp_all)
text ‹ Scene union requires that the two scenes are at least compatible. If they are not, the
result is the bottom scene. ›
lift_definition sup_scene :: "'s scene ⇒ 's scene ⇒ 's scene"
is "λ F G. if (∀ s⇩1 s⇩2. G (F s⇩1 s⇩2) s⇩2 = F (G s⇩1 s⇩2) s⇩2) then (λ s⇩1 s⇩2. G (F s⇩1 s⇩2) s⇩2) else (λ s⇩1 s⇩2. s⇩1)"
by (unfold_locales, auto, metis overrider.ovr_overshadow_right)
definition inf_scene :: "'s scene ⇒ 's scene ⇒ 's scene" where
[lens_defs]: "inf_scene X Y = - (sup (- X) (- Y))"
instance ..
end
abbreviation union_scene :: "'s scene ⇒ 's scene ⇒ 's scene" (infixl "⊔⇩S" 65)
where "union_scene ≡ sup"
abbreviation inter_scene :: "'s scene ⇒ 's scene ⇒ 's scene" (infixl "⊓⇩S" 70)
where "inter_scene ≡ inf"
abbreviation top_scene :: "'s scene" ("⊤⇩S")
where "top_scene ≡ top"
abbreviation bot_scene :: "'s scene" ("⊥⇩S")
where "bot_scene ≡ bot"
lemma uminus_scene_twice: "- (- (X :: 's scene)) = X"
by (transfer, simp)
lemma scene_override_id [simp]: "S⇩1 ⊕⇩S S⇩2 on ⊤⇩S = S⇩2"
by (transfer, simp)
lemma scene_override_unit [simp]: "S⇩1 ⊕⇩S S⇩2 on ⊥⇩S = S⇩1"
by (transfer, simp)
lemma scene_override_commute: "S⇩2 ⊕⇩S S⇩1 on (- X) = S⇩1 ⊕⇩S S⇩2 on X"
by (transfer, simp)
lemma scene_union_incompat: "¬ X ##⇩S Y ⟹ X ⊔⇩S Y = ⊥⇩S"
by (transfer, auto)
lemma scene_override_union: "X ##⇩S Y ⟹ S⇩1 ⊕⇩S S⇩2 on (X ⊔⇩S Y) = (S⇩1 ⊕⇩S S⇩2 on X) ⊕⇩S S⇩2 on Y"
by (transfer, auto)
lemma scene_union_unit: "X ⊔⇩S ⊥⇩S = X"
by (transfer, simp)
lemma idem_scene_union [simp]: "⟦ idem_scene A; idem_scene B ⟧ ⟹ idem_scene (A ⊔⇩S B)"
apply (transfer, auto)
apply (unfold_locales, auto)
apply (metis overrider.ovr_overshadow_left)
apply (metis overrider.ovr_overshadow_right)
done
lemma scene_union_annhil: "idem_scene X ⟹ X ⊔⇩S ⊤⇩S = ⊤⇩S"
by (transfer, simp)
lemma scene_union_pres_compat: "⟦ A ##⇩S B; A ##⇩S C ⟧ ⟹ A ##⇩S (B ⊔⇩S C)"
by (transfer, auto)
lemma scene_indep_self_compl: "A ⨝⇩S -A"
by (transfer, simp)
lemma scene_compat_self_compl: "A ##⇩S -A"
by (transfer, simp)
lemma scene_union_assoc:
assumes "X ##⇩S Y" "X ##⇩S Z" "Y ##⇩S Z"
shows "X ⊔⇩S (Y ⊔⇩S Z) = (X ⊔⇩S Y) ⊔⇩S Z"
using assms by (transfer, auto)
lemma scene_inter_indep:
assumes "idem_scene X" "idem_scene Y" "X ⨝⇩S Y"
shows "X ⊓⇩S Y = ⊥⇩S"
using assms
unfolding lens_defs
apply (transfer, auto)
apply (metis (no_types, hide_lams) idem_overrider.ovr_idem overrider.ovr_assoc overrider.ovr_overshadow_right)
apply (metis (no_types, hide_lams) idem_overrider.ovr_idem overrider.ovr_overshadow_right)
done
lemma scene_union_idem: "X ⊔⇩S X = X"
by (transfer, simp)
lemma scene_union_compl: "idem_scene X ⟹ X ⊔⇩S - X = ⊤⇩S"
by (transfer, auto)
lemma scene_inter_idem: "X ⊓⇩S X = X"
by (simp add: inf_scene_def, transfer, auto)
lemma scene_union_commute: "X ⊔⇩S Y = Y ⊔⇩S X"
by (transfer, auto)
lemma scene_inter_compl: "idem_scene X ⟹ X ⊓⇩S - X = ⊥⇩S"
by (simp add: inf_scene_def, transfer, auto)
lemma scene_demorgan1: "-(X ⊔⇩S Y) = -X ⊓⇩S -Y"
by (simp add: inf_scene_def, transfer, auto)
lemma scene_demorgan2: "-(X ⊓⇩S Y) = -X ⊔⇩S -Y"
by (simp add: inf_scene_def, transfer, auto)
lemma scene_compat_top: "idem_scene X ⟹ X ##⇩S ⊤⇩S"
by (transfer, simp)
instantiation scene :: (type) ord
begin
text ‹ $X$ is a subscene of $Y$ provided that overriding with first $Y$ and then $X$ can
be rewritten using the complement of $X$. ›
definition less_eq_scene :: "'a scene ⇒ 'a scene ⇒ bool" where
[lens_defs]: "less_eq_scene X Y = (∀ s⇩1 s⇩2 s⇩3. s⇩1 ⊕⇩S s⇩2 on Y ⊕⇩S s⇩3 on X = s⇩1 ⊕⇩S (s⇩2 ⊕⇩S s⇩3 on X) on Y)"
definition less_scene :: "'a scene ⇒ 'a scene ⇒ bool" where
[lens_defs]: "less_scene x y = (x ≤ y ∧ ¬ y ≤ x)"
instance ..
end
abbreviation subscene :: "'a scene ⇒ 'a scene ⇒ bool" (infix "⊆⇩S" 55)
where "subscene X Y ≡ X ≤ Y"
lemma subscene_refl: "X ⊆⇩S X"
by (simp add: less_eq_scene_def)
lemma subscene_trans: "⟦ idem_scene Y; X ⊆⇩S Y; Y ⊆⇩S Z ⟧ ⟹ X ⊆⇩S Z"
by (simp add: less_eq_scene_def, transfer, auto, metis (no_types, hide_lams) idem_overrider.ovr_idem)
lemma subscene_antisym: "⟦ idem_scene Y; X ⊆⇩S Y; Y ⊆⇩S X ⟧ ⟹ X = Y"
apply (simp add: less_eq_scene_def, transfer, auto)
apply (rule ext)
apply (rule ext)
apply (metis (full_types) idem_overrider.ovr_idem overrider.ovr_overshadow_left)
done
lemma subscene_copy_def:
assumes "idem_scene X" "idem_scene Y"
shows "X ⊆⇩S Y = (∀ s⇩1 s⇩2. cp⇘X⇙ s⇩1 ∘ cp⇘Y⇙ s⇩2 = cp⇘Y⇙ (cp⇘X⇙ s⇩1 s⇩2))"
using assms
by (simp add: less_eq_scene_def fun_eq_iff, transfer, auto)
lemma subscene_eliminate:
"⟦ idem_scene Y; X ≤ Y ⟧ ⟹ s⇩1 ⊕⇩S s⇩2 on X ⊕⇩S s⇩3 on Y = s⇩1 ⊕⇩S s⇩3 on Y"
by (metis less_eq_scene_def scene_override_overshadow_left scene_override_idem)
lemma scene_bot_least: "⊥⇩S ≤ X"
unfolding less_eq_scene_def by (transfer, auto)
lemma scene_top_greatest: "X ≤ ⊤⇩S"
unfolding less_eq_scene_def by (transfer, auto)
lemma scene_union_ub: "⟦ idem_scene Y; X ⨝⇩S Y ⟧ ⟹ X ≤ (X ⊔⇩S Y)"
by (simp add: less_eq_scene_def, transfer, auto)
(metis (no_types, hide_lams) idem_overrider.ovr_idem overrider.ovr_overshadow_right)
lemma scene_le_then_compat: "⟦ idem_scene X; idem_scene Y; X ≤ Y ⟧ ⟹ X ##⇩S Y"
unfolding less_eq_scene_def
by (transfer, auto, metis (no_types, lifting) idem_overrider.ovr_idem overrider_def)
lemma indep_then_compl_in: "A ⨝⇩S B ⟹ A ≤ -B"
unfolding less_eq_scene_def by (transfer, simp)
lift_definition scene_comp :: "'a scene ⇒ ('a ⟹ 'b) ⇒ 'b scene" (infixl ";⇩S" 80)
is "λ S X a b. if (vwb_lens X) then put⇘X⇙ a (S (get⇘X⇙ a) (get⇘X⇙ b)) else a"
by (unfold_locales, auto)
lemma scene_comp_idem [simp]: "idem_scene S ⟹ idem_scene (S ;⇩S X)"
by (transfer, unfold_locales, simp_all)
lemma scene_comp_lens_indep [simp]: "X ⨝ Y ⟹ (A ;⇩S X) ⨝⇩S (A ;⇩S Y)"
by (transfer, auto simp add: lens_indep.lens_put_comm lens_indep.lens_put_irr2)
lemma scene_comp_indep [simp]: "A ⨝⇩S B ⟹ (A ;⇩S X) ⨝⇩S (B ;⇩S X)"
by (transfer, auto)
subsection ‹ Linking Scenes and Lenses ›
text ‹ The following function extracts a scene from a very well behaved lens ›
lift_definition lens_scene :: "('v ⟹ 's) ⇒ 's scene" ("⟦_⟧⇩∼") is
"λ X s⇩1 s⇩2. if (mwb_lens X) then s⇩1 ⊕⇩L s⇩2 on X else s⇩1"
by (unfold_locales, auto simp add: lens_override_def)
lemma vwb_impl_idem_scene [simp]:
"vwb_lens X ⟹ idem_scene ⟦X⟧⇩∼"
by (transfer, unfold_locales, auto simp add: lens_override_overshadow_left lens_override_overshadow_right)
lemma idem_scene_impl_vwb:
"⟦ mwb_lens X; idem_scene ⟦X⟧⇩∼ ⟧ ⟹ vwb_lens X"
apply (cases "mwb_lens X")
apply (transfer, unfold idem_overrider_def overrider_def, auto)
apply (simp add: idem_overrider_axioms_def override_idem_implies_vwb)
done
lemma lens_compat_scene: "⟦ mwb_lens X; mwb_lens Y ⟧ ⟹ X ##⇩L Y ⟷ ⟦X⟧⇩∼ ##⇩S ⟦Y⟧⇩∼"
by (auto simp add: lens_scene.rep_eq scene_compat.rep_eq lens_defs)
text ‹ Next we show some important congruence properties ›
lemma zero_lens_scene: "⟦0⇩L⟧⇩∼ = ⊥⇩S"
by (transfer, simp)
lemma one_lens_scene: "⟦1⇩L⟧⇩∼ = ⊤⇩S"
by (transfer, simp)
lemma lens_scene_override:
"mwb_lens X ⟹ s⇩1 ⊕⇩S s⇩2 on ⟦X⟧⇩∼ = s⇩1 ⊕⇩L s⇩2 on X"
by (transfer, simp)
lemma lens_indep_scene:
assumes "vwb_lens X" "vwb_lens Y"
shows "(X ⨝ Y) ⟷ ⟦X⟧⇩∼ ⨝⇩S ⟦Y⟧⇩∼"
using assms
by (auto, (simp add: scene_indep_override, transfer, simp add: lens_indep_override_def)+)
lemma lens_indep_impl_scene_indep [simp]:
"(X ⨝ Y) ⟹ ⟦X⟧⇩∼ ⨝⇩S ⟦Y⟧⇩∼"
by (transfer, auto simp add: lens_indep_comm lens_override_def)
lemma lens_plus_scene:
"⟦ vwb_lens X; vwb_lens Y; X ⨝ Y ⟧ ⟹ ⟦X +⇩L Y⟧⇩∼ = ⟦X⟧⇩∼ ⊔⇩S ⟦Y⟧⇩∼"
by (transfer, auto simp add: lens_override_plus lens_indep_override_def lens_indep_overrideI plus_vwb_lens)
lemma subscene_implies_sublens': "⟦ vwb_lens X; vwb_lens Y ⟧ ⟹ ⟦X⟧⇩∼ ≤ ⟦Y⟧⇩∼ ⟷ X ⊆⇩L' Y"
by (simp add: lens_defs less_eq_scene_def, transfer, simp add: lens_override_def)
lemma sublens'_implies_subscene: "⟦ vwb_lens X; vwb_lens Y; X ⊆⇩L' Y ⟧ ⟹ ⟦X⟧⇩∼ ≤ ⟦Y⟧⇩∼"
by (simp add: lens_defs less_eq_scene_def, auto simp add: lens_override_def lens_scene_override)
lemma sublens_iff_subscene:
assumes "vwb_lens X" "vwb_lens Y"
shows "X ⊆⇩L Y ⟷ ⟦X⟧⇩∼ ≤ ⟦Y⟧⇩∼"
by (simp add: assms sublens_iff_sublens' subscene_implies_sublens')
text ‹ Equality on scenes is sound and complete with respect to lens equivalence. ›
lemma lens_equiv_scene:
assumes "vwb_lens X" "vwb_lens Y"
shows "X ≈⇩L Y ⟷ ⟦X⟧⇩∼ = ⟦Y⟧⇩∼"
proof
assume a: "X ≈⇩L Y"
show "⟦X⟧⇩∼ = ⟦Y⟧⇩∼"
by (meson a assms lens_equiv_def sublens_iff_subscene subscene_antisym vwb_impl_idem_scene)
next
assume b: "⟦X⟧⇩∼ = ⟦Y⟧⇩∼"
show "X ≈⇩L Y"
by (simp add: assms b lens_equiv_def sublens_iff_subscene subscene_refl)
qed
definition lens_insert :: "('a ⟹ 'b) ⇒ 'b scene ⇒ 'b scene" ("insert⇩S") where
"lens_insert x A = (if (⟦x⟧⇩∼ ≤ A) then ⟦x⟧⇩∼ ⊔⇩S A else A)"
lemma lens_insert_idem: "insert⇩S x (insert⇩S x A) = insert⇩S x A"
apply (auto simp add: lens_insert_def less_eq_scene_def)
apply (transfer)
apply (auto simp add: lens_override_overshadow_left)
apply (metis lens_override_overshadow_left)
done
text ‹ Membership operations. These have slightly hacky definitions at the moment in order to
cater for @{term mwb_lens}. See if they can be generalised? ›
definition lens_member :: "('a ⟹ 'b) ⇒ 'b scene ⇒ bool" (infix "∈⇩S" 50) where
[lens_defs]:
"lens_member x A = ((∀ s⇩1 s⇩2 s⇩3. s⇩1 ⊕⇩S s⇩2 on A ⊕⇩L s⇩3 on x = s⇩1 ⊕⇩S (s⇩2 ⊕⇩L s⇩3 on x) on A) ∧
(∀ b b'. get⇘x⇙ (b ⊕⇩S b' on A) = get⇘x⇙ b'))"
lemma lens_member_override: "x ∈⇩S A ⟹ s⇩1 ⊕⇩S s⇩2 on A ⊕⇩L s⇩3 on x = s⇩1 ⊕⇩S (s⇩2 ⊕⇩L s⇩3 on x) on A"
using lens_member_def by force
lemma lens_member_put:
assumes "vwb_lens x" "idem_scene a" "x ∈⇩S a"
shows "put⇘x⇙ s v ⊕⇩S s on a = s"
by (metis (full_types) assms lens_member_override lens_override_def scene_override_idem vwb_lens.put_eq)
lemma lens_member_top: "x ∈⇩S ⊤⇩S"
by (auto simp add: lens_member_def)
abbreviation lens_not_member :: "('a ⟹ 'b) ⇒ 'b scene ⇒ bool" (infix "∉⇩S" 50) where
"x ∉⇩S A ≡ (x ∈⇩S - A)"
lemma lens_member_get_override [simp]: "x ∈⇩S a ⟹ get⇘x⇙ (b ⊕⇩S b' on a) = get⇘x⇙ b'"
by (simp add: lens_member_def)
lemma lens_not_member_get_override [simp]: "x ∉⇩S a ⟹ get⇘x⇙ (b ⊕⇩S b' on a) = get⇘x⇙ b"
by (simp add: lens_member_def scene_override_commute)
subsection ‹ Function Domain Scene ›
lift_definition fun_dom_scene :: "'a set ⇒ ('a ⇒ 'b::two) scene" ("fds") is
"λ A f g. override_on f g A" by (unfold_locales, simp_all add: override_on_def fun_eq_iff)
lemma fun_dom_scene_empty: "fds({}) = ⊥⇩S"
by (transfer, simp)
lemma fun_dom_scene_union: "fds(A ∪ B) = fds(A) ⊔⇩S fds(B)"
by (transfer, auto simp add: fun_eq_iff override_on_def)
lemma fun_dom_scene_compl: "fds(- A) = - fds(A)"
by (transfer, auto simp add: fun_eq_iff override_on_def)
lemma fun_dom_scene_inter: "fds(A ∩ B) = fds(A) ⊓⇩S fds(B)"
by (simp add: inf_scene_def fun_dom_scene_union[THEN sym] fun_dom_scene_compl[THEN sym])
lemma fun_dom_scene_UNIV: "fds(UNIV) = ⊤⇩S"
by (transfer, auto simp add: fun_eq_iff override_on_def)
lemma fun_dom_scene_indep [simp]:
"fds(A) ⨝⇩S fds(B) ⟷ A ∩ B = {}"
by (transfer, auto simp add: override_on_def fun_eq_iff, meson two_diff)
lemma fun_dom_scene_always_compat [simp]: "fds(A) ##⇩S fds(B)"
by (transfer, simp add: override_on_def fun_eq_iff)
lemma fun_dom_scene_le [simp]: "fds(A) ⊆⇩S fds(B) ⟷ A ⊆ B"
unfolding less_eq_scene_def
by (transfer, auto simp add: override_on_def fun_eq_iff, meson two_diff)
text ‹ Hide implementation details for scenes ›
lifting_update scene.lifting
lifting_forget scene.lifting
end
Theory Lens_Record_Example
theory Lens_Record_Example
imports Optics
begin
text ‹The alphabet command supports syntax illustrated in the following comments.›
alphabet mylens =
x :: nat
y :: string
thm base_more_bij_lens
thm indeps
thm equivs
thm sublenses
thm quotients
thm compositions
lemma mylens_composition:
"x +⇩L y +⇩L more⇩L ≈⇩L 1⇩L" (is "?P ≈⇩L ?Q")
proof -
have "?Q ≈⇩L base⇩L +⇩L more⇩L"
by (simp add: lens_equiv_sym)
also have "... ≈⇩L (x +⇩L y) +⇩L more⇩L"
by (simp add: lens_plus_eq_left)
also have "... ≈⇩L x +⇩L y +⇩L more⇩L"
by (simp add: lens_plus_assoc)
finally show ?thesis
using lens_equiv_sym by auto
qed
lemma mylens_bij_lens:
"bij_lens (x +⇩L y +⇩L more⇩L)"
using bij_lens_equiv_id mylens_composition by auto
alphabet mylens_2 = mylens +
z :: int
k :: "string list"
thm lens_defs
thm base_more_bij_lens
thm indeps
thm equivs
thm sublenses
alphabet mylens_3 = mylens_2 +
n :: real
h :: nat
thm base_more_bij_lens
thm indeps
thm equivs
thm sublenses
end
Theory Lens_State
section ‹State and Lens integration›
theory Lens_State
imports
"HOL-Library.State_Monad"
Lens_Algebra
begin
text ‹Inspired by Haskell's lens package›
definition zoom :: "('a ⟹ 'b) ⇒ ('a, 'c) state ⇒ ('b, 'c) state" where
"zoom l m = State (λb. case run_state m (lens_get l b) of (c, a) ⇒ (c, lens_put l b a))"
definition use :: "('a ⟹ 'b) ⇒ ('b, 'a) state" where
"use l = zoom l State_Monad.get"
definition modify :: "('a ⟹ 'b) ⇒ ('a ⇒ 'a) ⇒ ('b, unit) state" where
"modify l f = zoom l (State_Monad.update f)"
definition assign :: "('a ⟹ 'b) ⇒ 'a ⇒ ('b, unit) state" where
"assign l b = zoom l (State_Monad.set b)"
context begin
qualified abbreviation "add l n ≡ modify l (λx. x + n)"
qualified abbreviation "sub l n ≡ modify l (λx. x - n)"
qualified abbreviation "mul l n ≡ modify l (λx. x * n)"
qualified abbreviation "inc l ≡ add l 1"
qualified abbreviation "dec l ≡ sub l 1"
end
bundle lens_state_notation begin
notation zoom (infixr "⊳" 80)
notation modify (infix "%=" 80)
notation assign (infix ".=" 80)
notation Lens_State.add (infix "+=" 80)
notation Lens_State.sub (infix "-=" 80)
notation Lens_State.mul (infix "*=" 80)
notation Lens_State.inc ("_ ++")
notation Lens_State.dec ("_ --")
end
context includes lens_state_notation begin
lemma zoom_comp1: "l1 ⊳ l2 ⊳ s = (l2 ;⇩L l1) ⊳ s"
unfolding zoom_def lens_comp_def
by (auto split: prod.splits)
lemma zoom_zero[simp]: "zero_lens ⊳ s = s"
unfolding zoom_def zero_lens_def
by simp
lemma zoom_id[simp]: "id_lens ⊳ s = s"
unfolding zoom_def id_lens_def
by simp
end
lemma (in mwb_lens) zoom_comp2[simp]: "zoom x m ⤜ (λa. zoom x (n a)) = zoom x (m ⤜ n)"
unfolding zoom_def State_Monad.bind_def
by (auto split: prod.splits simp: put_get put_put)
lemma (in wb_lens) use_alt_def: "use x = map_state (lens_get x) State_Monad.get"
unfolding State_Monad.get_def use_def zoom_def
by (simp add: comp_def get_put)
lemma (in wb_lens) modify_alt_def: "modify x f = State_Monad.update (update f)"
unfolding modify_def zoom_def lens_update_def State_Monad.update_def State_Monad.get_def State_Monad.set_def State_Monad.bind_def
by (auto)
lemma (in wb_lens) modify_id[simp]: "modify x (λx. x) = State_Monad.return ()"
unfolding lens_update_def modify_alt_def
by (simp add: get_put)
lemma (in mwb_lens) modify_comp[simp]: "bind (modify x f) (λ_. modify x g) = modify x (g ∘ f)"
unfolding modify_def
by simp
end