Theory HOMLinHOLonlyK
subsection‹HOMLinHOLonlyK.thy (slight variation of Figure 3 of \cite{J75})›
text‹Shallow embedding of higher-order modal logic (HOML) in the classical higher-order logic (HOL)
of Isabelle/HOL utilizing the LogiKEy methodology. Here logic K is introduced.›
theory HOMLinHOLonlyK imports Main
begin
nitpick_params[user_axioms,expect=genuine,show_all,format=2,max_genuine=3]
declare[[syntax_ambiguity_warning=false]]
typedecl i
typedecl e
type_synonym σ = "i⇒bool"
type_synonym τ = "e⇒σ"
consts R::"i⇒i⇒bool" ("_❙r_")
abbreviation Mbot::σ ("❙⊥") where "❙⊥ ≡ λw. False"
abbreviation Mtop::σ ("❙⊤") where "❙⊤ ≡ λw. True"
abbreviation Mneg::"σ⇒σ" ("❙¬_" [52]53) where "❙¬φ ≡ λw. ¬(φ w)"
abbreviation Mand::"σ⇒σ⇒σ" (infixl "❙∧" 50) where "φ❙∧ψ ≡ λw. φ w ∧ ψ w"
abbreviation Mor::"σ⇒σ⇒σ" (infixl "❙∨" 49) where "φ❙∨ψ ≡ λw. φ w ∨ ψ w "
abbreviation Mimp::"σ⇒σ⇒σ" (infixr "❙⊃" 48) where "φ❙⊃ψ ≡ λw. φ w ⟶ ψ w"
abbreviation Mequiv::"σ⇒σ⇒σ" (infixl "❙↔" 47) where "φ❙↔ψ ≡ λw. φ w ⟷ ψ w"
abbreviation Mbox::"σ⇒σ" ("❙□_" [54]55) where "❙□φ ≡ λw.∀v. w ❙r v ⟶ φ v"
abbreviation Mdia::"σ⇒σ" ("❙◇_" [54]55) where "❙◇φ ≡ λw.∃v. w ❙r v ∧ φ v"
abbreviation Mprimeq::"'a⇒'a⇒σ" ("_❙=_") where "x❙=y ≡ λw. x=y"
abbreviation Mprimneg::"'a⇒'a⇒σ" ("_❙≠_") where "x❙≠y ≡ λw. x≠y"
abbreviation Mnegpred::"τ⇒τ" ("❙~_") where "❙~Φ ≡ λx.λw. ¬Φ x w"
abbreviation Mconpred::"τ⇒τ⇒τ" (infixl "❙." 50) where "Φ❙.Ψ ≡ λx.λw. Φ x w ∧ Ψ x w"
abbreviation Mexclor::"σ⇒σ⇒σ" (infixl "❙∨⇧e" 49) where "φ❙∨⇧eψ ≡ (φ ❙∨ ψ) ❙∧ ❙¬(φ ❙∧ ψ)"
abbreviation Mallposs::"('a⇒σ)⇒σ" ("❙∀") where "❙∀Φ ≡ λw.∀x. Φ x w"
abbreviation Mallpossb (binder "❙∀" [8]9) where "❙∀x. φ(x) ≡ ❙∀φ"
abbreviation Mexiposs::"('a⇒σ)⇒σ" ("❙∃") where "❙∃Φ ≡ λw.∃x. Φ x w"
abbreviation Mexipossb (binder "❙∃" [8]9) where "❙∃x. φ(x) ≡ ❙∃φ"
consts existsAt::"e⇒σ" ("_❙@_")
abbreviation Mallact::"(e⇒σ)⇒σ" ("❙∀⇧E") where "❙∀⇧EΦ ≡ λw.∀x. x❙@w ⟶ Φ x w"
abbreviation Mallactb (binder "❙∀⇧E" [8]9) where "❙∀⇧Ex. φ(x) ≡ ❙∀⇧Eφ"
abbreviation Mexiact::"(e⇒σ)⇒σ" ("❙∃⇧E") where "❙∃⇧EΦ ≡ λw.∃x. x❙@w ∧ Φ x w"
abbreviation Mexiactb (binder "❙∃⇧E" [8]9) where "❙∃⇧Ex. φ(x) ≡ ❙∃⇧Eφ"
abbreviation Mleibeq::"'a⇒'a⇒σ" ("_❙≡_") where "x❙≡y ≡ ❙∀P. P x ❙⊃ P y"
abbreviation Mvalid::"σ⇒bool" ("⌊_⌋") where "⌊ψ⌋ ≡ ∀w. ψ w"
end