Built with Alectryon. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑ Ctrl+↓ to navigate, Ctrl+πŸ–±οΈ to focus. On Mac, use ⌘ instead of Ctrl.
Hover-Settings: Show types: Show goals:
-- https://leanprover.zulipchat.com/#narrow/stream/270676-lean4/topic/'Who.20Killed.20Aunt.20Agatha'.20puzzle
-- https://gist.github.com/utensil/17159725e5335f6a9688b7443fe0b3aa
-- https://www.tptp.org/cgi-bin/SystemOnTPTP

import Batteries.Logic
-- import Batteries.Tactic.RCases
-- import Mathlib.Tactic

variable (u : Type)
variable (lives : u β†’ Prop)
variable (killed hates richer : u β†’ u β†’ Prop)
variable (agatha butler charles : u)

variable (pel55_1 : βˆƒ x : u, lives x ∧ killed x agatha)
-- variable (pel55_2_1 : lives agatha)
-- variable (pel55_2_2 : lives butler)
-- variable (pel55_2_3 : lives charles)
variable (pel55_3 : βˆ€ x, lives x β†’ x = agatha ∨ x = butler ∨ x = charles)
variable (pel55_4 : βˆ€ x y, killed x y β†’ hates x y)
variable (pel55_5 : βˆ€ x y, killed x y β†’ Β¬ richer x y)
variable (pel55_6 : βˆ€ x, hates agatha x β†’ Β¬ hates charles x)

variable (pel55_7 : βˆ€ x, x β‰  butler β†’ hates agatha x)
variable (pel55_8 : βˆ€ x, Β¬ richer x agatha β†’ hates butler x)
variable (pel55_9 : βˆ€ x, hates agatha x β†’ hates butler x)

variable (pel55_10 : βˆ€ x, βˆƒ y, Β¬ hates x y)
-- variable (pel55_10a : βˆ€ x, βˆƒ y, lives y ∧ Β¬ hates x y)
variable (pel55_11 : agatha β‰  butler)
-- variable (pel55_12 : agatha β‰  charles)
-- variable (pel55_13 : charles β‰  butler)

theorem result : killed agatha agatha := 

Goals accomplished! πŸ™
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_10: βˆ€ (x : u), βˆƒ y, Β¬hates x y
pel55_11: agatha β‰  butler
haa: hates agatha agatha

killed agatha agatha
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_10: βˆ€ (x : u), βˆƒ y, Β¬hates x y
pel55_11: agatha β‰  butler
haa: hates agatha agatha

killed agatha agatha

Goals accomplished! πŸ™
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_10: βˆ€ (x : u), βˆƒ y, Β¬hates x y
pel55_11: agatha β‰  butler
haa: hates agatha agatha
kca: killed charles agatha

False
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_10: βˆ€ (x : u), βˆƒ y, Β¬hates x y
pel55_11: agatha β‰  butler
haa: hates agatha agatha
kca: killed charles agatha
hca: hates charles agatha

False
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_10: βˆ€ (x : u), βˆƒ y, Β¬hates x y
pel55_11: agatha β‰  butler
haa: hates agatha agatha
kca: killed charles agatha
hca: hates charles agatha
nhca: Β¬hates charles agatha

False

Goals accomplished! πŸ™
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_10: βˆ€ (x : u), βˆƒ y, Β¬hates x y
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed n agatha

killed agatha agatha
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_10: βˆ€ (x : u), βˆƒ y, Β¬hates x y
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed n agatha
la: n = agatha

inl
killed agatha agatha
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_10: βˆ€ (x : u), βˆƒ y, Β¬hates x y
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed n agatha
lb: n = butler
killed agatha agatha
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_10: βˆ€ (x : u), βˆƒ y, Β¬hates x y
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed n agatha
lc: n = charles
killed agatha agatha
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_10: βˆ€ (x : u), βˆƒ y, Β¬hates x y
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed n agatha
la: n = agatha

inl
killed agatha agatha
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_10: βˆ€ (x : u), βˆƒ y, Β¬hates x y
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed n agatha
la: n = agatha

inl
killed agatha agatha
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_10: βˆ€ (x : u), βˆƒ y, Β¬hates x y
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed agatha agatha
la: n = agatha

inl
killed agatha agatha
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_10: βˆ€ (x : u), βˆƒ y, Β¬hates x y
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed agatha agatha
la: n = agatha

inl
killed agatha agatha
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_10: βˆ€ (x : u), βˆƒ y, Β¬hates x y
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed agatha agatha
la: n = agatha

inl
killed agatha agatha

Goals accomplished! πŸ™
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_10: βˆ€ (x : u), βˆƒ y, Β¬hates x y
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed n agatha
lb: n = butler

inr.inl
killed agatha agatha
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_10: βˆ€ (x : u), βˆƒ y, Β¬hates x y
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed n agatha
lb: n = butler

inr.inl
killed agatha agatha
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_10: βˆ€ (x : u), βˆƒ y, Β¬hates x y
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed butler agatha
lb: n = butler

inr.inl
killed agatha agatha
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_10: βˆ€ (x : u), βˆƒ y, Β¬hates x y
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed butler agatha
lb: n = butler

inr.inl
killed agatha agatha
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_10: βˆ€ (x : u), βˆƒ y, Β¬hates x y
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed butler agatha
lb: n = butler

inr.inl
killed agatha agatha
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_10: βˆ€ (x : u), βˆƒ y, Β¬hates x y
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed butler agatha
lb: n = butler
nrba: Β¬richer butler agatha

inr.inl
killed agatha agatha
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_10: βˆ€ (x : u), βˆƒ y, Β¬hates x y
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed butler agatha
lb: n = butler
nrba: Β¬richer butler agatha
hbb: hates butler butler

inr.inl
killed agatha agatha
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed butler agatha
lb: n = butler
nrba: Β¬richer butler agatha
hbb: hates butler butler
pel55_10: βˆƒ f, βˆ€ (x : u), Β¬hates x (f x)

inr.inl
killed agatha agatha
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed butler agatha
lb: n = butler
nrba: Β¬richer butler agatha
hbb: hates butler butler
pel55_10: βˆƒ f, βˆ€ (x : u), Β¬hates x (f x)
sk: u β†’ u
nhsk: βˆ€ (x : u), Β¬hates x (sk x)

inr.inl
killed agatha agatha
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed butler agatha
lb: n = butler
nrba: Β¬richer butler agatha
hbb: hates butler butler
pel55_10: βˆƒ f, βˆ€ (x : u), Β¬hates x (f x)
sk: u β†’ u
nhsk: βˆ€ (x : u), Β¬hates x (sk x)
nhbsk: Β¬hates butler (sk butler)

inr.inl
killed agatha agatha
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed butler agatha
lb: n = butler
nrba: Β¬richer butler agatha
hbb: hates butler butler
pel55_10: βˆƒ f, βˆ€ (x : u), Β¬hates x (f x)
sk: u β†’ u
nhsk: βˆ€ (x : u), Β¬hates x (sk x)
nhbsk: Β¬hates butler (sk butler)

inr.inl
killed agatha agatha

Goals accomplished! πŸ™
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed butler agatha
lb: n = butler
nrba: Β¬richer butler agatha
hbb: hates butler butler
pel55_10: βˆƒ f, βˆ€ (x : u), Β¬hates x (f x)
sk: u β†’ u
nhsk: βˆ€ (x : u), Β¬hates x (sk x)
nhbsk: Β¬hates butler (sk butler)
hask: hates agatha (sk butler)

False
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed butler agatha
lb: n = butler
nrba: Β¬richer butler agatha
hbb: hates butler butler
pel55_10: βˆƒ f, βˆ€ (x : u), Β¬hates x (f x)
sk: u β†’ u
nhsk: βˆ€ (x : u), Β¬hates x (sk x)
nhbsk: Β¬hates butler (sk butler)
hask: hates agatha (sk butler)
hbsk: hates butler (sk butler)

False

Goals accomplished! πŸ™
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed butler agatha
lb: n = butler
nrba: Β¬richer butler agatha
hbb: hates butler butler
pel55_10: βˆƒ f, βˆ€ (x : u), Β¬hates x (f x)
sk: u β†’ u
nhsk: βˆ€ (x : u), Β¬hates x (sk x)
nhbsk: Β¬hates butler (sk butler)
nhask: Β¬hates agatha (sk butler)

inr.inl
killed agatha agatha

Goals accomplished! πŸ™
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed butler agatha
lb: n = butler
nrba: Β¬richer butler agatha
hbb: hates butler butler
pel55_10: βˆƒ f, βˆ€ (x : u), Β¬hates x (f x)
sk: u β†’ u
nhsk: βˆ€ (x : u), Β¬hates x (sk x)
nhbsk: Β¬hates butler (sk butler)
nhask: Β¬hates agatha (sk butler)
nesk: Β¬sk butler = butler

False
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed butler agatha
lb: n = butler
nrba: Β¬richer butler agatha
hbb: hates butler butler
pel55_10: βˆƒ f, βˆ€ (x : u), Β¬hates x (f x)
sk: u β†’ u
nhsk: βˆ€ (x : u), Β¬hates x (sk x)
nhbsk: Β¬hates butler (sk butler)
nhask: Β¬hates agatha (sk butler)
nesk: Β¬sk butler = butler
hbsk: hates agatha (sk butler)

False

Goals accomplished! πŸ™
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed butler agatha
lb: n = butler
nrba: Β¬richer butler agatha
hbb: hates butler butler
pel55_10: βˆƒ f, βˆ€ (x : u), Β¬hates x (f x)
sk: u β†’ u
nhsk: βˆ€ (x : u), Β¬hates x (sk x)
nhbsk: Β¬hates butler (sk butler)
nhask: Β¬hates agatha (sk butler)
eqsk: sk butler = butler
nhbb: Β¬hates butler butler

inr.inl
killed agatha agatha

Goals accomplished! πŸ™
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_10: βˆ€ (x : u), βˆƒ y, Β¬hates x y
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed n agatha
lc: n = charles

inr.inr
killed agatha agatha
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_10: βˆ€ (x : u), βˆƒ y, Β¬hates x y
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed n agatha
lc: n = charles

inr.inr
killed agatha agatha
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_10: βˆ€ (x : u), βˆƒ y, Β¬hates x y
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed charles agatha
lc: n = charles

inr.inr
killed agatha agatha
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_10: βˆ€ (x : u), βˆƒ y, Β¬hates x y
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed charles agatha
lc: n = charles

inr.inr
killed agatha agatha
u: Type
lives: u β†’ Prop
killed, hates, richer: u β†’ u β†’ Prop
agatha, butler, charles: u
pel55_1: βˆƒ x, lives x ∧ killed x agatha
pel55_3: βˆ€ (x : u), lives x β†’ x = agatha ∨ x = butler ∨ x = charles
pel55_4: βˆ€ (x y : u), killed x y β†’ hates x y
pel55_5: βˆ€ (x y : u), killed x y β†’ Β¬richer x y
pel55_6: βˆ€ (x : u), hates agatha x β†’ Β¬hates charles x
pel55_7: βˆ€ (x : u), x β‰  butler β†’ hates agatha x
pel55_8: βˆ€ (x : u), Β¬richer x agatha β†’ hates butler x
pel55_9: βˆ€ (x : u), hates agatha x β†’ hates butler x
pel55_10: βˆ€ (x : u), βˆƒ y, Β¬hates x y
pel55_11: agatha β‰  butler
haa: hates agatha agatha
nkca: Β¬killed charles agatha
n: u
ln: lives n
kna: killed charles agatha
lc: n = charles

inr.inr
killed agatha agatha

Goals accomplished! πŸ™

Goals accomplished! πŸ™