Theory CJDDLplus

(*<*)
theory CJDDLplus
  imports Main
begin
nitpick_params[user_axioms=true, show_all, expect=genuine, format = 3]
(*>*)

section ‹Introduction›
text‹\noindent{We present an encoding of an ambitious ethical theory ---Alan Gewirth's "Principle of Generic Consistency (PGC)"--- 
in Isabelle/HOL. The PGC has stirred much attention in philosophy and ethics cite"Beyleveld" and has been proposed as a
potential means to bound the impact of artificial general intelligence (AGI) cite"Kornai".
With our contribution we make a first, important step towards formally assessing the PGC and its potential applications in AI.
Our formalisation utilises the shallow semantical embedding approach cite"J23"
and adapts a recent embedding of dyadic deontic logic in HOL cite"C71" cite"BenzmuellerDDL".
}›

section ‹Semantic Embedding of Carmo and Jones' Dyadic Deontic Logic (DDL) augmented with Kaplanian contexts›

text‹\noindent{We introduce a modification of the semantic embedding developed by Benzm\"uller et al. cite"C71" cite"BenzmuellerDDL"
for the Dyadic Deontic Logic originally presented by Carmo and Jones cite"CJDDL". We extend this embedding
to a two-dimensional semantics as originally presented by David Kaplan cite"Kaplan1979" cite"Kaplan1989".}›

subsection ‹Definition of Types›

typedecl w   ― ‹  Type for possible worlds (Kaplan's "circumstances of evaluation" or "counterfactual situations")  ›
typedecl e   ― ‹  Type for individuals (entities eligible to become agents) ›
typedecl c   ― ‹  Type for Kaplanian "contexts of use" ›
type_synonym wo = "wbool" ― ‹  contents/propositions are identified with their truth-sets ›
type_synonym cwo = "cwo"  ― ‹  sentence meaning (Kaplan's "character") is a function from contexts to contents ›
type_synonym m = "cwo"      ― ‹  we use the letter 'm' for characters (reminiscent of "meaning") ›

subsection ‹Semantic Characterisation of DDL› (*cf. original Carmo and Jones Paper @{cite "CJDDL"} p.290ff*)

subsubsection ‹Basic Set Operations›
abbreviation subset::"wowobool" (infix "" 46) where "α  β  w. α w   β w"
abbreviation intersection::"wowowo" (infixr "" 48) where "α  β  λx. α x  β x"
abbreviation union::"wowowo" (infixr "" 48) where "α  β  λx. α x  β x"
abbreviation complement::"wowo" ("_"[45]46) where "α  λx. ¬α x"
abbreviation instantiated::"wobool" ("_"[45]46) where " φ  x. φ x"
abbreviation setEq::"wowobool" (infix "=s" 46) where "α =s β  x. α x  β x"
abbreviation univSet :: "wo" ("") where "  λw. True"
abbreviation emptySet :: "wo" ("") where "  λw. False"

subsubsection ‹Set-Theoretic Conditions for DDL›

consts
av::"wwo"   ― ‹ set of worlds that are open alternatives (aka. actual versions) of w ›
pv::"wwo"   ― ‹ set of worlds that are possible alternatives (aka. potential versions) of w ›
ob::"wowobool" ― ‹ set of propositions which are obligatory in a given context (of type wo)  ›

axiomatization where
sem_3a: "w. (av w)" and ― ‹  av is serial: in every situation there is always an open alternative ›
sem_4a: "w. av w  pv w" and ― ‹  open alternatives are possible alternatives ›
sem_4b: "w. pv w w" and ― ‹  pv is reflexive: every situation is a possible alternative to itself ›
sem_5a: "X. ¬(ob X )" and ― ‹  contradictions cannot be obligatory ›
sem_5b: "X Y Z. (X  Y) =s (X  Z)  (ob X Y  ob X Z)" and
sem_5c: "X Y Z. (X  Y  Z)  ob X Y  ob X Z  ob X (Y  Z)" and
sem_5d: "X Y Z. (Y  X  ob X Y  X  Z)  ob Z ((Z  (X))  Y)" and
sem_5e: "X Y Z. Y  X  ob X Z  (Y  Z)  ob Y Z"

lemma True nitpick[satisfy] oops ― ‹ model found: axioms are consistent ›

subsubsection ‹Verifying Semantic Conditions›

lemma sem_5b1: "ob X Y  ob X (Y  X)" by (metis (no_types, lifting) sem_5b)
lemma sem_5b2: "(ob X (Y  X)  ob X Y)" by (metis (no_types, lifting) sem_5b) 
lemma sem_5ab: "ob X Y  (X  Y)" by (metis (full_types) sem_5a sem_5b)
lemma sem_5bd1: "Y  X  ob X Y  X  Z  ob Z ((X)  Y)" using sem_5b sem_5d by smt
lemma sem_5bd2: "ob X Y  X  Z  ob Z ((Z  (X))  Y)"  using sem_5b sem_5d  by (smt sem_5b1)  
lemma sem_5bd3: "ob X Y  X  Z  ob Z ((X)  Y)"  by (smt sem_5bd2 sem_5b) 
lemma sem_5bd4: "ob X Y  X  Z  ob Z ((X)  (X   Y))" using sem_5bd3 by auto
lemma sem_5bcd: "(ob X Z  ob Y Z)  ob (X  Y) Z" using sem_5b sem_5c sem_5d oops

(* 5e and 5ab justify redefinition of @{cite "O⟨φ|σ⟩"} as (ob A B)*)
lemma "ob A B   ((A  B)  (X. X  A  (X  B)  ob X B))" using sem_5e sem_5ab by blast

subsection ‹(Shallow) Semantic Embedding of DDL›

subsubsection ‹Basic Propositional Logic›
abbreviation pand::"mmm" (infixr"" 51) where "φψ  λc w. (φ c w)(ψ c w)"
abbreviation por::"mmm" (infixr"" 50) where "φψ  λc w. (φ c w)(ψ c w)"
abbreviation pimp::"mmm" (infix"" 49) where "φψ  λc w. (φ c w)(ψ c w)"
abbreviation pequ::"mmm" (infix"" 48) where "φψ  λc w. (φ c w)(ψ c w)"
abbreviation pnot::"mm" ("¬_" [52]53) where "¬φ  λc w. ¬(φ c w)"

subsubsection ‹Modal Operators›
abbreviation cjboxa :: "mm" ("a_" [52]53) where "aφ  λc w. v. (av w) v  (φ c v)"
abbreviation cjdiaa :: "mm" ("a_" [52]53) where "aφ  λc w. v. (av w) v  (φ c v)"
abbreviation cjboxp :: "mm" ("p_" [52]53) where "pφ  λc w. v. (pv w) v  (φ c v)"
abbreviation cjdiap :: "mm" ("p_" [52]53) where "pφ  λc w. v. (pv w) v  (φ c v)"
abbreviation cjtaut :: "m" ("") where "  λc w. True"
abbreviation cjcontr :: "m" ("") where "  λc w. False"

subsubsection ‹Deontic Operators›
abbreviation cjod :: "mmm" ("O_|_"54) where "Oφ|σ  λc w. ob (σ c) (φ c)"
abbreviation cjoa :: "mm" ("Oa_" [53]54) where "Oaφ  λc w. (ob (av w)) (φ c)  (x. (av w) x  ¬(φ c x))"
abbreviation cjop :: "mm" ("Oi_" [53]54) where "Oiφ  λc w. (ob (pv w)) (φ c)  (x. (pv w) x  ¬(φ c x))"

subsubsection ‹Logical Validity (Classical)›
abbreviation modvalidctx :: "mcbool" ("_M") where "φM  λc. w. φ c w" ― ‹ context-dependent modal validity ›
abbreviation modvalid :: "mbool" ("_") where "φ  c. φM c" ― ‹ general modal validity (modally valid in each context) ›

(*
If we introduce the alternative definition of logical validity below (from Kaplan's LD) instead of the previous one,
we can prove valid most of the following theorems excepting only CJ_7 and CJ_8 and the necessitation rule.

consts World::"c⇒w"  ― ‹ function retrieving the world corresponding to context c (Kaplanian contexts are world-centered) ›        
abbreviation ldtruectx::"m⇒c⇒bool" ("⌊_⌋_") where "⌊φ⌋c ≡ φ c (World c)" ― ‹  truth in the given context ›
abbreviation ldvalid::"m⇒bool" ("⌊_⌋") where "⌊φ⌋ ≡ ∀c. ⌊φ⌋c"    ― ‹  LD validity (true in every context) ›
*)

subsection ‹Verifying the Embedding›

subsubsection ‹Avoiding Modal Collapse›
lemma "P  OaP" nitpick oops ― ‹ (actual) deontic modal collapse is countersatisfiable ›
lemma "P  OiP" nitpick oops ― ‹ (ideal) deontic modal collapse is countersatisfiable ›
lemma "P  aP" nitpick oops ― ‹ alethic modal collapse is countersatisfiable (implies all other necessity operators) ›

subsubsection ‹Necessitation Rule›
lemma NecDDLa: "A  aA"  by simp (* Valid only using classical (not LD) validity*)
lemma NecDDLp:  "A  pA" by simp (* Valid only using classical (not LD) validity*)

subsubsection ‹Lemmas for Semantic Conditions› (* extracted from Benzmüller et al. paper @{cite "BenzmuellerDDL"}*)

abbreviation mboxS5 :: "mm" ("S5_" [52]53) where "S5φ  λc w. v. φ c v"
abbreviation mdiaS5 :: "mm" ("S5_" [52]53) where "S5φ  λc w. v. φ c v"

lemma C_2: "OA | B  S5(B  A)" by (simp add: sem_5ab)
lemma C_3:  "((S5(A  B  C))  OB|A  OC|A)  O(B  C)| A" by (simp add: sem_5c)
lemma C_4: "(S5(A  B)  S5(A  C)  OC|B)  OC|A" using sem_5e by blast
lemma C_5: "S5(A  B)  (OC|A  OC|B)" using C_2 sem_5e by blast
lemma C_6: "S5(C  (A  B))  (OA|C  OB|C)" by (metis sem_5b)
lemma C_7: "OB|A  S5OB|A" by blast 
lemma C_8: "OB|A  OA  B| " using sem_5bd4 by presburger

subsubsection ‹Verifying Axiomatic Characterisation›

text‹\noindent{The following theorems have been taken from the original Carmo and Jones' paper (cite"CJDDL" p.293ff).}›

lemma CJ_3: "pA  aA" by (simp add: sem_4a)
lemma CJ_4: "¬O|A" by (simp add: sem_5a)

lemma CJ_5: "(OB|A  OC|A)  OBC|A" nitpick oops ― ‹ countermodel found ›
lemma CJ_5_minus: "S5(A  B  C)  (OB|A  OC|A)  OBC|A" by (simp add: sem_5c)

lemma CJ_6: "OB|A  OB|AB" by (smt C_2 C_4)
lemma CJ_7: "A  B  OC|A  OC|B" using sem_5ab sem_5e by blast (* Valid only using classical (not Kaplan's indexical) validity*)
lemma CJ_8: "C  (A  B)  OA|C  OB|C" using C_6 by simp (* Valid only using classical (not Kaplan's indexical) validity*)

lemma CJ_9a: "pOB|A  pOB|A" by simp
lemma CJ_9p: "aOB|A  aOB|A" by simp
lemma CJ_9_var_a: "OB|A  aOB|A" by simp
lemma CJ_9_var_b: "OB|A  pOB|A" by simp
lemma CJ_10: "p(A  B  C)  OC|B  OC|AB" by (smt C_4)

lemma CJ_11a: "(OaA  OaB)  Oa(A  B)" nitpick oops ― ‹  countermodel found ›
lemma CJ_11a_var: "a(A  B)  (OaA  OaB)  Oa(A  B)" using sem_5c by auto

lemma CJ_11p: "(OiA  OiB)  Oi(A  B)" nitpick oops ― ‹  countermodel found ›
lemma CJ_11p_var: "p(A  B)  (OiA  OiB)  Oi(A  B)" using sem_5c by auto

lemma CJ_12a: "aA  (¬OaA  ¬Oa(¬A))" using sem_5ab by blast (*using C_2 by blast *)
lemma CJ_12p: "pA  (¬OiA  ¬Oi(¬A))" using sem_5ab by blast (*using C_2 by blast*) 

lemma CJ_13a: "a(A  B)  (OaA  OaB)" using sem_5b by metis (*using C_6 by blast *)
lemma CJ_13p: "p(A  B)  (OiA  OiB)" using sem_5b by metis (*using C_6 by blast *)

lemma CJ_O_O: "OB|A  OA  B|" using sem_5bd4 by presburger

text‹\noindent{An ideal obligation which is actually possible both to fulfill and to violate entails an actual obligation (cite"CJDDL" p.319).}›
lemma CJ_Oi_Oa: "(OiA  aA  a(¬A))  OaA" using sem_5e sem_4a by blast

text‹\noindent{Bridge relations between conditional obligations and actual/ideal obligations:}›
lemma CJ_14a: "OB|A  aA  aB  a¬B  OaB" using sem_5e by blast
lemma CJ_14p: "OB|A  pA  pB  p¬B  OiB" using sem_5e by blast

lemma CJ_15a: "(OB|A  a(A  B)  a(A  ¬B))   Oa(A  B)" using CJ_O_O sem_5e by fastforce (*using CJ_O_O CJ_14a by blast*)
lemma CJ_15p: "(OB|A  p(A  B)  p(A  ¬B))   Oi(A  B)" using CJ_O_O sem_5e by fastforce (*using CJ_O_O CJ_14p by blast*)

(*<*)
end
(*>*)