From 6092e129915c62c6dc51c532a6f3888139eca571 Mon Sep 17 00:00:00 2001 From: dhalilov Date: Mon, 27 Feb 2023 14:01:44 +0100 Subject: [PATCH 01/68] Add group definitions --- .../scala/lisa/mathematics/GroupTheory.scala | 66 +++++++++++++++++++ 1 file changed, 66 insertions(+) create mode 100644 src/main/scala/lisa/mathematics/GroupTheory.scala diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala new file mode 100644 index 00000000..714ced9c --- /dev/null +++ b/src/main/scala/lisa/mathematics/GroupTheory.scala @@ -0,0 +1,66 @@ +package lisa.mathematics + +import lisa.automation.kernel.SimpleSimplifier.Substitution +import lisa.automation.kernel.OLPropositionalSolver.Tautology +import lisa.mathematics.GroupTheory.* +import lisa.mathematics.SetTheory.* + +/** + * Group theory, following Chapter 2 of S. Lang "Undergraduate Algebra". + * + * Book : [[https://link.springer.com/book/10.1007/978-1-4684-9234-7]] + */ +object GroupTheory extends lisa.Main { + + private val G, * = variable + private val x, y, z = variable + private val e, f = variable + + /** + * Bounded quantifiers --- These express the usual `∀x ∈ G ϕ` and `∃x ∈ G ϕ`, for some variables (sets) `x` and `G`, which + * are shorthands for `∀x (x ∈ G ==> ϕ)` and `∃x (x ∈ G ==> ϕ)`, respectively. + */ + def ∀(x: VariableLabel, range: VariableLabel, inner: Formula): Formula = forall(x, in(x, range) ==> inner) + + def ∃(x: VariableLabel, range: VariableLabel, inner: Formula): Formula = exists(x, in(x, range) ==> inner) + + def ∃!(x: VariableLabel, range: VariableLabel, inner: Formula): Formula = existsOne(x, in(x, range) ==> inner) + + /** + * Binary relation --- `*` is a binary relation on `G` if it associates to each pair of elements of `G` + * a unique element in `G`. In other words, `*` is a function `G x G -> G`. + */ + val binaryRelation = DEF(G, *) --> functionFrom(*, cartesianProduct(G, G), G) + + /** + * Shorthand for `x * y`. + */ + val op = DEF(*, x, y) --> app(*, pair(x, y)) + + /** + * Associativity --- `*` is associative (in G) if `(x * y) * z = x * (y * z)` for all `x, y, z` in G. + */ + val associativity = DEF(G, *) --> + ∀(x, G, ∀(y, G, ∀(z, G, op(*, op(*, x, y), z) === op(*, x, op(*, y, z))))) + + /** + * Neutral element --- We say that an element `e` in G is neutral if `e * x = x * e = x` for all `x` in G. + */ + val isNeutral = DEF(G, *, e) --> (in(e, G) /\ ∀(x, G, (op(*, e, x) === x) /\ (op(*, x, e) === x))) + + /** + * Identity existence --- There exists a neutral element `e` in G. + */ + val identityExistence = DEF(G, *) --> ∃(e, G, isNeutral(G, *, e)) + + /** + * Inverse existence --- For all `x` in G, there exists an element `y` in G such that `x * y = y * x = e`. + */ + val inverseExistence = DEF(G, *) --> ∀(x, G, ∃(y, G, isNeutral(G, *, op(*, x, y)) /\ isNeutral(G, *, op(*, y, x)))) + + /** + * Group --- A group (G, *) is a set along with a law of composition `*`, satisfying [[associativity]], [[identityExistence]] + * and [[inverseExistence]]. + */ + val group = DEF(G, *) --> binaryRelation(G, *) /\ associativity(G, *) /\ identityExistence(G, *) /\ inverseExistence(G, *) +} From 032ee855aabc348c19ce13f8c69276abfbfaede4 Mon Sep 17 00:00:00 2001 From: dhalilov Date: Tue, 28 Feb 2023 10:43:54 +0100 Subject: [PATCH 02/68] Correct bounded quantifiers semantics --- src/main/scala/lisa/mathematics/GroupTheory.scala | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala index 714ced9c..0df38e5f 100644 --- a/src/main/scala/lisa/mathematics/GroupTheory.scala +++ b/src/main/scala/lisa/mathematics/GroupTheory.scala @@ -18,13 +18,13 @@ object GroupTheory extends lisa.Main { /** * Bounded quantifiers --- These express the usual `∀x ∈ G ϕ` and `∃x ∈ G ϕ`, for some variables (sets) `x` and `G`, which - * are shorthands for `∀x (x ∈ G ==> ϕ)` and `∃x (x ∈ G ==> ϕ)`, respectively. + * are shorthands for `∀x (x ∈ G ==> ϕ)` and `∃x (x ∈ G /\ ϕ)`, respectively. */ def ∀(x: VariableLabel, range: VariableLabel, inner: Formula): Formula = forall(x, in(x, range) ==> inner) - def ∃(x: VariableLabel, range: VariableLabel, inner: Formula): Formula = exists(x, in(x, range) ==> inner) + def ∃(x: VariableLabel, range: VariableLabel, inner: Formula): Formula = exists(x, in(x, range) /\ inner) - def ∃!(x: VariableLabel, range: VariableLabel, inner: Formula): Formula = existsOne(x, in(x, range) ==> inner) + def ∃!(x: VariableLabel, range: VariableLabel, inner: Formula): Formula = existsOne(x, in(x, range) /\ inner) /** * Binary relation --- `*` is a binary relation on `G` if it associates to each pair of elements of `G` From e8c816c38b046fac6639feb8bcf52ea7c20285b0 Mon Sep 17 00:00:00 2001 From: dhalilov Date: Tue, 28 Feb 2023 10:49:14 +0100 Subject: [PATCH 03/68] Proof of identity uniqueness --- .../scala/lisa/mathematics/GroupTheory.scala | 66 ++++++++++++++++++- 1 file changed, 64 insertions(+), 2 deletions(-) diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala index 0df38e5f..2cb57217 100644 --- a/src/main/scala/lisa/mathematics/GroupTheory.scala +++ b/src/main/scala/lisa/mathematics/GroupTheory.scala @@ -2,8 +2,10 @@ package lisa.mathematics import lisa.automation.kernel.SimpleSimplifier.Substitution import lisa.automation.kernel.OLPropositionalSolver.Tautology +import lisa.mathematics.FirstOrderLogic.equalityTransitivity import lisa.mathematics.GroupTheory.* -import lisa.mathematics.SetTheory.* +import lisa.mathematics.SetTheory.{thenHave, z, *} +import lisa.utils.KernelHelpers.{∀, ∃} /** * Group theory, following Chapter 2 of S. Lang "Undergraduate Algebra". @@ -51,7 +53,7 @@ object GroupTheory extends lisa.Main { /** * Identity existence --- There exists a neutral element `e` in G. */ - val identityExistence = DEF(G, *) --> ∃(e, G, isNeutral(G, *, e)) + val identityExistence = DEF(G, *) --> ∃(e, isNeutral(G, *, e)) /** * Inverse existence --- For all `x` in G, there exists an element `y` in G such that `x * y = y * x = e`. @@ -63,4 +65,64 @@ object GroupTheory extends lisa.Main { * and [[inverseExistence]]. */ val group = DEF(G, *) --> binaryRelation(G, *) /\ associativity(G, *) /\ identityExistence(G, *) /\ inverseExistence(G, *) + + /** + * Identity uniqueness --- In a group (G, *), an identity element is unique, i.e. if both `e * x = x * e = x` and + * `f * x = x * f = x` for all `x`, then `e = f`. + * This justifies calling `e` the identity element. + */ + val identityUniqueness = Theorem( + group(G, *) |- ∃!(e, isNeutral(G, *, e)) + ) { + // Forward direction (trivial) + // Maybe changing [[RightExistsOne]] would remove this trivial step + val forward = have(isNeutral(G, *, e) |- ((e === f) ==> isNeutral(G, *, f))) subproof { + assume(isNeutral(G, *, e)) + thenHave((e === f) |- isNeutral(G, *, f)) by Substitution + thenHave((e === f) ==> isNeutral(G, *, f)) by Restate + } + + // Backward direction + // We prove that if e and f are neutral elements then ef = f = e, where the first equality comes from e's left neutrality, + // and the second equality from f's right neutrality + have((isNeutral(G, *, e), isNeutral(G, *, f)) |- (e === f)) subproof { + // We prove that neutral elements are elements of G, such that * can be applied. + val eMembership = have(isNeutral(G, *, e) |- in(e, G)) subproof { + assume(isNeutral(G, *, e)) + have(thesis) by Tautology.from(isNeutral.definition) + } + val fMembership = have(isNeutral(G, *, f) |- in(f, G)) by Tautology.from(eMembership of e -> f) + + assume(isNeutral(G, *, e)) + assume(isNeutral(G, *, f)) + + // First equality : ef = f + have(∀(x, G, (op(*, e, x) === x) /\ (op(*, x, e) === x))) by Tautology.from(isNeutral.definition) + thenHave(in(f, G) ==> ((op(*, e, f) === f) /\ (op(*, f, e) === f))) by InstantiateForall(f) + val cut1 = thenHave(in(f, G) |- ((op(*, e, f) === f) /\ (op(*, f, e) === f))) by Restate + + have((op(*, e, f) === f) /\ (op(*, f, e) === f)) by Cut(fMembership, cut1) + val firstEq = thenHave(op(*, e, f) === f) by Tautology + + // Second equality : ef = e + val cut2 = have(in(e, G) |- ((op(*, f, e) === e) /\ (op(*, e, f) === e))) by InstFunSchema(Map(e -> f, f -> e))(cut1) + have((op(*, f, e) === e) /\ (op(*, e, f) === e)) by Cut(eMembership, cut2) + val secondEq = thenHave(op(*, e, f) === e) by Tautology + + val eqs = have((op(*, e, f) === f) /\ (op(*, e, f) === e)) by RightAnd(firstEq, secondEq) + val cut3 = have(((op(*, e, f) === f) /\ (op(*, e, f) === e)) |- (e === f)) by Tautology.from(equalityTransitivity of (x -> f, y -> op(*, e, f), z -> e)) + have(e === f) by Cut(eqs, cut3) + } + val backward = thenHave(isNeutral(G, *, e) |- (isNeutral(G, *, f) ==> (e === f))) by Restate + + val existence = have(group(G, *) |- ∃(e, isNeutral(G, *, e))) by Tautology.from(group.definition, identityExistence.definition) + + have(isNeutral(G, *, e) |- ((e === f) <=> isNeutral(G, *, f))) by RightIff(forward, backward) + thenHave(isNeutral(G, *, e) |- ∀(f, (e === f) <=> isNeutral(G, *, f))) by RightForall + thenHave(isNeutral(G, *, e) |- ∃(e, ∀(f, (e === f) <=> isNeutral(G, *, f)))) by RightExists + thenHave(∃(e, isNeutral(G, *, e)) |- ∃(e, ∀(f, (e === f) <=> isNeutral(G, *, f)))) by LeftExists + val uniqueness = thenHave(∃(e, isNeutral(G, *, e)) |- ∃!(e, isNeutral(G, *, e))) by RightExistsOne + + have(group(G, *) |- ∃!(e, isNeutral(G, *, e))) by Cut(existence, uniqueness) + } } From b65230c50050555853e727565e2dcf972c2a4cfd Mon Sep 17 00:00:00 2001 From: dhalilov Date: Fri, 3 Mar 2023 16:54:17 +0100 Subject: [PATCH 04/68] Implement ExistenceAndUniqueness tactic --- .../grouptheory/GroupTheoryTactics.scala | 100 ++++++++++++++++++ .../scala/lisa/mathematics/GroupTheory.scala | 23 +--- 2 files changed, 104 insertions(+), 19 deletions(-) create mode 100644 src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala diff --git a/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala b/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala new file mode 100644 index 00000000..f1652d7c --- /dev/null +++ b/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala @@ -0,0 +1,100 @@ +package lisa.automation.grouptheory + +import lisa.automation.kernel.SimplePropositionalSolver.* +import lisa.automation.kernel.SimpleSimplifier.* +import lisa.mathematics.SetTheory +import lisa.prooflib.BasicStepTactic.* +import lisa.prooflib.Library +import lisa.prooflib.ProofTacticLib.{*, given} +import lisa.prooflib.* +import lisa.settheory.SetTheoryLibrary +import lisa.utils.KernelHelpers.* + +object GroupTheoryTactics { + import lisa.settheory.SetTheoryLibrary.{_, given} + + + /** + *
+   * Γ |- ∃x. φ, Δ   Γ', φ, φ[y/x] |- x = y, Δ'
+   * -------------------------------------------
+   * Γ, Γ' |- ∃!x. φ, Δ, Δ'
+   * 
+ * + * This tactic separates the existence and the uniqueness proofs, which are often easier to prove independently, + * and thus is easier to use than [[RightExistsOne]]. + */ + object ExistenceAndUniqueness extends ProofTactic { + def withParameters(using proof: SetTheoryLibrary.Proof, om: OutputManager)(phi: Formula, x: VariableLabel, y: VariableLabel) + (existence: proof.Fact, uniqueness: proof.Fact)(bot: Sequent): proof.ProofTacticJudgement = { + val existenceSeq = proof.getSequent(existence) + val uniquenessSeq = proof.getSequent(uniqueness) + + lazy val substPhi = substituteVariables(phi, Map[VariableLabel, Term](x -> y)) + lazy val existenceFormula = ∃(x, phi) + lazy val uniqueExistenceFormula = ∃!(x, phi) + + // Checking that all formulas are present + if (x == y) { + proof.InvalidProofTactic("x and y can not be equal.") + } else if (!contains(existenceSeq.right, existenceFormula)) { + proof.InvalidProofTactic(s"Existence sequent conclusion does not contain ∃$x. φ.") + } else if (!contains(uniquenessSeq.left, phi)) { + proof.InvalidProofTactic("Uniqueness sequent premises do not contain φ.") + } else if (!contains(uniquenessSeq.left, substPhi)) { + proof.InvalidProofTactic(s"Uniqueness sequent premises do not contain φ[$y/$x].") + } else if (!contains(uniquenessSeq.right, x === y) && !contains(uniquenessSeq.right, y === x)) { + proof.InvalidProofTactic(s"Uniqueness sequent premises do not contain $x = $y") + } else if (!contains(bot.right, uniqueExistenceFormula)) { + proof.InvalidProofTactic(s"Bottom sequent conclusion does not contain ∃!$x. φ") + } + + // Checking pivots + else if (!isSameSet(existenceSeq.left ++ uniquenessSeq.left, bot.left + phi + substPhi)) { + proof.InvalidProofTactic("Could not infer correct left pivots.") + } else if (!isSameSet(existenceSeq.right ++ uniquenessSeq.right + uniqueExistenceFormula, bot.right + existenceFormula + (x === y))) { + proof.InvalidProofTactic("Could not infer correct right pivots.") + } else { + val steps = SUBPROOF(using proof)(Some(bot)) { + val forward = have(phi |- ((x === y) ==> substPhi)) subproof { + assume(phi) + thenHave((x === y) |- substPhi) by Substitution + thenHave((x === y) ==> substPhi) by Restate + } + + val backward = have(phi |- (substPhi ==> (x === y))) by Restate.from(uniqueness) + + have(phi |- ((x === y) <=> substPhi)) + have(phi |- ((x === y) <=> substPhi)) by RightIff(forward, backward) + thenHave(phi |- ∀(y, (x === y) <=> substPhi)) by RightForall + thenHave(phi |- ∃(x, ∀(y, (x === y) <=> substPhi))) by RightExists + thenHave(∃(x, phi) |- ∃(x, ∀(y, (x === y) <=> substPhi))) by LeftExists + val cut = thenHave(∃(x, phi) |- ∃!(x, phi)) by RightExistsOne + + have(bot) by Cut(existence, cut) + } + + unwrapTactic(steps.judgement.asInstanceOf[proof.ProofTacticJudgement])("Subproof for ExistenceAndUniqueness failed.") + } + } + + def apply(using proof: SetTheoryLibrary.Proof, om: OutputManager)(phi: Formula)(existence: proof.Fact, uniqueness: proof.Fact)(bot: Sequent): proof.ProofTacticJudgement = { + val existenceSeq = proof.getSequent(existence) + val uniquenessSeq = proof.getSequent(uniqueness) + + if (!contains(uniquenessSeq.left, phi)) { + return proof.InvalidProofTactic("Uniqueness sequent premises do not contain φ.") + } + + // Try to infer x and y from the uniqueness sequent right-hand side equality + uniquenessSeq.right.collect { + case PredicateFormula(`equality`, List(Term(x: VariableLabel, _), Term(y: VariableLabel, _))) if x != y => (x, y) + }.collectFirst { + case (x, y) if contains(uniquenessSeq.left, substituteVariables(phi, Map[VariableLabel, Term](x -> y))) => + ExistenceAndUniqueness.withParameters(phi, x, y)(existence, uniqueness)(bot) + case (x, y) if contains(uniquenessSeq.left, substituteVariables(phi, Map[VariableLabel, Term](y -> x))) => + ExistenceAndUniqueness.withParameters(phi, y, x)(existence, uniqueness)(bot) + }.getOrElse(proof.InvalidProofTactic("Could not infer correct variables in uniqueness sequent.")) + } + } +} diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala index 2cb57217..5bdd0be9 100644 --- a/src/main/scala/lisa/mathematics/GroupTheory.scala +++ b/src/main/scala/lisa/mathematics/GroupTheory.scala @@ -1,10 +1,11 @@ package lisa.mathematics +import lisa.automation.grouptheory.GroupTheoryTactics.ExistenceAndUniqueness import lisa.automation.kernel.SimpleSimplifier.Substitution import lisa.automation.kernel.OLPropositionalSolver.Tautology import lisa.mathematics.FirstOrderLogic.equalityTransitivity import lisa.mathematics.GroupTheory.* -import lisa.mathematics.SetTheory.{thenHave, z, *} +import lisa.mathematics.SetTheory.* import lisa.utils.KernelHelpers.{∀, ∃} /** @@ -74,18 +75,9 @@ object GroupTheory extends lisa.Main { val identityUniqueness = Theorem( group(G, *) |- ∃!(e, isNeutral(G, *, e)) ) { - // Forward direction (trivial) - // Maybe changing [[RightExistsOne]] would remove this trivial step - val forward = have(isNeutral(G, *, e) |- ((e === f) ==> isNeutral(G, *, f))) subproof { - assume(isNeutral(G, *, e)) - thenHave((e === f) |- isNeutral(G, *, f)) by Substitution - thenHave((e === f) ==> isNeutral(G, *, f)) by Restate - } - - // Backward direction // We prove that if e and f are neutral elements then ef = f = e, where the first equality comes from e's left neutrality, // and the second equality from f's right neutrality - have((isNeutral(G, *, e), isNeutral(G, *, f)) |- (e === f)) subproof { + val uniqueness = have((isNeutral(G, *, e), isNeutral(G, *, f)) |- (e === f)) subproof { // We prove that neutral elements are elements of G, such that * can be applied. val eMembership = have(isNeutral(G, *, e) |- in(e, G)) subproof { assume(isNeutral(G, *, e)) @@ -113,16 +105,9 @@ object GroupTheory extends lisa.Main { val cut3 = have(((op(*, e, f) === f) /\ (op(*, e, f) === e)) |- (e === f)) by Tautology.from(equalityTransitivity of (x -> f, y -> op(*, e, f), z -> e)) have(e === f) by Cut(eqs, cut3) } - val backward = thenHave(isNeutral(G, *, e) |- (isNeutral(G, *, f) ==> (e === f))) by Restate val existence = have(group(G, *) |- ∃(e, isNeutral(G, *, e))) by Tautology.from(group.definition, identityExistence.definition) - have(isNeutral(G, *, e) |- ((e === f) <=> isNeutral(G, *, f))) by RightIff(forward, backward) - thenHave(isNeutral(G, *, e) |- ∀(f, (e === f) <=> isNeutral(G, *, f))) by RightForall - thenHave(isNeutral(G, *, e) |- ∃(e, ∀(f, (e === f) <=> isNeutral(G, *, f)))) by RightExists - thenHave(∃(e, isNeutral(G, *, e)) |- ∃(e, ∀(f, (e === f) <=> isNeutral(G, *, f)))) by LeftExists - val uniqueness = thenHave(∃(e, isNeutral(G, *, e)) |- ∃!(e, isNeutral(G, *, e))) by RightExistsOne - - have(group(G, *) |- ∃!(e, isNeutral(G, *, e))) by Cut(existence, uniqueness) + have(group(G, *) |- ∃!(e, isNeutral(G, *, e))) by ExistenceAndUniqueness(isNeutral(G, *, e))(existence, uniqueness) } } From 3f0bc81a56a24427b93d94499c0358a5231dfd29 Mon Sep 17 00:00:00 2001 From: dhalilov Date: Mon, 13 Mar 2023 15:55:25 +0100 Subject: [PATCH 05/68] Fix ExistenceAndUniqueness's premises in subproof --- .../lisa/automation/grouptheory/GroupTheoryTactics.scala | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala b/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala index f1652d7c..4c624a5b 100644 --- a/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala +++ b/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala @@ -55,6 +55,8 @@ object GroupTheoryTactics { } else if (!isSameSet(existenceSeq.right ++ uniquenessSeq.right + uniqueExistenceFormula, bot.right + existenceFormula + (x === y))) { proof.InvalidProofTactic("Could not infer correct right pivots.") } else { + val gammaPrime = uniquenessSeq.left.filter(f => !isSame(f, phi) && !isSame(f, substPhi)) + val steps = SUBPROOF(using proof)(Some(bot)) { val forward = have(phi |- ((x === y) ==> substPhi)) subproof { assume(phi) @@ -62,6 +64,10 @@ object GroupTheoryTactics { thenHave((x === y) ==> substPhi) by Restate } + for (f <- gammaPrime) { + assume(f) + } + val backward = have(phi |- (substPhi ==> (x === y))) by Restate.from(uniqueness) have(phi |- ((x === y) <=> substPhi)) From 689e00891e03a6a3b72eeb7b80c81ab9c8737abe Mon Sep 17 00:00:00 2001 From: dhalilov Date: Mon, 13 Mar 2023 15:59:22 +0100 Subject: [PATCH 06/68] Prove uniqueness of inverse element --- .../scala/lisa/mathematics/GroupTheory.scala | 58 +++++++++++++++++++ 1 file changed, 58 insertions(+) diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala index 5bdd0be9..1570a6f0 100644 --- a/src/main/scala/lisa/mathematics/GroupTheory.scala +++ b/src/main/scala/lisa/mathematics/GroupTheory.scala @@ -17,6 +17,7 @@ object GroupTheory extends lisa.Main { private val G, * = variable private val x, y, z = variable + private val u, v, w = variable private val e, f = variable /** @@ -110,4 +111,61 @@ object GroupTheory extends lisa.Main { have(group(G, *) |- ∃!(e, isNeutral(G, *, e))) by ExistenceAndUniqueness(isNeutral(G, *, e))(existence, uniqueness) } + + /** + * Theorem --- The inverse of an element `x` (i.e. `y` such that `x * y = y * x = e`) in `G` is unique. + */ + val inverseUniqueness = Theorem( + group(G, *) |- ∀(x, G, ∃!(y, G, isNeutral(G, *, op(*, x, y)) /\ isNeutral(G, *, op(*, y, x)))) + ) { + have(group(G, *) |- ∀(x, G, ∃(y, G, isNeutral(G, *, op(*, x, y)) /\ isNeutral(G, *, op(*, y, x))))) by Tautology.from(group.definition, inverseExistence.definition) + thenHave(group(G, *) |- (in(x, G) ==> ∃(y, G, isNeutral(G, *, op(*, x, y)) /\ isNeutral(G, *, op(*, y, x))))) by InstantiateForall(x) + val existence = thenHave((group(G, *), in(x, G)) |- ∃(y, G, isNeutral(G, *, op(*, x, y)) /\ isNeutral(G, *, op(*, y, x)))) by Restate + + val phi = in(y, G) /\ (isNeutral(G, *, op(*, x, y)) /\ isNeutral(G, *, op(*, y, x))) + val substPhi = substituteVariables(phi, Map[VariableLabel, Term](y -> z)) + + val uniqueness = have((group(G, *), in(x, G), phi, substPhi) |- (y === z)) subproof { + assume(group(G, *)) + assume(in(x, G)) + assume(phi) + assume(substPhi) + + // We prove the following chain of equalities: + // z = (yx)z = y(xz) = y + // where the first and last equalities come from neutrality, and the second equality from associativity + + // z = (yx)z + have(∀(u, G, (op(*, op(*, y, x), u) === u) /\ (op(*, u, op(*, y, x)) === u))) by Tautology.from(isNeutral.definition of (e -> op(*, y, x))) + thenHave(in(z, G) ==> ((op(*, op(*, y, x), z) === z) /\ (op(*, z, op(*, y, x)) === z))) by InstantiateForall(z) + val eq1 = thenHave(op(*, op(*, y, x), z) === z) by Tautology + + // (yx)z = y(xz) + have(∀(u, G, ∀(v, G, ∀(w, G, op(*, op(*, u, v), w) === op(*, u, op(*, v, w)))))) by Tautology.from(group.definition, associativity.definition) + thenHave(in(y, G) ==> ∀(v, G, ∀(w, G, op(*, op(*, y, v), w) === op(*, y, op(*, v, w))))) by InstantiateForall(y) + thenHave(∀(v, G, ∀(w, G, op(*, op(*, y, v), w) === op(*, y, op(*, v, w))))) by Tautology + thenHave(in(x, G) ==> ∀(w, G, op(*, op(*, y, x), w) === op(*, y, op(*, x, w)))) by InstantiateForall(x) + thenHave(∀(w, G, op(*, op(*, y, x), w) === op(*, y, op(*, x, w)))) by Tautology + thenHave(in(z, G) ==> (op(*, op(*, y, x), z) === op(*, y, op(*, x, z)))) by InstantiateForall(z) + val eq2 = thenHave(op(*, op(*, y, x), z) === op(*, y, op(*, x, z))) by Tautology + + // y(xz) = y + have(∀(u, G, (op(*, op(*, x, z), u) === u) /\ (op(*, u, op(*, x, z)) === u))) by Tautology.from(isNeutral.definition of (e -> op(*, x, z))) + thenHave(in(y, G) ==> ((op(*, op(*, x, z), y) === y) /\ (op(*, y, op(*, x, z)) === y))) by InstantiateForall(y) + val eq3 = thenHave(op(*, y, op(*, x, z)) === y) by Tautology + + // z = y(xz) + val eq4 = have(z === op(*, y, op(*, x, z))) by Tautology.from( + eq1, eq2, equalityTransitivity of (x -> z, y -> op(*, op(*, y, x), z), z -> op(*, y, op(*, x, z))) + ) + + have(z === y) by Tautology.from( + eq3, eq4, equalityTransitivity of (x -> z, y -> op(*, y, op(*, x, z)), z -> y) + ) + } + + have((group(G, *), in(x, G)) |- ∃!(y, phi)) by ExistenceAndUniqueness.withParameters(phi, y, z)(existence, uniqueness) + thenHave(group(G, *) |- (in(x, G) ==> ∃!(y, phi))) by Restate + thenHave(thesis) by RightForall + } } From 4b0e036a128070b9de72a386d5f160f13a676e98 Mon Sep 17 00:00:00 2001 From: dhalilov Date: Wed, 15 Mar 2023 21:51:19 +0100 Subject: [PATCH 07/68] Add existential quantifiers substitutions --- .../lisa/mathematics/FirstOrderLogic.scala | 34 +++++++++++++++++++ 1 file changed, 34 insertions(+) diff --git a/src/main/scala/lisa/mathematics/FirstOrderLogic.scala b/src/main/scala/lisa/mathematics/FirstOrderLogic.scala index d58efda8..400d0e60 100644 --- a/src/main/scala/lisa/mathematics/FirstOrderLogic.scala +++ b/src/main/scala/lisa/mathematics/FirstOrderLogic.scala @@ -17,6 +17,7 @@ object FirstOrderLogic extends lisa.Main { private val c = variable private val p = formulaVariable private val P = predicate(1) + private val Q = predicate(1) /** * Theorem --- A formula is equivalent to itself universally quantified if @@ -65,6 +66,39 @@ object FirstOrderLogic extends lisa.Main { thenHave(thesis) by Restate } + val substitutionInExistenceQuantifier = Theorem( + (∃(x, P(x)), ∀(y, P(y) <=> Q(y))) |- ∃(x, Q(x)) + ) { + have(P(x) |- P(x)) by Hypothesis + val substitution = thenHave((P(x), P(x) <=> Q(x)) |- Q(x)) by RightSubstIff(List((P(x), Q(x))), lambda(p, p)) + + have(∀(y, P(y) <=> Q(y)) |- ∀(y, P(y) <=> Q(y))) by Hypothesis + val equiv = thenHave(∀(y, P(y) <=> Q(y)) |- P(x) <=> Q(x)) by InstantiateForall(x) + + have((P(x), ∀(y, P(y) <=> Q(y))) |- Q(x)) by Cut(equiv, substitution) + thenHave((P(x), ∀(y, P(y) <=> Q(y))) |- ∃(x, Q(x))) by RightExists + thenHave(thesis) by LeftExists + } + + val substitutionInUniquenessQuantifier = Theorem( + (∃!(x, P(x)), ∀(y, P(y) <=> Q(y))) |- ∃!(x, Q(x)) + ) { + have((x === y) <=> P(y) |- (x === y) <=> P(y)) by Hypothesis + thenHave(∀(y, (x === y) <=> P(y)) |- (x === y) <=> P(y)) by LeftForall + val substitution = thenHave((∀(y, (x === y) <=> P(y)), P(y) <=> Q(y)) |- (x === y) <=> Q(y)) by RightSubstIff( + List((P(y), Q(y))), lambda(p, (x === y) <=> p) + ) + + have(∀(y, P(y) <=> Q(y)) |- ∀(y, P(y) <=> Q(y))) by Hypothesis + val equiv = thenHave(∀(y, P(y) <=> Q(y)) |- P(y) <=> Q(y)) by InstantiateForall(y) + + have((∀(y, (x === y) <=> P(y)), ∀(y, P(y) <=> Q(y))) |- (x === y) <=> Q(y)) by Cut(equiv, substitution) + thenHave((∀(y, (x === y) <=> P(y)), ∀(y, P(y) <=> Q(y))) |- ∀(y, (x === y) <=> Q(y))) by RightForall + thenHave((∀(y, (x === y) <=> P(y)), ∀(y, P(y) <=> Q(y))) |- ∃(x, ∀(y, (x === y) <=> Q(y)))) by RightExists + thenHave((∃(x, ∀(y, (x === y) <=> P(y))), ∀(y, P(y) <=> Q(y))) |- ∃(x, ∀(y, (x === y) <=> Q(y)))) by LeftExists + thenHave(thesis) by Restate + } + val equalityTransitivity = Theorem( (x === y) /\ (y === z) |- (x === z) ) { From 53a808a83b0cd6237c14fff4a05d7f3daa26d178 Mon Sep 17 00:00:00 2001 From: dhalilov Date: Wed, 15 Mar 2023 21:52:57 +0100 Subject: [PATCH 08/68] Create Cases tactic --- .../grouptheory/GroupTheoryTactics.scala | 51 +++++++++++++++++-- 1 file changed, 46 insertions(+), 5 deletions(-) diff --git a/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala b/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala index 4c624a5b..3429c88b 100644 --- a/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala +++ b/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala @@ -1,17 +1,17 @@ package lisa.automation.grouptheory +import lisa.automation.kernel.OLPropositionalSolver.Tautology import lisa.automation.kernel.SimplePropositionalSolver.* import lisa.automation.kernel.SimpleSimplifier.* import lisa.mathematics.SetTheory +import lisa.prooflib.* import lisa.prooflib.BasicStepTactic.* -import lisa.prooflib.Library import lisa.prooflib.ProofTacticLib.{*, given} -import lisa.prooflib.* import lisa.settheory.SetTheoryLibrary import lisa.utils.KernelHelpers.* object GroupTheoryTactics { - import lisa.settheory.SetTheoryLibrary.{_, given} + import lisa.settheory.SetTheoryLibrary.{*, given} /** @@ -70,7 +70,6 @@ object GroupTheoryTactics { val backward = have(phi |- (substPhi ==> (x === y))) by Restate.from(uniqueness) - have(phi |- ((x === y) <=> substPhi)) have(phi |- ((x === y) <=> substPhi)) by RightIff(forward, backward) thenHave(phi |- ∀(y, (x === y) <=> substPhi)) by RightForall thenHave(phi |- ∃(x, ∀(y, (x === y) <=> substPhi))) by RightExists @@ -85,7 +84,6 @@ object GroupTheoryTactics { } def apply(using proof: SetTheoryLibrary.Proof, om: OutputManager)(phi: Formula)(existence: proof.Fact, uniqueness: proof.Fact)(bot: Sequent): proof.ProofTacticJudgement = { - val existenceSeq = proof.getSequent(existence) val uniquenessSeq = proof.getSequent(uniqueness) if (!contains(uniquenessSeq.left, phi)) { @@ -103,4 +101,47 @@ object GroupTheoryTactics { }.getOrElse(proof.InvalidProofTactic("Could not infer correct variables in uniqueness sequent.")) } } + + /** + *
+   * Γ, φ |- Δ, Σ   Γ, ¬φ |- Δ, Σ'  
+   * -----------------------------
+   * Γ |- Δ
+   * 
+ * + * + * TODO: Extending the tactic to more general pivots + */ + object Cases extends ProofTactic { + def withParameters(using proof: SetTheoryLibrary.Proof, om: OutputManager)(phi: Formula)(pos: proof.Fact, neg: proof.Fact)(bot: Sequent): proof.ProofTacticJudgement = { + val seqPos = proof.getSequent(pos) + val seqNeg = proof.getSequent(neg) + + if (!(contains(seqPos.left, phi) && contains(seqNeg.left, !phi) && !contains(seqNeg.left, phi)) && + !(contains(seqPos.left, !phi) && contains(seqNeg.left, phi) && !contains(seqNeg.left, !phi))) { + proof.InvalidProofTactic("The given sequent do not contain φ or ¬φ.") + } else { + val gamma = bot.left + val steps = SUBPROOF(using proof)(Some(bot)) { + val excludedMiddle = have(phi \/ !phi) by Tautology + val toCut = have((gamma + (phi \/ !phi)) |- bot.right) by LeftOr(pos, neg) + + have(thesis) by Cut(excludedMiddle, toCut) + } + + unwrapTactic(steps.judgement.asInstanceOf[proof.ProofTacticJudgement])("Subproof for Cases failed.") + } + } + + def apply(using proof: SetTheoryLibrary.Proof, om: OutputManager)(pos: proof.Fact, neg: proof.Fact)(bot: Sequent): proof.ProofTacticJudgement = { + val seqPos = proof.getSequent(pos) + val pivot = seqPos.left -- bot.left + + if (pivot.size != 1) { + proof.InvalidProofTactic("Could not infer correct formula φ.") + } else { + Cases.withParameters(pivot.head)(pos, neg)(bot) + } + } + } } From 07e95106c557a5c3fa9127c9e35770a9d4ebdba8 Mon Sep 17 00:00:00 2001 From: dhalilov Date: Wed, 15 Mar 2023 21:54:17 +0100 Subject: [PATCH 09/68] Create conditional definitions --- .../scala/lisa/mathematics/GroupTheory.scala | 101 +++++++++++++++++- 1 file changed, 97 insertions(+), 4 deletions(-) diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala index 1570a6f0..9cfd78bc 100644 --- a/src/main/scala/lisa/mathematics/GroupTheory.scala +++ b/src/main/scala/lisa/mathematics/GroupTheory.scala @@ -1,12 +1,11 @@ package lisa.mathematics -import lisa.automation.grouptheory.GroupTheoryTactics.ExistenceAndUniqueness -import lisa.automation.kernel.SimpleSimplifier.Substitution +import lisa.automation.grouptheory.GroupTheoryTactics.{Cases, ExistenceAndUniqueness} import lisa.automation.kernel.OLPropositionalSolver.Tautology -import lisa.mathematics.FirstOrderLogic.equalityTransitivity +import lisa.automation.kernel.SimpleSimplifier.Substitution +import lisa.mathematics.FirstOrderLogic.{equalityTransitivity, existsOneImpliesExists, substitutionInUniquenessQuantifier} import lisa.mathematics.GroupTheory.* import lisa.mathematics.SetTheory.* -import lisa.utils.KernelHelpers.{∀, ∃} /** * Group theory, following Chapter 2 of S. Lang "Undergraduate Algebra". @@ -30,6 +29,100 @@ object GroupTheory extends lisa.Main { def ∃!(x: VariableLabel, range: VariableLabel, inner: Formula): Formula = existsOne(x, in(x, range) /\ inner) + /** + * Defines the element that is uniquely given by the uniqueness theorem, or falls back to the error element if the + * assumptions of the theorem are not satisfied. + * + * This is useful in defining specific elements in groups, where their uniqueness (and existence) strongly rely + * on the assumption of the group structure. + * + * TODO This should probably be merged with [[The]] with an additional `orElse` method, to be discussed + */ + def TheConditional(u: VariableLabel, f: Formula)(just: runningSetTheory.Theorem, defaultValue: Term = ∅): The = { + val seq = just.proposition + + if (seq.left.isEmpty) { + The(u, f)(just) + } else { + val prem = if (seq.left.size == 1) seq.left.head else And(seq.left.toSeq: _*) + val completeDef = (prem ==> f) /\ (!prem ==> (u === defaultValue)) + val substF = substituteVariables(completeDef, Map[VariableLabel, Term](u -> defaultValue)) + val substDef = substituteVariables(completeDef, Map[VariableLabel, Term](u -> v)) + + val completeUniquenessTheorem = Theorem( + ∃!(u, completeDef) + ) { + val case1 = have(prem |- ∃!(u, completeDef)) subproof { + // We prove the equivalence f <=> completeDef so that we can substitute it in the uniqueness quantifier + val equiv = have(prem |- ∀(u, f <=> completeDef)) subproof { + have(f |- f) by Hypothesis + thenHave((prem, f) |- f) by Weakening + val left = thenHave(f |- (prem ==> f)) by Restate + + have(prem |- prem) by Hypothesis + thenHave((prem, !prem) |- ()) by LeftNot + thenHave((prem, !prem) |- (u === defaultValue)) by Weakening + val right = thenHave(prem |- (!prem ==> (u === defaultValue))) by Restate + + have((prem, f) |- completeDef) by RightAnd(left, right) + val forward = thenHave(prem |- f ==> completeDef) by Restate + + have(completeDef |- completeDef) by Hypothesis + thenHave((prem, completeDef) |- completeDef) by Weakening + thenHave((prem, completeDef) |- f) by Tautology + val backward = thenHave(prem |- completeDef ==> f) by Restate + + have(prem |- f <=> completeDef) by RightIff(forward, backward) + thenHave(thesis) by RightForall + } + + val substitution = have((∃!(u, f), ∀(u, f <=> completeDef)) |- ∃!(u, completeDef)) by Restate.from( + substitutionInUniquenessQuantifier of(P -> lambda(u, f), Q -> lambda(u, completeDef)) + ) + + val implication = have((prem, ∃!(u, f)) |- ∃!(u, completeDef)) by Cut(equiv, substitution) + have(prem |- ∃!(u, completeDef)) by Cut(just, implication) + } + + val case2 = have(!prem |- ∃!(u, completeDef)) subproof { + val existence = have(!prem |- ∃(u, completeDef)) subproof { + have(!prem |- !prem) by Hypothesis + thenHave((prem, !prem) |- ()) by LeftNot + thenHave((prem, !prem) |- substF) by Weakening + val left = thenHave(!prem |- (prem ==> substF)) by Restate + + have(defaultValue === defaultValue) by RightRefl + thenHave(!prem |- defaultValue === defaultValue) by Weakening + val right = thenHave(!prem ==> (defaultValue === defaultValue)) by Restate + + have(!prem |- (prem ==> substF) /\ (!prem ==> (defaultValue === defaultValue))) by RightAnd(left, right) + thenHave(thesis) by RightExists.withParameters(defaultValue) + } + + val uniqueness = have((!prem, completeDef, substDef) |- (u === v)) subproof { + assume(!prem) + assume(completeDef) + assume(substDef) + + val eq1 = have(u === defaultValue) by Tautology + val eq2 = have(defaultValue === v) by Tautology + val p = have((u === defaultValue) /\ (defaultValue === v)) by RightAnd(eq1, eq2) + + val transitivity = equalityTransitivity of (x -> u, y -> defaultValue, z -> v) + have(thesis) by Cut(p, transitivity) + } + + have(thesis) by ExistenceAndUniqueness(completeDef)(existence, uniqueness) + } + + have(thesis) by Cases(case1, case2) + } + show + + The(u, completeDef)(completeUniquenessTheorem) + } + } + /** * Binary relation --- `*` is a binary relation on `G` if it associates to each pair of elements of `G` * a unique element in `G`. In other words, `*` is a function `G x G -> G`. From 62ece1db5a0eacf7c452f48a2519a58e3fc665a8 Mon Sep 17 00:00:00 2001 From: dhalilov Date: Thu, 16 Mar 2023 14:48:55 +0100 Subject: [PATCH 10/68] Fix TheConditional cut --- src/main/scala/lisa/mathematics/GroupTheory.scala | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala index 9cfd78bc..cee81a30 100644 --- a/src/main/scala/lisa/mathematics/GroupTheory.scala +++ b/src/main/scala/lisa/mathematics/GroupTheory.scala @@ -81,7 +81,8 @@ object GroupTheory extends lisa.Main { ) val implication = have((prem, ∃!(u, f)) |- ∃!(u, completeDef)) by Cut(equiv, substitution) - have(prem |- ∃!(u, completeDef)) by Cut(just, implication) + val uniqueness = have(prem |- ∃!(u, f)) by Restate.from(just) + have(prem |- ∃!(u, completeDef)) by Cut(uniqueness, implication) } val case2 = have(!prem |- ∃!(u, completeDef)) subproof { From 4ae4f1832e4d3bbe381a147275491f7bea419680 Mon Sep 17 00:00:00 2001 From: dhalilov Date: Thu, 16 Mar 2023 14:49:42 +0100 Subject: [PATCH 11/68] Define identity element --- src/main/scala/lisa/mathematics/GroupTheory.scala | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala index cee81a30..18fe74bc 100644 --- a/src/main/scala/lisa/mathematics/GroupTheory.scala +++ b/src/main/scala/lisa/mathematics/GroupTheory.scala @@ -206,6 +206,11 @@ object GroupTheory extends lisa.Main { have(group(G, *) |- ∃!(e, isNeutral(G, *, e))) by ExistenceAndUniqueness(isNeutral(G, *, e))(existence, uniqueness) } + /** + * Defines the identity element of `(G, *)`. + */ + val identity = DEF(G, *) --> TheConditional(e, isNeutral(G, *, e))(identityUniqueness) + /** * Theorem --- The inverse of an element `x` (i.e. `y` such that `x * y = y * x = e`) in `G` is unique. */ From d6b7524aa0e1c59098a681fa707c8b3aff66959e Mon Sep 17 00:00:00 2001 From: dhalilov Date: Thu, 16 Mar 2023 14:50:49 +0100 Subject: [PATCH 12/68] Define inverse element --- src/main/scala/lisa/mathematics/GroupTheory.scala | 11 +++++++---- 1 file changed, 7 insertions(+), 4 deletions(-) diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala index 18fe74bc..50feab8c 100644 --- a/src/main/scala/lisa/mathematics/GroupTheory.scala +++ b/src/main/scala/lisa/mathematics/GroupTheory.scala @@ -215,7 +215,7 @@ object GroupTheory extends lisa.Main { * Theorem --- The inverse of an element `x` (i.e. `y` such that `x * y = y * x = e`) in `G` is unique. */ val inverseUniqueness = Theorem( - group(G, *) |- ∀(x, G, ∃!(y, G, isNeutral(G, *, op(*, x, y)) /\ isNeutral(G, *, op(*, y, x)))) + (group(G, *), in(x, G)) |- ∃!(y, G, isNeutral(G, *, op(*, x, y)) /\ isNeutral(G, *, op(*, y, x))) ) { have(group(G, *) |- ∀(x, G, ∃(y, G, isNeutral(G, *, op(*, x, y)) /\ isNeutral(G, *, op(*, y, x))))) by Tautology.from(group.definition, inverseExistence.definition) thenHave(group(G, *) |- (in(x, G) ==> ∃(y, G, isNeutral(G, *, op(*, x, y)) /\ isNeutral(G, *, op(*, y, x))))) by InstantiateForall(x) @@ -263,8 +263,11 @@ object GroupTheory extends lisa.Main { ) } - have((group(G, *), in(x, G)) |- ∃!(y, phi)) by ExistenceAndUniqueness.withParameters(phi, y, z)(existence, uniqueness) - thenHave(group(G, *) |- (in(x, G) ==> ∃!(y, phi))) by Restate - thenHave(thesis) by RightForall + have(thesis) by ExistenceAndUniqueness.withParameters(phi, y, z)(existence, uniqueness) } + + /** + * Defines the inverse of an element `x` in a group `(G, *)`. + */ + val inverse = DEF(x, G, *) --> TheConditional(y, in(y, G) /\ isNeutral(G, *, op(*, x, y)) /\ isNeutral(G, *, op(*, y, x)))(inverseUniqueness) } From 59f5657b5e48791271dd5553e88487c2d4137c62 Mon Sep 17 00:00:00 2001 From: dhalilov Date: Mon, 20 Mar 2023 10:39:47 +0100 Subject: [PATCH 13/68] Fix missing predicates --- src/main/scala/lisa/mathematics/GroupTheory.scala | 1 + 1 file changed, 1 insertion(+) diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala index 50feab8c..6086e444 100644 --- a/src/main/scala/lisa/mathematics/GroupTheory.scala +++ b/src/main/scala/lisa/mathematics/GroupTheory.scala @@ -18,6 +18,7 @@ object GroupTheory extends lisa.Main { private val x, y, z = variable private val u, v, w = variable private val e, f = variable + private val P, Q = predicate(1) /** * Bounded quantifiers --- These express the usual `∀x ∈ G ϕ` and `∃x ∈ G ϕ`, for some variables (sets) `x` and `G`, which From a2f738c0cdf44a81ff4bc8362692573592656099 Mon Sep 17 00:00:00 2001 From: dhalilov Date: Tue, 21 Mar 2023 22:01:38 +0100 Subject: [PATCH 14/68] Simplify proof of identity uniqueness --- .../scala/lisa/mathematics/GroupTheory.scala | 27 +++++++++---------- 1 file changed, 13 insertions(+), 14 deletions(-) diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala index 6086e444..d1632fe6 100644 --- a/src/main/scala/lisa/mathematics/GroupTheory.scala +++ b/src/main/scala/lisa/mathematics/GroupTheory.scala @@ -171,39 +171,38 @@ object GroupTheory extends lisa.Main { val identityUniqueness = Theorem( group(G, *) |- ∃!(e, isNeutral(G, *, e)) ) { + val existence = have(group(G, *) |- ∃(e, isNeutral(G, *, e))) by Tautology.from(group.definition, identityExistence.definition) + // We prove that if e and f are neutral elements then ef = f = e, where the first equality comes from e's left neutrality, // and the second equality from f's right neutrality val uniqueness = have((isNeutral(G, *, e), isNeutral(G, *, f)) |- (e === f)) subproof { // We prove that neutral elements are elements of G, such that * can be applied. - val eMembership = have(isNeutral(G, *, e) |- in(e, G)) subproof { - assume(isNeutral(G, *, e)) - have(thesis) by Tautology.from(isNeutral.definition) - } - val fMembership = have(isNeutral(G, *, f) |- in(f, G)) by Tautology.from(eMembership of e -> f) + val membership = have(isNeutral(G, *, e) |- in(e, G)) by Tautology.from(isNeutral.definition) assume(isNeutral(G, *, e)) assume(isNeutral(G, *, f)) - // First equality : ef = f + // 1. ef = f have(∀(x, G, (op(*, e, x) === x) /\ (op(*, x, e) === x))) by Tautology.from(isNeutral.definition) thenHave(in(f, G) ==> ((op(*, e, f) === f) /\ (op(*, f, e) === f))) by InstantiateForall(f) val cut1 = thenHave(in(f, G) |- ((op(*, e, f) === f) /\ (op(*, f, e) === f))) by Restate - have((op(*, e, f) === f) /\ (op(*, f, e) === f)) by Cut(fMembership, cut1) + have((op(*, e, f) === f) /\ (op(*, f, e) === f)) by Cut(membership of (e -> f), cut1) val firstEq = thenHave(op(*, e, f) === f) by Tautology - // Second equality : ef = e + // 2. ef = e val cut2 = have(in(e, G) |- ((op(*, f, e) === e) /\ (op(*, e, f) === e))) by InstFunSchema(Map(e -> f, f -> e))(cut1) - have((op(*, f, e) === e) /\ (op(*, e, f) === e)) by Cut(eMembership, cut2) - val secondEq = thenHave(op(*, e, f) === e) by Tautology - val eqs = have((op(*, e, f) === f) /\ (op(*, e, f) === e)) by RightAnd(firstEq, secondEq) - val cut3 = have(((op(*, e, f) === f) /\ (op(*, e, f) === e)) |- (e === f)) by Tautology.from(equalityTransitivity of (x -> f, y -> op(*, e, f), z -> e)) + have((op(*, f, e) === e) /\ (op(*, e, f) === e)) by Cut(membership of (e -> e), cut2) + val secondEq = thenHave(e === op(*, e, f)) by Tautology + + // 3. Conclude by transitivity + val eqs = have((e === op(*, e, f)) /\ (op(*, e, f) === f)) by RightAnd(secondEq, firstEq) + val cut3 = have(((e === op(*, e, f)) /\ (op(*, e, f) === f)) |- (e === f)) by Tautology.from(equalityTransitivity of (x -> e, y -> op(*, e, f), z -> f)) + have(e === f) by Cut(eqs, cut3) } - val existence = have(group(G, *) |- ∃(e, isNeutral(G, *, e))) by Tautology.from(group.definition, identityExistence.definition) - have(group(G, *) |- ∃!(e, isNeutral(G, *, e))) by ExistenceAndUniqueness(isNeutral(G, *, e))(existence, uniqueness) } From 3e6151af380de6e0a463ccd0a26a3924ab6567a6 Mon Sep 17 00:00:00 2001 From: dhalilov Date: Tue, 21 Mar 2023 23:13:08 +0100 Subject: [PATCH 15/68] Fix ExistenceAndUniqueness variable inference --- .../grouptheory/GroupTheoryTactics.scala | 42 +++++++++++++------ 1 file changed, 30 insertions(+), 12 deletions(-) diff --git a/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala b/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala index 3429c88b..70f207e4 100644 --- a/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala +++ b/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala @@ -83,22 +83,40 @@ object GroupTheoryTactics { } } - def apply(using proof: SetTheoryLibrary.Proof, om: OutputManager)(phi: Formula)(existence: proof.Fact, uniqueness: proof.Fact)(bot: Sequent): proof.ProofTacticJudgement = { + def apply(using proof: SetTheoryLibrary.Proof, om: OutputManager)(phi: Formula)(existence: proof.Fact, uniqueness: proof.Fact) + (bot: Sequent): proof.ProofTacticJudgement = { + val existenceSeq = proof.getSequent(existence) val uniquenessSeq = proof.getSequent(uniqueness) - if (!contains(uniquenessSeq.left, phi)) { - return proof.InvalidProofTactic("Uniqueness sequent premises do not contain φ.") + // Try to infer x from the premises + // Specifically, find variables in the correct quantifiers, common to all three sequents + val existsVars = existenceSeq.right.collect { + case BinderFormula(Exists, x, f) if isSame(f, phi) => x + } + if (existsVars.isEmpty) { + return proof.InvalidProofTactic("Missing existential quantifier in the existence sequent.") + } + + val commonVars = bot.right.collect { + case BinderFormula(ExistsOne, x, f) if isSame(f, phi) && existsVars.contains(x) => x } + if (commonVars.isEmpty || commonVars.size > 1) { + return proof.InvalidProofTactic("Could not infer correct variable x in quantifiers.") + } + + val x = commonVars.head - // Try to infer x and y from the uniqueness sequent right-hand side equality - uniquenessSeq.right.collect { - case PredicateFormula(`equality`, List(Term(x: VariableLabel, _), Term(y: VariableLabel, _))) if x != y => (x, y) - }.collectFirst { - case (x, y) if contains(uniquenessSeq.left, substituteVariables(phi, Map[VariableLabel, Term](x -> y))) => - ExistenceAndUniqueness.withParameters(phi, x, y)(existence, uniqueness)(bot) - case (x, y) if contains(uniquenessSeq.left, substituteVariables(phi, Map[VariableLabel, Term](y -> x))) => - ExistenceAndUniqueness.withParameters(phi, y, x)(existence, uniqueness)(bot) - }.getOrElse(proof.InvalidProofTactic("Could not infer correct variables in uniqueness sequent.")) + // Infer y from the equalities in the uniqueness sequent + uniquenessSeq.right.collectFirst { + case PredicateFormula(`equality`, List(Term(`x`, _), Term(y: VariableLabel, _))) + if x != y && contains(uniquenessSeq.left, substituteVariables(phi, Map[VariableLabel, Term](x -> y))) => y + + case PredicateFormula(`equality`, List(Term(y: VariableLabel, _), Term(`x`, _))) + if x != y && contains(uniquenessSeq.left, substituteVariables(phi, Map[VariableLabel, Term](x -> y))) => y + } match { + case Some(y) => ExistenceAndUniqueness.withParameters(phi, x, y)(existence, uniqueness)(bot) + case None => proof.InvalidProofTactic("Could not infer correct variable y in uniqueness sequent.") + } } } From 5fca597dae2109736db5df7d9043c8b9f6333915 Mon Sep 17 00:00:00 2001 From: dhalilov Date: Tue, 21 Mar 2023 23:15:44 +0100 Subject: [PATCH 16/68] Use TacticSubproof --- .../automation/grouptheory/GroupTheoryTactics.scala | 10 +++------- 1 file changed, 3 insertions(+), 7 deletions(-) diff --git a/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala b/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala index 70f207e4..4b98a0a1 100644 --- a/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala +++ b/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala @@ -57,7 +57,7 @@ object GroupTheoryTactics { } else { val gammaPrime = uniquenessSeq.left.filter(f => !isSame(f, phi) && !isSame(f, substPhi)) - val steps = SUBPROOF(using proof)(Some(bot)) { + TacticSubproof { val forward = have(phi |- ((x === y) ==> substPhi)) subproof { assume(phi) thenHave((x === y) |- substPhi) by Substitution @@ -78,8 +78,6 @@ object GroupTheoryTactics { have(bot) by Cut(existence, cut) } - - unwrapTactic(steps.judgement.asInstanceOf[proof.ProofTacticJudgement])("Subproof for ExistenceAndUniqueness failed.") } } @@ -140,14 +138,12 @@ object GroupTheoryTactics { proof.InvalidProofTactic("The given sequent do not contain φ or ¬φ.") } else { val gamma = bot.left - val steps = SUBPROOF(using proof)(Some(bot)) { + TacticSubproof { val excludedMiddle = have(phi \/ !phi) by Tautology val toCut = have((gamma + (phi \/ !phi)) |- bot.right) by LeftOr(pos, neg) - have(thesis) by Cut(excludedMiddle, toCut) + have(bot) by Cut(excludedMiddle, toCut) } - - unwrapTactic(steps.judgement.asInstanceOf[proof.ProofTacticJudgement])("Subproof for Cases failed.") } } From 70e959c00b28a71596cc7491246cebee90d81319 Mon Sep 17 00:00:00 2001 From: dhalilov Date: Tue, 21 Mar 2023 23:55:34 +0100 Subject: [PATCH 17/68] Create isInverse and rewrite inverse uniqueness proof --- .../scala/lisa/mathematics/GroupTheory.scala | 117 ++++++++++++------ 1 file changed, 77 insertions(+), 40 deletions(-) diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala index d1632fe6..a86da24d 100644 --- a/src/main/scala/lisa/mathematics/GroupTheory.scala +++ b/src/main/scala/lisa/mathematics/GroupTheory.scala @@ -152,10 +152,15 @@ object GroupTheory extends lisa.Main { */ val identityExistence = DEF(G, *) --> ∃(e, isNeutral(G, *, e)) + /** + * Inverse element --- `y` is called an inverse of `x` if `x * y = y * x = e`. + */ + val isInverse = DEF(y, x, G, *) --> in(y, G) /\ isNeutral(G, *, op(*, x, y)) /\ isNeutral(G, *, op(*, y, x)) + /** * Inverse existence --- For all `x` in G, there exists an element `y` in G such that `x * y = y * x = e`. */ - val inverseExistence = DEF(G, *) --> ∀(x, G, ∃(y, G, isNeutral(G, *, op(*, x, y)) /\ isNeutral(G, *, op(*, y, x)))) + val inverseExistence = DEF(G, *) --> ∀(x, G, ∃(y, isInverse(y, x, G, *))) /** * Group --- A group (G, *) is a set along with a law of composition `*`, satisfying [[associativity]], [[identityExistence]] @@ -215,59 +220,91 @@ object GroupTheory extends lisa.Main { * Theorem --- The inverse of an element `x` (i.e. `y` such that `x * y = y * x = e`) in `G` is unique. */ val inverseUniqueness = Theorem( - (group(G, *), in(x, G)) |- ∃!(y, G, isNeutral(G, *, op(*, x, y)) /\ isNeutral(G, *, op(*, y, x))) + (group(G, *), in(x, G)) |- ∃!(y, isInverse(y, x, G, *)) ) { - have(group(G, *) |- ∀(x, G, ∃(y, G, isNeutral(G, *, op(*, x, y)) /\ isNeutral(G, *, op(*, y, x))))) by Tautology.from(group.definition, inverseExistence.definition) - thenHave(group(G, *) |- (in(x, G) ==> ∃(y, G, isNeutral(G, *, op(*, x, y)) /\ isNeutral(G, *, op(*, y, x))))) by InstantiateForall(x) - val existence = thenHave((group(G, *), in(x, G)) |- ∃(y, G, isNeutral(G, *, op(*, x, y)) /\ isNeutral(G, *, op(*, y, x)))) by Restate + have(group(G, *) |- ∀(x, G, ∃(y, isInverse(y, x, G, *)))) by Tautology.from(group.definition, inverseExistence.definition) + thenHave(group(G, *) |- (in(x, G) ==> ∃(y, isInverse(y, x, G, *)))) by InstantiateForall(x) + val existence = thenHave((group(G, *), in(x, G)) |- ∃(y, isInverse(y, x, G, *))) by Restate + + // Assume y and z are inverses of x. + // We prove the following chain of equalities: + // z =1= (yx)z =2= y(xz) =3= y + // where equalities come from + // 1. Left neutrality of yx + // 2. Associativity + // 3. Right neutrality of xz + val uniqueness = have((group(G, *), in(x, G), isInverse(y, x, G, *), isInverse(z, x, G, *)) |- (y === z)) subproof { + val inverseMembership = have(isInverse(y, x, G, *) |- in(y, G)) by Tautology.from(isInverse.definition) - val phi = in(y, G) /\ (isNeutral(G, *, op(*, x, y)) /\ isNeutral(G, *, op(*, y, x))) - val substPhi = substituteVariables(phi, Map[VariableLabel, Term](y -> z)) - val uniqueness = have((group(G, *), in(x, G), phi, substPhi) |- (y === z)) subproof { assume(group(G, *)) assume(in(x, G)) - assume(phi) - assume(substPhi) - - // We prove the following chain of equalities: - // z = (yx)z = y(xz) = y - // where the first and last equalities come from neutrality, and the second equality from associativity - - // z = (yx)z - have(∀(u, G, (op(*, op(*, y, x), u) === u) /\ (op(*, u, op(*, y, x)) === u))) by Tautology.from(isNeutral.definition of (e -> op(*, y, x))) - thenHave(in(z, G) ==> ((op(*, op(*, y, x), z) === z) /\ (op(*, z, op(*, y, x)) === z))) by InstantiateForall(z) - val eq1 = thenHave(op(*, op(*, y, x), z) === z) by Tautology - - // (yx)z = y(xz) - have(∀(u, G, ∀(v, G, ∀(w, G, op(*, op(*, u, v), w) === op(*, u, op(*, v, w)))))) by Tautology.from(group.definition, associativity.definition) - thenHave(in(y, G) ==> ∀(v, G, ∀(w, G, op(*, op(*, y, v), w) === op(*, y, op(*, v, w))))) by InstantiateForall(y) - thenHave(∀(v, G, ∀(w, G, op(*, op(*, y, v), w) === op(*, y, op(*, v, w))))) by Tautology - thenHave(in(x, G) ==> ∀(w, G, op(*, op(*, y, x), w) === op(*, y, op(*, x, w)))) by InstantiateForall(x) - thenHave(∀(w, G, op(*, op(*, y, x), w) === op(*, y, op(*, x, w)))) by Tautology - thenHave(in(z, G) ==> (op(*, op(*, y, x), z) === op(*, y, op(*, x, z)))) by InstantiateForall(z) - val eq2 = thenHave(op(*, op(*, y, x), z) === op(*, y, op(*, x, z))) by Tautology - - // y(xz) = y - have(∀(u, G, (op(*, op(*, x, z), u) === u) /\ (op(*, u, op(*, x, z)) === u))) by Tautology.from(isNeutral.definition of (e -> op(*, x, z))) - thenHave(in(y, G) ==> ((op(*, op(*, x, z), y) === y) /\ (op(*, y, op(*, x, z)) === y))) by InstantiateForall(y) - val eq3 = thenHave(op(*, y, op(*, x, z)) === y) by Tautology - - // z = y(xz) - val eq4 = have(z === op(*, y, op(*, x, z))) by Tautology.from( - eq1, eq2, equalityTransitivity of (x -> z, y -> op(*, op(*, y, x), z), z -> op(*, y, op(*, x, z))) + assume(isInverse(y, x, G, *)) + assume(isInverse(z, x, G, *)) + + // 1. (yx)z = z + val leftNeutrality = have((group(G, *), in(x, G), isInverse(y, x, G, *), in(z, G)) |- (op(*, op(*, y, x), z) === z)) subproof { + assume(group(G, *)) + assume(in(x, G)) + assume(isInverse(y, x, G, *)) + assume(in(z, G)) + + have(∀(u, G, (op(*, op(*, y, x), u) === u) /\ (op(*, u, op(*, y, x)) === u))) by Tautology.from(isInverse.definition, isNeutral.definition of (e -> op(*, y, x))) + thenHave(in(z, G) ==> ((op(*, op(*, y, x), z) === z) /\ (op(*, z, op(*, y, x)) === z))) by InstantiateForall(z) + thenHave(op(*, op(*, y, x), z) === z) by Tautology + } + val firstEq = have(op(*, op(*, y, x), z) === z) by Cut(inverseMembership of (y -> z), leftNeutrality) + + // 2. (yx)z = y(xz) + val permuteParentheses = have((group(G, *), in(x, G), in(y, G), in(z, G)) |- (op(*, op(*, y, x), z) === op(*, y, op(*, x, z)))) subproof { + assume(group(G, *)) + assume(in(x, G)) + assume(in(y, G)) + assume(in(z, G)) + + have(∀(u, G, ∀(v, G, ∀(w, G, op(*, op(*, u, v), w) === op(*, u, op(*, v, w)))))) by Tautology.from(group.definition, associativity.definition) + thenHave(in(y, G) ==> ∀(v, G, ∀(w, G, op(*, op(*, y, v), w) === op(*, y, op(*, v, w))))) by InstantiateForall(y) + thenHave(∀(v, G, ∀(w, G, op(*, op(*, y, v), w) === op(*, y, op(*, v, w))))) by Tautology + thenHave(in(x, G) ==> ∀(w, G, op(*, op(*, y, x), w) === op(*, y, op(*, x, w)))) by InstantiateForall(x) + thenHave(∀(w, G, op(*, op(*, y, x), w) === op(*, y, op(*, x, w)))) by Tautology + thenHave(in(z, G) ==> (op(*, op(*, y, x), z) === op(*, y, op(*, x, z)))) by InstantiateForall(z) + thenHave(op(*, op(*, y, x), z) === op(*, y, op(*, x, z))) by Tautology + } + val associativityCut = have((group(G, *), in(x, G) /\ in(y, G) /\ in(z, G)) |- (op(*, op(*, y, x), z) === op(*, y, op(*, x, z)))) by Restate.from(permuteParentheses) + val memberships = have(in(x, G) /\ in(y, G) /\ in(z, G)) by Tautology.from(inverseMembership of (y -> y), inverseMembership of (y -> z)) + val secondEq = have(op(*, op(*, y, x), z) === op(*, y, op(*, x, z))) by Cut(memberships, associativityCut) + + // 3. y(xz) = y + val rightNeutrality = have((group(G, *), in(x, G), in(y, G), isInverse(z, x, G, *)) |- (op(*, y, op(*, x, z)) === y)) subproof { + assume(group(G, *)) + assume(in(x, G)) + assume(in(y, G)) + assume(isInverse(z, x, G, *)) + + have(∀(u, G, (op(*, op(*, x, z), u) === u) /\ (op(*, u, op(*, x, z)) === u))) by Tautology.from(isInverse.definition of (y -> z), isNeutral.definition of (e -> op(*, x, z))) + thenHave(in(y, G) ==> ((op(*, op(*, x, z), y) === y) /\ (op(*, y, op(*, x, z)) === y))) by InstantiateForall(y) + thenHave(op(*, y, op(*, x, z)) === y) by Tautology + } + val thirdEq = have(op(*, y, op(*, x, z)) === y) by Cut(inverseMembership of (y -> y), rightNeutrality) + + // Conclude by transitivity + + // 4. z = y(xz) + val fourthEq = have(z === op(*, y, op(*, x, z))) by Tautology.from( + firstEq, secondEq, equalityTransitivity of (x -> z, y -> op(*, op(*, y, x), z), z -> op(*, y, op(*, x, z))) ) + // 5. z = y have(z === y) by Tautology.from( - eq3, eq4, equalityTransitivity of (x -> z, y -> op(*, y, op(*, x, z)), z -> y) + thirdEq, fourthEq, equalityTransitivity of (x -> z, y -> op(*, y, op(*, x, z)), z -> y) ) } - have(thesis) by ExistenceAndUniqueness.withParameters(phi, y, z)(existence, uniqueness) + have(thesis) by ExistenceAndUniqueness.withParameters(isInverse(y, x, G, *), y, z)(existence, uniqueness) } /** * Defines the inverse of an element `x` in a group `(G, *)`. */ - val inverse = DEF(x, G, *) --> TheConditional(y, in(y, G) /\ isNeutral(G, *, op(*, x, y)) /\ isNeutral(G, *, op(*, y, x)))(inverseUniqueness) + val inverse = DEF(x, G, *) --> TheConditional(y, isInverse(y, x, G, *))(inverseUniqueness) } From 2a0cb721de91f9a4163d110a210bc7ff5c3b12d9 Mon Sep 17 00:00:00 2001 From: dhalilov Date: Wed, 22 Mar 2023 22:12:14 +0100 Subject: [PATCH 18/68] Switch to infix notation for op --- .../scala/lisa/mathematics/GroupTheory.scala | 92 +++++++++---------- 1 file changed, 46 insertions(+), 46 deletions(-) diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala index a86da24d..802b4c1a 100644 --- a/src/main/scala/lisa/mathematics/GroupTheory.scala +++ b/src/main/scala/lisa/mathematics/GroupTheory.scala @@ -134,28 +134,28 @@ object GroupTheory extends lisa.Main { /** * Shorthand for `x * y`. */ - val op = DEF(*, x, y) --> app(*, pair(x, y)) + val op = DEF(x, *, y) --> app(*, pair(x, y)) /** * Associativity --- `*` is associative (in G) if `(x * y) * z = x * (y * z)` for all `x, y, z` in G. */ val associativity = DEF(G, *) --> - ∀(x, G, ∀(y, G, ∀(z, G, op(*, op(*, x, y), z) === op(*, x, op(*, y, z))))) + ∀(x, G, ∀(y, G, ∀(z, G, op(op(x, *, y), *, z) === op(x, *, op(y, *, z))))) /** * Neutral element --- We say that an element `e` in G is neutral if `e * x = x * e = x` for all `x` in G. */ - val isNeutral = DEF(G, *, e) --> (in(e, G) /\ ∀(x, G, (op(*, e, x) === x) /\ (op(*, x, e) === x))) + val isNeutral = DEF(e, G, *) --> (in(e, G) /\ ∀(x, G, (op(e, *, x) === x) /\ (op(x, *, e) === x))) /** * Identity existence --- There exists a neutral element `e` in G. */ - val identityExistence = DEF(G, *) --> ∃(e, isNeutral(G, *, e)) + val identityExistence = DEF(G, *) --> ∃(e, isNeutral(e, G, *)) /** * Inverse element --- `y` is called an inverse of `x` if `x * y = y * x = e`. */ - val isInverse = DEF(y, x, G, *) --> in(y, G) /\ isNeutral(G, *, op(*, x, y)) /\ isNeutral(G, *, op(*, y, x)) + val isInverse = DEF(y, x, G, *) --> in(y, G) /\ isNeutral(op(x, *, y), G, *) /\ isNeutral(op(y, *, x), G, *) /** * Inverse existence --- For all `x` in G, there exists an element `y` in G such that `x * y = y * x = e`. @@ -174,47 +174,47 @@ object GroupTheory extends lisa.Main { * This justifies calling `e` the identity element. */ val identityUniqueness = Theorem( - group(G, *) |- ∃!(e, isNeutral(G, *, e)) + group(G, *) |- ∃!(e, isNeutral(e, G, *)) ) { - val existence = have(group(G, *) |- ∃(e, isNeutral(G, *, e))) by Tautology.from(group.definition, identityExistence.definition) + val existence = have(group(G, *) |- ∃(e, isNeutral(e, G, *))) by Tautology.from(group.definition, identityExistence.definition) // We prove that if e and f are neutral elements then ef = f = e, where the first equality comes from e's left neutrality, // and the second equality from f's right neutrality - val uniqueness = have((isNeutral(G, *, e), isNeutral(G, *, f)) |- (e === f)) subproof { + val uniqueness = have((isNeutral(e, G, *), isNeutral(f, G, *)) |- (e === f)) subproof { // We prove that neutral elements are elements of G, such that * can be applied. - val membership = have(isNeutral(G, *, e) |- in(e, G)) by Tautology.from(isNeutral.definition) + val membership = have(isNeutral(e, G, *) |- in(e, G)) by Tautology.from(isNeutral.definition) - assume(isNeutral(G, *, e)) - assume(isNeutral(G, *, f)) + assume(isNeutral(e, G, *)) + assume(isNeutral(f, G, *)) // 1. ef = f - have(∀(x, G, (op(*, e, x) === x) /\ (op(*, x, e) === x))) by Tautology.from(isNeutral.definition) - thenHave(in(f, G) ==> ((op(*, e, f) === f) /\ (op(*, f, e) === f))) by InstantiateForall(f) - val cut1 = thenHave(in(f, G) |- ((op(*, e, f) === f) /\ (op(*, f, e) === f))) by Restate + have(∀(x, G, (op(e, *, x) === x) /\ (op(x, *, e) === x))) by Tautology.from(isNeutral.definition) + thenHave(in(f, G) ==> ((op(e, *, f) === f) /\ (op(f, *, e) === f))) by InstantiateForall(f) + val cut1 = thenHave(in(f, G) |- ((op(e, *, f) === f) /\ (op(f, *, e) === f))) by Restate - have((op(*, e, f) === f) /\ (op(*, f, e) === f)) by Cut(membership of (e -> f), cut1) - val firstEq = thenHave(op(*, e, f) === f) by Tautology + have((op(e, *, f) === f) /\ (op(f, *, e) === f)) by Cut(membership of (e -> f), cut1) + val firstEq = thenHave(op(e, *, f) === f) by Tautology // 2. ef = e - val cut2 = have(in(e, G) |- ((op(*, f, e) === e) /\ (op(*, e, f) === e))) by InstFunSchema(Map(e -> f, f -> e))(cut1) + val cut2 = have(in(e, G) |- ((op(f, *, e) === e) /\ (op(e, *, f) === e))) by InstFunSchema(Map(e -> f, f -> e))(cut1) - have((op(*, f, e) === e) /\ (op(*, e, f) === e)) by Cut(membership of (e -> e), cut2) - val secondEq = thenHave(e === op(*, e, f)) by Tautology + have((op(f, *, e) === e) /\ (op(e, *, f) === e)) by Cut(membership of (e -> e), cut2) + val secondEq = thenHave(e === op(e, *, f)) by Tautology // 3. Conclude by transitivity - val eqs = have((e === op(*, e, f)) /\ (op(*, e, f) === f)) by RightAnd(secondEq, firstEq) - val cut3 = have(((e === op(*, e, f)) /\ (op(*, e, f) === f)) |- (e === f)) by Tautology.from(equalityTransitivity of (x -> e, y -> op(*, e, f), z -> f)) + val eqs = have((e === op(e, *, f)) /\ (op(e, *, f) === f)) by RightAnd(secondEq, firstEq) + val cut3 = have(((e === op(e, *, f)) /\ (op(e, *, f) === f)) |- (e === f)) by Tautology.from(equalityTransitivity of (x -> e, y -> op(e, *, f), z -> f)) have(e === f) by Cut(eqs, cut3) } - have(group(G, *) |- ∃!(e, isNeutral(G, *, e))) by ExistenceAndUniqueness(isNeutral(G, *, e))(existence, uniqueness) + have(group(G, *) |- ∃!(e, isNeutral(e, G, *))) by ExistenceAndUniqueness(isNeutral(e, G, *))(existence, uniqueness) } /** * Defines the identity element of `(G, *)`. */ - val identity = DEF(G, *) --> TheConditional(e, isNeutral(G, *, e))(identityUniqueness) + val identity = DEF(G, *) --> TheConditional(e, isNeutral(e, G, *))(identityUniqueness) /** * Theorem --- The inverse of an element `x` (i.e. `y` such that `x * y = y * x = e`) in `G` is unique. @@ -243,60 +243,60 @@ object GroupTheory extends lisa.Main { assume(isInverse(z, x, G, *)) // 1. (yx)z = z - val leftNeutrality = have((group(G, *), in(x, G), isInverse(y, x, G, *), in(z, G)) |- (op(*, op(*, y, x), z) === z)) subproof { + val leftNeutrality = have((group(G, *), in(x, G), isInverse(y, x, G, *), in(z, G)) |- (op(op(y, *, x), *, z) === z)) subproof { assume(group(G, *)) assume(in(x, G)) assume(isInverse(y, x, G, *)) assume(in(z, G)) - have(∀(u, G, (op(*, op(*, y, x), u) === u) /\ (op(*, u, op(*, y, x)) === u))) by Tautology.from(isInverse.definition, isNeutral.definition of (e -> op(*, y, x))) - thenHave(in(z, G) ==> ((op(*, op(*, y, x), z) === z) /\ (op(*, z, op(*, y, x)) === z))) by InstantiateForall(z) - thenHave(op(*, op(*, y, x), z) === z) by Tautology + have(∀(u, G, (op(op(y, *, x), *, u) === u) /\ (op(u, *, op(y, *, x)) === u))) by Tautology.from(isInverse.definition, isNeutral.definition of (e -> op(y, *, x))) + thenHave(in(z, G) ==> ((op(op(y, *, x), *, z) === z) /\ (op(z, *, op(y, *, x)) === z))) by InstantiateForall(z) + thenHave(op(op(y, *, x), *, z) === z) by Tautology } - val firstEq = have(op(*, op(*, y, x), z) === z) by Cut(inverseMembership of (y -> z), leftNeutrality) + val firstEq = have(op(op(y, *, x), *, z) === z) by Cut(inverseMembership of (y -> z), leftNeutrality) // 2. (yx)z = y(xz) - val permuteParentheses = have((group(G, *), in(x, G), in(y, G), in(z, G)) |- (op(*, op(*, y, x), z) === op(*, y, op(*, x, z)))) subproof { + val permuteParentheses = have((group(G, *), in(x, G), in(y, G), in(z, G)) |- (op(op(y, *, x), *, z) === op(y, *, op(x, *, z)))) subproof { assume(group(G, *)) assume(in(x, G)) assume(in(y, G)) assume(in(z, G)) - have(∀(u, G, ∀(v, G, ∀(w, G, op(*, op(*, u, v), w) === op(*, u, op(*, v, w)))))) by Tautology.from(group.definition, associativity.definition) - thenHave(in(y, G) ==> ∀(v, G, ∀(w, G, op(*, op(*, y, v), w) === op(*, y, op(*, v, w))))) by InstantiateForall(y) - thenHave(∀(v, G, ∀(w, G, op(*, op(*, y, v), w) === op(*, y, op(*, v, w))))) by Tautology - thenHave(in(x, G) ==> ∀(w, G, op(*, op(*, y, x), w) === op(*, y, op(*, x, w)))) by InstantiateForall(x) - thenHave(∀(w, G, op(*, op(*, y, x), w) === op(*, y, op(*, x, w)))) by Tautology - thenHave(in(z, G) ==> (op(*, op(*, y, x), z) === op(*, y, op(*, x, z)))) by InstantiateForall(z) - thenHave(op(*, op(*, y, x), z) === op(*, y, op(*, x, z))) by Tautology + have(∀(u, G, ∀(v, G, ∀(w, G, op(op(u, *, v), *, w) === op(u, *, op(v, *, w)))))) by Tautology.from(group.definition, associativity.definition) + thenHave(in(y, G) ==> ∀(v, G, ∀(w, G, op(op(y, *, v), *, w) === op(y, *, op(v, *, w))))) by InstantiateForall(y) + thenHave(∀(v, G, ∀(w, G, op(op(y, *, v), *, w) === op(y, *, op(v, *, w))))) by Tautology + thenHave(in(x, G) ==> ∀(w, G, op(op(y, *, x), *, w) === op(y, *, op(x, *, w)))) by InstantiateForall(x) + thenHave(∀(w, G, op(op(y, *, x), *, w) === op(y, *, op(x, *, w)))) by Tautology + thenHave(in(z, G) ==> (op(op(y, *, x), *, z) === op(y, *, op(x, *, z)))) by InstantiateForall(z) + thenHave(op(op(y, *, x), *, z) === op(y, *, op(x, *, z))) by Tautology } - val associativityCut = have((group(G, *), in(x, G) /\ in(y, G) /\ in(z, G)) |- (op(*, op(*, y, x), z) === op(*, y, op(*, x, z)))) by Restate.from(permuteParentheses) + val associativityCut = have((group(G, *), in(x, G) /\ in(y, G) /\ in(z, G)) |- (op(op(y, *, x), *, z) === op(y, *, op(x, *, z)))) by Restate.from(permuteParentheses) val memberships = have(in(x, G) /\ in(y, G) /\ in(z, G)) by Tautology.from(inverseMembership of (y -> y), inverseMembership of (y -> z)) - val secondEq = have(op(*, op(*, y, x), z) === op(*, y, op(*, x, z))) by Cut(memberships, associativityCut) + val secondEq = have(op(op(y, *, x), *, z) === op(y, *, op(x, *, z))) by Cut(memberships, associativityCut) // 3. y(xz) = y - val rightNeutrality = have((group(G, *), in(x, G), in(y, G), isInverse(z, x, G, *)) |- (op(*, y, op(*, x, z)) === y)) subproof { + val rightNeutrality = have((group(G, *), in(x, G), in(y, G), isInverse(z, x, G, *)) |- (op(y, *, op(x, *, z)) === y)) subproof { assume(group(G, *)) assume(in(x, G)) assume(in(y, G)) assume(isInverse(z, x, G, *)) - have(∀(u, G, (op(*, op(*, x, z), u) === u) /\ (op(*, u, op(*, x, z)) === u))) by Tautology.from(isInverse.definition of (y -> z), isNeutral.definition of (e -> op(*, x, z))) - thenHave(in(y, G) ==> ((op(*, op(*, x, z), y) === y) /\ (op(*, y, op(*, x, z)) === y))) by InstantiateForall(y) - thenHave(op(*, y, op(*, x, z)) === y) by Tautology + have(∀(u, G, (op(op(x, *, z), *, u) === u) /\ (op(u, *, op(x, *, z)) === u))) by Tautology.from(isInverse.definition of (y -> z), isNeutral.definition of (e -> op(x, *, z))) + thenHave(in(y, G) ==> ((op(op(x, *, z), *, y) === y) /\ (op(y, *, op(x, *, z)) === y))) by InstantiateForall(y) + thenHave(op(y, *, op(x, *, z)) === y) by Tautology } - val thirdEq = have(op(*, y, op(*, x, z)) === y) by Cut(inverseMembership of (y -> y), rightNeutrality) + val thirdEq = have(op(y, *, op(x, *, z)) === y) by Cut(inverseMembership of (y -> y), rightNeutrality) // Conclude by transitivity // 4. z = y(xz) - val fourthEq = have(z === op(*, y, op(*, x, z))) by Tautology.from( - firstEq, secondEq, equalityTransitivity of (x -> z, y -> op(*, op(*, y, x), z), z -> op(*, y, op(*, x, z))) + val fourthEq = have(z === op(y, *, op(x, *, z))) by Tautology.from( + firstEq, secondEq, equalityTransitivity of (x -> z, y -> op(op(y, *, x), *, z), z -> op(y, *, op(x, *, z))) ) // 5. z = y have(z === y) by Tautology.from( - thirdEq, fourthEq, equalityTransitivity of (x -> z, y -> op(*, y, op(*, x, z)), z -> y) + thirdEq, fourthEq, equalityTransitivity of (x -> z, y -> op(y, *, op(x, *, z)), z -> y) ) } From fe62a0e20164ceef665072899e149fde57e93cde Mon Sep 17 00:00:00 2001 From: dhalilov Date: Wed, 22 Mar 2023 22:13:22 +0100 Subject: [PATCH 19/68] Prove symmetry of inverse --- .../scala/lisa/mathematics/GroupTheory.scala | 22 +++++++++++++++++++ 1 file changed, 22 insertions(+) diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala index 802b4c1a..0a0c1cbd 100644 --- a/src/main/scala/lisa/mathematics/GroupTheory.scala +++ b/src/main/scala/lisa/mathematics/GroupTheory.scala @@ -307,4 +307,26 @@ object GroupTheory extends lisa.Main { * Defines the inverse of an element `x` in a group `(G, *)`. */ val inverse = DEF(x, G, *) --> TheConditional(y, isInverse(y, x, G, *))(inverseUniqueness) + + /** + * Theorem --- `y` is the inverse of `x` iff `x` is the inverse of `y` + */ + val inverseSymmetry = Theorem( + group(G, *) |- ∀(x, G, isInverse(y, x, G, *) ==> isInverse(x, y, G, *)) + ) { + have((group(G, *), in(x, G), isInverse(y, x, G, *)) |- isInverse(x, y, G, *)) subproof { + assume(group(G, *)) + assume(in(x, G)) + assume(isInverse(y, x, G, *)) + + have(in(y, G) /\ isNeutral(op(x, *, y), G, *) /\ isNeutral(op(y, *, x), G, *)) by Tautology.from(isInverse.definition) + thenHave(isNeutral(op(y, *, x), G, *) /\ isNeutral(op(x, *, y), G, *)) by Tautology + val definition = thenHave(in(x, G) /\ isNeutral(op(y, *, x), G, *) /\ isNeutral(op(x, *, y), G, *)) by Tautology + + have(thesis) by Tautology.from(definition, isInverse.definition of (y -> x, x -> y)) + } + thenHave((group(G, *), in(x, G)) |- isInverse(y, x, G, *) ==> isInverse(x, y, G, *)) by Restate + thenHave(group(G, *) |- in(x, G) ==> (isInverse(y, x, G, *) ==> isInverse(x, y, G, *))) by Restate + thenHave(thesis) by RightForall + } } From 19055b8a772d28142fd86677e8072363d6b19c48 Mon Sep 17 00:00:00 2001 From: dhalilov Date: Fri, 24 Mar 2023 18:34:16 +0100 Subject: [PATCH 20/68] Prove involution of inverse --- .../scala/lisa/mathematics/GroupTheory.scala | 34 +++++++++++++++++++ 1 file changed, 34 insertions(+) diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala index 0a0c1cbd..fd883777 100644 --- a/src/main/scala/lisa/mathematics/GroupTheory.scala +++ b/src/main/scala/lisa/mathematics/GroupTheory.scala @@ -329,4 +329,38 @@ object GroupTheory extends lisa.Main { thenHave(group(G, *) |- in(x, G) ==> (isInverse(y, x, G, *) ==> isInverse(x, y, G, *))) by Restate thenHave(thesis) by RightForall } + + /** + * Involution of the inverse -- For all `x`, `inverse(inverse(x)) = x`. + */ + val inverseIsInvolutive = Theorem( + group(G, *) |- ∀(x, G, inverse(inverse(x, G, *), G, *) === x) + ) { + assume(group(G, *)) + + // We use the symmetry of the inverse to prove that x is an inverse of inverse(x) + // We finally conclude by uniqueness + + have(in(x, G) |- ∀(y, (y === inverse(x, G, *)) <=> isInverse(y, x, G, *))) by Tautology.from(inverseUniqueness, inverse.definition) + val uniqueness = thenHave(in(x, G) |- ((y === inverse(x, G, *)) <=> isInverse(y, x, G, *))) by InstantiateForall(y) + + have(in(x, G) |- ((inverse(x, G, *) === inverse(x, G, *)) <=> isInverse(inverse(x, G, *), x, G, *))) by Restate.from(uniqueness of (y -> inverse(x, G, *))) + val byDef = thenHave(in(x, G) |- isInverse(inverse(x, G, *), x, G, *)) by Restate + val membership = have(in(x, G) |- in(inverse(x, G, *), G)) by Tautology.from(byDef, isInverse.definition of (y -> inverse(x, G, *))) + + have(∀(z, G, isInverse(inverse(x, G, *), z, G, *) ==> isInverse(z, inverse(x, G, *), G, *))) by Tautology.from(inverseSymmetry of (y -> inverse(x, G, *))) + thenHave(in(x, G) ==> (isInverse(inverse(x, G, *), x, G, *) ==> isInverse(x, inverse(x, G, *), G, *))) by InstantiateForall(x) + val symmetry = thenHave((in(x, G), isInverse(inverse(x, G, *), x, G, *)) |- isInverse(x, inverse(x, G, *), G, *)) by Restate + + val symmetricInverse = have(in(x, G) |- isInverse(x, inverse(x, G, *), G, *)) by Cut(byDef, symmetry) + + // Conclude by uniqueness + val u = inverse(inverse(x, G, *), G, *) + val iff = have(in(inverse(x, G, *), G) |- ((x === u) <=> isInverse(x, inverse(x, G, *), G, *))) by Restate.from(uniqueness of (x -> inverse(x, G, *), y -> x)) + val eq = have((in(x, G), in(inverse(x, G, *), G)) |- (x === u)) by Tautology.from(symmetricInverse, iff) + + have(in(x, G) |- (x === u)) by Cut(membership, eq) + thenHave(in(x, G) ==> (x === u)) by Restate + thenHave(thesis) by RightForall + } } From 522273bc28c585c7a12a8b1c80f5d34005337508 Mon Sep 17 00:00:00 2001 From: dhalilov Date: Fri, 24 Mar 2023 18:34:55 +0100 Subject: [PATCH 21/68] Optimize imports --- src/main/scala/lisa/mathematics/GroupTheory.scala | 2 -- 1 file changed, 2 deletions(-) diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala index fd883777..9304bf69 100644 --- a/src/main/scala/lisa/mathematics/GroupTheory.scala +++ b/src/main/scala/lisa/mathematics/GroupTheory.scala @@ -2,9 +2,7 @@ package lisa.mathematics import lisa.automation.grouptheory.GroupTheoryTactics.{Cases, ExistenceAndUniqueness} import lisa.automation.kernel.OLPropositionalSolver.Tautology -import lisa.automation.kernel.SimpleSimplifier.Substitution import lisa.mathematics.FirstOrderLogic.{equalityTransitivity, existsOneImpliesExists, substitutionInUniquenessQuantifier} -import lisa.mathematics.GroupTheory.* import lisa.mathematics.SetTheory.* /** From f205f09501883ea8be0842d7dc92119da13d6262 Mon Sep 17 00:00:00 2001 From: dhalilov Date: Fri, 24 Mar 2023 19:20:30 +0100 Subject: [PATCH 22/68] Modularize proof of involution of inverse --- .../scala/lisa/mathematics/GroupTheory.scala | 43 +++++++++++++------ 1 file changed, 30 insertions(+), 13 deletions(-) diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala index 9304bf69..1afbc06c 100644 --- a/src/main/scala/lisa/mathematics/GroupTheory.scala +++ b/src/main/scala/lisa/mathematics/GroupTheory.scala @@ -336,26 +336,43 @@ object GroupTheory extends lisa.Main { ) { assume(group(G, *)) - // We use the symmetry of the inverse to prove that x is an inverse of inverse(x) - // We finally conclude by uniqueness + val inverseCharacterization = have((group(G, *), in(x, G)) |- ((y === inverse(x, G, *)) <=> isInverse(y, x, G, *))) subproof { + assume(group(G, *)) - have(in(x, G) |- ∀(y, (y === inverse(x, G, *)) <=> isInverse(y, x, G, *))) by Tautology.from(inverseUniqueness, inverse.definition) - val uniqueness = thenHave(in(x, G) |- ((y === inverse(x, G, *)) <=> isInverse(y, x, G, *))) by InstantiateForall(y) + have(in(x, G) |- ∀(y, (y === inverse(x, G, *)) <=> isInverse(y, x, G, *))) by Tautology.from(inverseUniqueness, inverse.definition) + thenHave(thesis) by InstantiateForall(y) + } - have(in(x, G) |- ((inverse(x, G, *) === inverse(x, G, *)) <=> isInverse(inverse(x, G, *), x, G, *))) by Restate.from(uniqueness of (y -> inverse(x, G, *))) - val byDef = thenHave(in(x, G) |- isInverse(inverse(x, G, *), x, G, *)) by Restate - val membership = have(in(x, G) |- in(inverse(x, G, *), G)) by Tautology.from(byDef, isInverse.definition of (y -> inverse(x, G, *))) + // Prove that inverse(x) is an inverse of x + val inverseIsInverse = have((group(G, *), in(x, G)) |- isInverse(inverse(x, G, *), x, G, *)) subproof { + assume(group(G, *)) - have(∀(z, G, isInverse(inverse(x, G, *), z, G, *) ==> isInverse(z, inverse(x, G, *), G, *))) by Tautology.from(inverseSymmetry of (y -> inverse(x, G, *))) - thenHave(in(x, G) ==> (isInverse(inverse(x, G, *), x, G, *) ==> isInverse(x, inverse(x, G, *), G, *))) by InstantiateForall(x) - val symmetry = thenHave((in(x, G), isInverse(inverse(x, G, *), x, G, *)) |- isInverse(x, inverse(x, G, *), G, *)) by Restate + have(in(x, G) |- ((inverse(x, G, *) === inverse(x, G, *)) <=> isInverse(inverse(x, G, *), x, G, *))) by Restate.from( + inverseCharacterization of (y -> inverse(x, G, *)) + ) + thenHave(thesis) by Restate + } - val symmetricInverse = have(in(x, G) |- isInverse(x, inverse(x, G, *), G, *)) by Cut(byDef, symmetry) + val membership = have((group(G, *), in(x, G)) |- in(inverse(x, G, *), G)) by Tautology.from( + inverseIsInverse, isInverse.definition of (y -> inverse(x, G, *)) + ) + + // Use the symmetry of the inverse relation to prove that x is an inverse of inverse(x) + val satisfaction = have((group(G, *), in(x, G)) |- isInverse(x, inverse(x, G, *), G, *)) subproof { + assume(group(G, *)) + + have(∀(z, G, isInverse(inverse(x, G, *), z, G, *) ==> isInverse(z, inverse(x, G, *), G, *))) by Tautology.from(inverseSymmetry of (y -> inverse(x, G, *))) + thenHave(in(x, G) ==> (isInverse(inverse(x, G, *), x, G, *) ==> isInverse(x, inverse(x, G, *), G, *))) by InstantiateForall(x) + val symmetryCut = thenHave((in(x, G), isInverse(inverse(x, G, *), x, G, *)) |- isInverse(x, inverse(x, G, *), G, *)) by Restate + + have(thesis) by Cut(inverseIsInverse, symmetryCut) + } // Conclude by uniqueness val u = inverse(inverse(x, G, *), G, *) - val iff = have(in(inverse(x, G, *), G) |- ((x === u) <=> isInverse(x, inverse(x, G, *), G, *))) by Restate.from(uniqueness of (x -> inverse(x, G, *), y -> x)) - val eq = have((in(x, G), in(inverse(x, G, *), G)) |- (x === u)) by Tautology.from(symmetricInverse, iff) + + val characterization = have(in(inverse(x, G, *), G) |- ((x === u) <=> isInverse(x, inverse(x, G, *), G, *))) by Restate.from(inverseCharacterization of (x -> inverse(x, G, *), y -> x)) + val eq = have((in(x, G), in(inverse(x, G, *), G)) |- (x === u)) by Tautology.from(satisfaction, characterization) have(in(x, G) |- (x === u)) by Cut(membership, eq) thenHave(in(x, G) ==> (x === u)) by Restate From dd4b657f20a101d101d41363a9de1c2a1e65494d Mon Sep 17 00:00:00 2001 From: dhalilov Date: Fri, 24 Mar 2023 19:21:41 +0100 Subject: [PATCH 23/68] Add minor comments tweaks --- .../scala/lisa/mathematics/GroupTheory.scala | 34 +++++++++++++------ 1 file changed, 24 insertions(+), 10 deletions(-) diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala index 1afbc06c..b72c599a 100644 --- a/src/main/scala/lisa/mathematics/GroupTheory.scala +++ b/src/main/scala/lisa/mathematics/GroupTheory.scala @@ -11,13 +11,26 @@ import lisa.mathematics.SetTheory.* * Book : [[https://link.springer.com/book/10.1007/978-1-4684-9234-7]] */ object GroupTheory extends lisa.Main { + // Groups + private val G = variable - private val G, * = variable + // Group laws + private val * = variable + + // Group elements private val x, y, z = variable private val u, v, w = variable + + // Identity elements private val e, f = variable + + // Predicates private val P, Q = predicate(1) + // + // 0. Notation + // + /** * Bounded quantifiers --- These express the usual `∀x ∈ G ϕ` and `∃x ∈ G ϕ`, for some variables (sets) `x` and `G`, which * are shorthands for `∀x (x ∈ G ==> ϕ)` and `∃x (x ∈ G /\ ϕ)`, respectively. @@ -117,15 +130,18 @@ object GroupTheory extends lisa.Main { have(thesis) by Cases(case1, case2) } - show The(u, completeDef)(completeUniquenessTheorem) } } + // + // 1. Basic definitions and results + // + /** * Binary relation --- `*` is a binary relation on `G` if it associates to each pair of elements of `G` - * a unique element in `G`. In other words, `*` is a function `G x G -> G`. + * a unique element in `G`. In other words, `*` is a function `G × G -> G`. */ val binaryRelation = DEF(G, *) --> functionFrom(*, cartesianProduct(G, G), G) @@ -135,18 +151,18 @@ object GroupTheory extends lisa.Main { val op = DEF(x, *, y) --> app(*, pair(x, y)) /** - * Associativity --- `*` is associative (in G) if `(x * y) * z = x * (y * z)` for all `x, y, z` in G. + * Associativity --- `*` is associative (in `G`) if `(x * y) * z = x * (y * z)` for all `x, y, z` in `G`. */ val associativity = DEF(G, *) --> ∀(x, G, ∀(y, G, ∀(z, G, op(op(x, *, y), *, z) === op(x, *, op(y, *, z))))) /** - * Neutral element --- We say that an element `e` in G is neutral if `e * x = x * e = x` for all `x` in G. + * Neutral element --- We say that an element `e` in `G` is neutral if `e * x = x * e = x` for all `x` in `G`. */ val isNeutral = DEF(e, G, *) --> (in(e, G) /\ ∀(x, G, (op(e, *, x) === x) /\ (op(x, *, e) === x))) /** - * Identity existence --- There exists a neutral element `e` in G. + * Identity existence --- There exists a neutral element `e` in `G`. */ val identityExistence = DEF(G, *) --> ∃(e, isNeutral(e, G, *)) @@ -226,7 +242,7 @@ object GroupTheory extends lisa.Main { // Assume y and z are inverses of x. // We prove the following chain of equalities: - // z =1= (yx)z =2= y(xz) =3= y + // z === (yx)z === y(xz) === y // where equalities come from // 1. Left neutrality of yx // 2. Associativity @@ -234,7 +250,6 @@ object GroupTheory extends lisa.Main { val uniqueness = have((group(G, *), in(x, G), isInverse(y, x, G, *), isInverse(z, x, G, *)) |- (y === z)) subproof { val inverseMembership = have(isInverse(y, x, G, *) |- in(y, G)) by Tautology.from(isInverse.definition) - assume(group(G, *)) assume(in(x, G)) assume(isInverse(y, x, G, *)) @@ -298,7 +313,7 @@ object GroupTheory extends lisa.Main { ) } - have(thesis) by ExistenceAndUniqueness.withParameters(isInverse(y, x, G, *), y, z)(existence, uniqueness) + have(thesis) by ExistenceAndUniqueness(isInverse(y, x, G, *))(existence, uniqueness) } /** @@ -323,7 +338,6 @@ object GroupTheory extends lisa.Main { have(thesis) by Tautology.from(definition, isInverse.definition of (y -> x, x -> y)) } - thenHave((group(G, *), in(x, G)) |- isInverse(y, x, G, *) ==> isInverse(x, y, G, *)) by Restate thenHave(group(G, *) |- in(x, G) ==> (isInverse(y, x, G, *) ==> isInverse(x, y, G, *))) by Restate thenHave(thesis) by RightForall } From 05d6a600c34360c7a4b14c8e4e7e6ca882150faf Mon Sep 17 00:00:00 2001 From: dhalilov Date: Fri, 31 Mar 2023 18:57:34 +0200 Subject: [PATCH 24/68] Add not equal sign --- lisa-utils/src/main/scala/lisa/utils/KernelHelpers.scala | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/lisa-utils/src/main/scala/lisa/utils/KernelHelpers.scala b/lisa-utils/src/main/scala/lisa/utils/KernelHelpers.scala index 285cebec..1c59ed2d 100644 --- a/lisa-utils/src/main/scala/lisa/utils/KernelHelpers.scala +++ b/lisa-utils/src/main/scala/lisa/utils/KernelHelpers.scala @@ -103,6 +103,10 @@ object KernelHelpers { extension (t: Term) { infix def ===(u: Term): Formula = PredicateFormula(equality, Seq(t, u)) infix def =(u: Term): Formula = PredicateFormula(equality, Seq(t, u)) + + infix def !==(u: Term): Formula = !(t === u) + + infix def ≠(u: Term): Formula = !(t === u) } /* Conversions */ From 974fa1df893bc0c0b50876cd27d614682bee54d7 Mon Sep 17 00:00:00 2001 From: dhalilov Date: Sun, 16 Apr 2023 14:59:27 +0200 Subject: [PATCH 25/68] =?UTF-8?q?Remove=20bounded=20quantifiers=20and=20us?= =?UTF-8?q?e=20Unicode=20=E2=88=88=20symbol?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- .../scala/lisa/mathematics/GroupTheory.scala | 126 ++++++++---------- 1 file changed, 58 insertions(+), 68 deletions(-) diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala index b72c599a..05ee3166 100644 --- a/src/main/scala/lisa/mathematics/GroupTheory.scala +++ b/src/main/scala/lisa/mathematics/GroupTheory.scala @@ -31,16 +31,6 @@ object GroupTheory extends lisa.Main { // 0. Notation // - /** - * Bounded quantifiers --- These express the usual `∀x ∈ G ϕ` and `∃x ∈ G ϕ`, for some variables (sets) `x` and `G`, which - * are shorthands for `∀x (x ∈ G ==> ϕ)` and `∃x (x ∈ G /\ ϕ)`, respectively. - */ - def ∀(x: VariableLabel, range: VariableLabel, inner: Formula): Formula = forall(x, in(x, range) ==> inner) - - def ∃(x: VariableLabel, range: VariableLabel, inner: Formula): Formula = exists(x, in(x, range) /\ inner) - - def ∃!(x: VariableLabel, range: VariableLabel, inner: Formula): Formula = existsOne(x, in(x, range) /\ inner) - /** * Defines the element that is uniquely given by the uniqueness theorem, or falls back to the error element if the * assumptions of the theorem are not satisfied. @@ -154,12 +144,12 @@ object GroupTheory extends lisa.Main { * Associativity --- `*` is associative (in `G`) if `(x * y) * z = x * (y * z)` for all `x, y, z` in `G`. */ val associativity = DEF(G, *) --> - ∀(x, G, ∀(y, G, ∀(z, G, op(op(x, *, y), *, z) === op(x, *, op(y, *, z))))) + ∀(x, x ∈ G ==> ∀(y, y ∈ G ==> ∀(z, z ∈ G ==> (op(op(x, *, y), *, z) === op(x, *, op(y, *, z)))))) /** * Neutral element --- We say that an element `e` in `G` is neutral if `e * x = x * e = x` for all `x` in `G`. */ - val isNeutral = DEF(e, G, *) --> (in(e, G) /\ ∀(x, G, (op(e, *, x) === x) /\ (op(x, *, e) === x))) + val isNeutral = DEF(e, G, *) --> (e ∈ G /\ ∀(x, (x ∈ G) ==> ((op(e, *, x) === x) /\ (op(x, *, e) === x)))) /** * Identity existence --- There exists a neutral element `e` in `G`. @@ -169,12 +159,12 @@ object GroupTheory extends lisa.Main { /** * Inverse element --- `y` is called an inverse of `x` if `x * y = y * x = e`. */ - val isInverse = DEF(y, x, G, *) --> in(y, G) /\ isNeutral(op(x, *, y), G, *) /\ isNeutral(op(y, *, x), G, *) + val isInverse = DEF(y, x, G, *) --> (y ∈ G) /\ isNeutral(op(x, *, y), G, *) /\ isNeutral(op(y, *, x), G, *) /** * Inverse existence --- For all `x` in G, there exists an element `y` in G such that `x * y = y * x = e`. */ - val inverseExistence = DEF(G, *) --> ∀(x, G, ∃(y, isInverse(y, x, G, *))) + val inverseExistence = DEF(G, *) --> ∀(x, (x ∈ G) ==> ∃(y, isInverse(y, x, G, *))) /** * Group --- A group (G, *) is a set along with a law of composition `*`, satisfying [[associativity]], [[identityExistence]] @@ -196,21 +186,21 @@ object GroupTheory extends lisa.Main { // and the second equality from f's right neutrality val uniqueness = have((isNeutral(e, G, *), isNeutral(f, G, *)) |- (e === f)) subproof { // We prove that neutral elements are elements of G, such that * can be applied. - val membership = have(isNeutral(e, G, *) |- in(e, G)) by Tautology.from(isNeutral.definition) + val membership = have(isNeutral(e, G, *) |- e ∈ G) by Tautology.from(isNeutral.definition) assume(isNeutral(e, G, *)) assume(isNeutral(f, G, *)) // 1. ef = f - have(∀(x, G, (op(e, *, x) === x) /\ (op(x, *, e) === x))) by Tautology.from(isNeutral.definition) - thenHave(in(f, G) ==> ((op(e, *, f) === f) /\ (op(f, *, e) === f))) by InstantiateForall(f) - val cut1 = thenHave(in(f, G) |- ((op(e, *, f) === f) /\ (op(f, *, e) === f))) by Restate + have(∀(x, x ∈ G ==> ((op(e, *, x) === x) /\ (op(x, *, e) === x)))) by Tautology.from(isNeutral.definition) + thenHave(f ∈ G ==> ((op(e, *, f) === f) /\ (op(f, *, e) === f))) by InstantiateForall(f) + val cut1 = thenHave(f ∈ G |- ((op(e, *, f) === f) /\ (op(f, *, e) === f))) by Restate have((op(e, *, f) === f) /\ (op(f, *, e) === f)) by Cut(membership of (e -> f), cut1) val firstEq = thenHave(op(e, *, f) === f) by Tautology // 2. ef = e - val cut2 = have(in(e, G) |- ((op(f, *, e) === e) /\ (op(e, *, f) === e))) by InstFunSchema(Map(e -> f, f -> e))(cut1) + val cut2 = have(e ∈ G |- ((op(f, *, e) === e) /\ (op(e, *, f) === e))) by InstFunSchema(Map(e -> f, f -> e))(cut1) have((op(f, *, e) === e) /\ (op(e, *, f) === e)) by Cut(membership of (e -> e), cut2) val secondEq = thenHave(e === op(e, *, f)) by Tautology @@ -234,11 +224,11 @@ object GroupTheory extends lisa.Main { * Theorem --- The inverse of an element `x` (i.e. `y` such that `x * y = y * x = e`) in `G` is unique. */ val inverseUniqueness = Theorem( - (group(G, *), in(x, G)) |- ∃!(y, isInverse(y, x, G, *)) + (group(G, *), x ∈ G) |- ∃!(y, isInverse(y, x, G, *)) ) { - have(group(G, *) |- ∀(x, G, ∃(y, isInverse(y, x, G, *)))) by Tautology.from(group.definition, inverseExistence.definition) - thenHave(group(G, *) |- (in(x, G) ==> ∃(y, isInverse(y, x, G, *)))) by InstantiateForall(x) - val existence = thenHave((group(G, *), in(x, G)) |- ∃(y, isInverse(y, x, G, *))) by Restate + have(group(G, *) |- ∀(x, x ∈ G ==> ∃(y, isInverse(y, x, G, *)))) by Tautology.from(group.definition, inverseExistence.definition) + thenHave(group(G, *) |- (x ∈ G ==> ∃(y, isInverse(y, x, G, *)))) by InstantiateForall(x) + val existence = thenHave((group(G, *), x ∈ G) |- ∃(y, isInverse(y, x, G, *))) by Restate // Assume y and z are inverses of x. // We prove the following chain of equalities: @@ -247,55 +237,55 @@ object GroupTheory extends lisa.Main { // 1. Left neutrality of yx // 2. Associativity // 3. Right neutrality of xz - val uniqueness = have((group(G, *), in(x, G), isInverse(y, x, G, *), isInverse(z, x, G, *)) |- (y === z)) subproof { - val inverseMembership = have(isInverse(y, x, G, *) |- in(y, G)) by Tautology.from(isInverse.definition) + val uniqueness = have((group(G, *), x ∈ G, isInverse(y, x, G, *), isInverse(z, x, G, *)) |- (y === z)) subproof { + val inverseMembership = have(isInverse(y, x, G, *) |- y ∈ G) by Tautology.from(isInverse.definition) assume(group(G, *)) - assume(in(x, G)) + assume(x ∈ G) assume(isInverse(y, x, G, *)) assume(isInverse(z, x, G, *)) // 1. (yx)z = z - val leftNeutrality = have((group(G, *), in(x, G), isInverse(y, x, G, *), in(z, G)) |- (op(op(y, *, x), *, z) === z)) subproof { + val leftNeutrality = have((group(G, *), x ∈ G, isInverse(y, x, G, *), z ∈ G) |- (op(op(y, *, x), *, z) === z)) subproof { assume(group(G, *)) - assume(in(x, G)) + assume(x ∈ G) assume(isInverse(y, x, G, *)) - assume(in(z, G)) + assume(z ∈ G) - have(∀(u, G, (op(op(y, *, x), *, u) === u) /\ (op(u, *, op(y, *, x)) === u))) by Tautology.from(isInverse.definition, isNeutral.definition of (e -> op(y, *, x))) - thenHave(in(z, G) ==> ((op(op(y, *, x), *, z) === z) /\ (op(z, *, op(y, *, x)) === z))) by InstantiateForall(z) + have(∀(u, u ∈ G ==> ((op(op(y, *, x), *, u) === u) /\ (op(u, *, op(y, *, x)) === u)))) by Tautology.from(isInverse.definition, isNeutral.definition of (e -> op(y, *, x))) + thenHave(z ∈ G ==> ((op(op(y, *, x), *, z) === z) /\ (op(z, *, op(y, *, x)) === z))) by InstantiateForall(z) thenHave(op(op(y, *, x), *, z) === z) by Tautology } val firstEq = have(op(op(y, *, x), *, z) === z) by Cut(inverseMembership of (y -> z), leftNeutrality) // 2. (yx)z = y(xz) - val permuteParentheses = have((group(G, *), in(x, G), in(y, G), in(z, G)) |- (op(op(y, *, x), *, z) === op(y, *, op(x, *, z)))) subproof { + val permuteParentheses = have((group(G, *), x ∈ G, y ∈ G, z ∈ G) |- (op(op(y, *, x), *, z) === op(y, *, op(x, *, z)))) subproof { assume(group(G, *)) - assume(in(x, G)) - assume(in(y, G)) - assume(in(z, G)) - - have(∀(u, G, ∀(v, G, ∀(w, G, op(op(u, *, v), *, w) === op(u, *, op(v, *, w)))))) by Tautology.from(group.definition, associativity.definition) - thenHave(in(y, G) ==> ∀(v, G, ∀(w, G, op(op(y, *, v), *, w) === op(y, *, op(v, *, w))))) by InstantiateForall(y) - thenHave(∀(v, G, ∀(w, G, op(op(y, *, v), *, w) === op(y, *, op(v, *, w))))) by Tautology - thenHave(in(x, G) ==> ∀(w, G, op(op(y, *, x), *, w) === op(y, *, op(x, *, w)))) by InstantiateForall(x) - thenHave(∀(w, G, op(op(y, *, x), *, w) === op(y, *, op(x, *, w)))) by Tautology - thenHave(in(z, G) ==> (op(op(y, *, x), *, z) === op(y, *, op(x, *, z)))) by InstantiateForall(z) + assume(x ∈ G) + assume(y ∈ G) + assume(z ∈ G) + + have(∀(u, u ∈ G ==> ∀(v, v ∈ G ==> ∀(w, w ∈ G ==> (op(op(u, *, v), *, w) === op(u, *, op(v, *, w))))))) by Tautology.from(group.definition, associativity.definition) + thenHave(y ∈ G ==> ∀(v, v ∈ G ==> ∀(w, w ∈ G ==> (op(op(y, *, v), *, w) === op(y, *, op(v, *, w)))))) by InstantiateForall(y) + thenHave(∀(v, v ∈ G ==> ∀(w, w ∈ G ==> (op(op(y, *, v), *, w) === op(y, *, op(v, *, w)))))) by Tautology + thenHave(x ∈ G ==> ∀(w, w ∈ G ==> (op(op(y, *, x), *, w) === op(y, *, op(x, *, w))))) by InstantiateForall(x) + thenHave(∀(w, w ∈ G ==> (op(op(y, *, x), *, w) === op(y, *, op(x, *, w))))) by Tautology + thenHave(z ∈ G ==> (op(op(y, *, x), *, z) === op(y, *, op(x, *, z)))) by InstantiateForall(z) thenHave(op(op(y, *, x), *, z) === op(y, *, op(x, *, z))) by Tautology } - val associativityCut = have((group(G, *), in(x, G) /\ in(y, G) /\ in(z, G)) |- (op(op(y, *, x), *, z) === op(y, *, op(x, *, z)))) by Restate.from(permuteParentheses) - val memberships = have(in(x, G) /\ in(y, G) /\ in(z, G)) by Tautology.from(inverseMembership of (y -> y), inverseMembership of (y -> z)) + val associativityCut = have((group(G, *), x ∈ G /\ y ∈ G /\ z ∈ G) |- (op(op(y, *, x), *, z) === op(y, *, op(x, *, z)))) by Restate.from(permuteParentheses) + val memberships = have(x ∈ G /\ y ∈ G /\ z ∈ G) by Tautology.from(inverseMembership of (y -> y), inverseMembership of (y -> z)) val secondEq = have(op(op(y, *, x), *, z) === op(y, *, op(x, *, z))) by Cut(memberships, associativityCut) // 3. y(xz) = y - val rightNeutrality = have((group(G, *), in(x, G), in(y, G), isInverse(z, x, G, *)) |- (op(y, *, op(x, *, z)) === y)) subproof { + val rightNeutrality = have((group(G, *), x ∈ G, y ∈ G, isInverse(z, x, G, *)) |- (op(y, *, op(x, *, z)) === y)) subproof { assume(group(G, *)) - assume(in(x, G)) - assume(in(y, G)) + assume(x ∈ G) + assume(y ∈ G) assume(isInverse(z, x, G, *)) - have(∀(u, G, (op(op(x, *, z), *, u) === u) /\ (op(u, *, op(x, *, z)) === u))) by Tautology.from(isInverse.definition of (y -> z), isNeutral.definition of (e -> op(x, *, z))) - thenHave(in(y, G) ==> ((op(op(x, *, z), *, y) === y) /\ (op(y, *, op(x, *, z)) === y))) by InstantiateForall(y) + have(∀(u, u ∈ G ==> ((op(op(x, *, z), *, u) === u) /\ (op(u, *, op(x, *, z)) === u)))) by Tautology.from(isInverse.definition of (y -> z), isNeutral.definition of (e -> op(x, *, z))) + thenHave(y ∈ G ==> ((op(op(x, *, z), *, y) === y) /\ (op(y, *, op(x, *, z)) === y))) by InstantiateForall(y) thenHave(op(y, *, op(x, *, z)) === y) by Tautology } val thirdEq = have(op(y, *, op(x, *, z)) === y) by Cut(inverseMembership of (y -> y), rightNeutrality) @@ -325,20 +315,20 @@ object GroupTheory extends lisa.Main { * Theorem --- `y` is the inverse of `x` iff `x` is the inverse of `y` */ val inverseSymmetry = Theorem( - group(G, *) |- ∀(x, G, isInverse(y, x, G, *) ==> isInverse(x, y, G, *)) + group(G, *) |- ∀(x, x ∈ G ==> (isInverse(y, x, G, *) ==> isInverse(x, y, G, *))) ) { - have((group(G, *), in(x, G), isInverse(y, x, G, *)) |- isInverse(x, y, G, *)) subproof { + have((group(G, *), x ∈ G, isInverse(y, x, G, *)) |- isInverse(x, y, G, *)) subproof { assume(group(G, *)) - assume(in(x, G)) + assume(x ∈ G) assume(isInverse(y, x, G, *)) - have(in(y, G) /\ isNeutral(op(x, *, y), G, *) /\ isNeutral(op(y, *, x), G, *)) by Tautology.from(isInverse.definition) + have(y ∈ G /\ isNeutral(op(x, *, y), G, *) /\ isNeutral(op(y, *, x), G, *)) by Tautology.from(isInverse.definition) thenHave(isNeutral(op(y, *, x), G, *) /\ isNeutral(op(x, *, y), G, *)) by Tautology - val definition = thenHave(in(x, G) /\ isNeutral(op(y, *, x), G, *) /\ isNeutral(op(x, *, y), G, *)) by Tautology + val definition = thenHave(x ∈ G /\ isNeutral(op(y, *, x), G, *) /\ isNeutral(op(x, *, y), G, *)) by Tautology have(thesis) by Tautology.from(definition, isInverse.definition of (y -> x, x -> y)) } - thenHave(group(G, *) |- in(x, G) ==> (isInverse(y, x, G, *) ==> isInverse(x, y, G, *))) by Restate + thenHave(group(G, *) |- x ∈ G ==> (isInverse(y, x, G, *) ==> isInverse(x, y, G, *))) by Restate thenHave(thesis) by RightForall } @@ -346,38 +336,38 @@ object GroupTheory extends lisa.Main { * Involution of the inverse -- For all `x`, `inverse(inverse(x)) = x`. */ val inverseIsInvolutive = Theorem( - group(G, *) |- ∀(x, G, inverse(inverse(x, G, *), G, *) === x) + group(G, *) |- ∀(x, x ∈ G ==> (inverse(inverse(x, G, *), G, *) === x)) ) { assume(group(G, *)) - val inverseCharacterization = have((group(G, *), in(x, G)) |- ((y === inverse(x, G, *)) <=> isInverse(y, x, G, *))) subproof { + val inverseCharacterization = have((group(G, *), x ∈ G) |- ((y === inverse(x, G, *)) <=> isInverse(y, x, G, *))) subproof { assume(group(G, *)) - have(in(x, G) |- ∀(y, (y === inverse(x, G, *)) <=> isInverse(y, x, G, *))) by Tautology.from(inverseUniqueness, inverse.definition) + have(x ∈ G |- ∀(y, (y === inverse(x, G, *)) <=> isInverse(y, x, G, *))) by Tautology.from(inverseUniqueness, inverse.definition) thenHave(thesis) by InstantiateForall(y) } // Prove that inverse(x) is an inverse of x - val inverseIsInverse = have((group(G, *), in(x, G)) |- isInverse(inverse(x, G, *), x, G, *)) subproof { + val inverseIsInverse = have((group(G, *), x ∈ G) |- isInverse(inverse(x, G, *), x, G, *)) subproof { assume(group(G, *)) - have(in(x, G) |- ((inverse(x, G, *) === inverse(x, G, *)) <=> isInverse(inverse(x, G, *), x, G, *))) by Restate.from( + have(x ∈ G |- ((inverse(x, G, *) === inverse(x, G, *)) <=> isInverse(inverse(x, G, *), x, G, *))) by Restate.from( inverseCharacterization of (y -> inverse(x, G, *)) ) thenHave(thesis) by Restate } - val membership = have((group(G, *), in(x, G)) |- in(inverse(x, G, *), G)) by Tautology.from( + val membership = have((group(G, *), x ∈ G) |- in(inverse(x, G, *), G)) by Tautology.from( inverseIsInverse, isInverse.definition of (y -> inverse(x, G, *)) ) // Use the symmetry of the inverse relation to prove that x is an inverse of inverse(x) - val satisfaction = have((group(G, *), in(x, G)) |- isInverse(x, inverse(x, G, *), G, *)) subproof { + val satisfaction = have((group(G, *), x ∈ G) |- isInverse(x, inverse(x, G, *), G, *)) subproof { assume(group(G, *)) - have(∀(z, G, isInverse(inverse(x, G, *), z, G, *) ==> isInverse(z, inverse(x, G, *), G, *))) by Tautology.from(inverseSymmetry of (y -> inverse(x, G, *))) - thenHave(in(x, G) ==> (isInverse(inverse(x, G, *), x, G, *) ==> isInverse(x, inverse(x, G, *), G, *))) by InstantiateForall(x) - val symmetryCut = thenHave((in(x, G), isInverse(inverse(x, G, *), x, G, *)) |- isInverse(x, inverse(x, G, *), G, *)) by Restate + have(∀(z, z ∈ G ==> (isInverse(inverse(x, G, *), z, G, *) ==> isInverse(z, inverse(x, G, *), G, *)))) by Tautology.from(inverseSymmetry of (y -> inverse(x, G, *))) + thenHave(x ∈ G ==> (isInverse(inverse(x, G, *), x, G, *) ==> isInverse(x, inverse(x, G, *), G, *))) by InstantiateForall(x) + val symmetryCut = thenHave((x ∈ G, isInverse(inverse(x, G, *), x, G, *)) |- isInverse(x, inverse(x, G, *), G, *)) by Restate have(thesis) by Cut(inverseIsInverse, symmetryCut) } @@ -386,10 +376,10 @@ object GroupTheory extends lisa.Main { val u = inverse(inverse(x, G, *), G, *) val characterization = have(in(inverse(x, G, *), G) |- ((x === u) <=> isInverse(x, inverse(x, G, *), G, *))) by Restate.from(inverseCharacterization of (x -> inverse(x, G, *), y -> x)) - val eq = have((in(x, G), in(inverse(x, G, *), G)) |- (x === u)) by Tautology.from(satisfaction, characterization) + val eq = have((x ∈ G, in(inverse(x, G, *), G)) |- (x === u)) by Tautology.from(satisfaction, characterization) - have(in(x, G) |- (x === u)) by Cut(membership, eq) - thenHave(in(x, G) ==> (x === u)) by Restate + have(x ∈ G |- (x === u)) by Cut(membership, eq) + thenHave(x ∈ G ==> (x === u)) by Restate thenHave(thesis) by RightForall } } From 7bc75f410015e1cae1de5cfdb30e575d87a651b6 Mon Sep 17 00:00:00 2001 From: dhalilov Date: Sat, 22 Apr 2023 18:41:22 +0200 Subject: [PATCH 26/68] Create Definition tactic --- .../grouptheory/GroupTheoryTactics.scala | 148 +++++++++++++++++- 1 file changed, 141 insertions(+), 7 deletions(-) diff --git a/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala b/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala index 4b98a0a1..07e9f31b 100644 --- a/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala +++ b/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala @@ -1,19 +1,16 @@ package lisa.automation.grouptheory import lisa.automation.kernel.OLPropositionalSolver.Tautology -import lisa.automation.kernel.SimplePropositionalSolver.* -import lisa.automation.kernel.SimpleSimplifier.* -import lisa.mathematics.SetTheory +import lisa.automation.kernel.SimpleSimplifier.Substitution +import lisa.mathematics.SetTheory.* +import lisa.settheory.SetTheoryLibrary import lisa.prooflib.* import lisa.prooflib.BasicStepTactic.* import lisa.prooflib.ProofTacticLib.{*, given} -import lisa.settheory.SetTheoryLibrary -import lisa.utils.KernelHelpers.* object GroupTheoryTactics { import lisa.settheory.SetTheoryLibrary.{*, given} - /** *
    * Γ |- ∃x. φ, Δ   Γ', φ, φ[y/x] |- x = y, Δ'
@@ -44,7 +41,7 @@ object GroupTheoryTactics {
       } else if (!contains(uniquenessSeq.left, substPhi)) {
         proof.InvalidProofTactic(s"Uniqueness sequent premises do not contain φ[$y/$x].")
       } else if (!contains(uniquenessSeq.right, x === y) && !contains(uniquenessSeq.right, y === x)) {
-        proof.InvalidProofTactic(s"Uniqueness sequent premises do not contain $x = $y")
+        proof.InvalidProofTactic(s"Uniqueness sequent conclusion does not contain $x = $y")
       } else if (!contains(bot.right, uniqueExistenceFormula)) {
         proof.InvalidProofTactic(s"Bottom sequent conclusion does not contain ∃!$x. φ")
       }
@@ -118,6 +115,143 @@ object GroupTheoryTactics {
     }
   }
 
+  /**
+   * 
+   *
+   * -------------  if f(xs) = The(y, P(y)) is a function definition
+   * |- P(f(xs))
+   * 
+ * Here `xs` is an arbitrary list of parameters. + * + * If `f(xs) = The(y, (φ ==> Q(y)) /\ (!φ ==> (y === t)))` is a conditional function definition, then: + *
+   *
+   * --------------
+   * φ |- Q(f(xs))
+   * 
+ */ + object Definition extends ProofTactic { + def apply(using proof: SetTheoryLibrary.Proof, om: OutputManager) + (f: ConstantFunctionLabel, uniqueness: proof.Fact)(xs: Term*) + (bot: Sequent): proof.ProofTacticJudgement = { + f.definition match { + case SetTheoryLibrary.theory.FunctionDefinition(_, _, expr) => + // Check if the definition is conditional + val method = expr(xs) match { + case ConnectorFormula(And, Seq( + ConnectorFormula(Implies, Seq(a, _)), + ConnectorFormula(Implies, Seq(b, _)) + )) if isSame(neg(a), b) => conditional + + case _ => unconditional + } + + method(f, uniqueness)(xs)(bot) + + case SetTheoryLibrary.theory.Axiom(_, _) => proof.InvalidProofTactic("Axiomatic definitions are not supported.") + case _ => proof.InvalidProofTactic("Could not get definition of function.") + } + } + + /** + *
+     *
+     * -------------  if f(xs) = The(y, P(y)) is a function definition
+     * |- P(f(xs))
+     * 
+ */ + def unconditional(using proof: SetTheoryLibrary.Proof, om: OutputManager) + (f: ConstantFunctionLabel, uniqueness: proof.Fact)(xs: Term*) + (bot: Sequent): proof.ProofTacticJudgement = { + f.definition match { + case definition@SetTheoryLibrary.theory.FunctionDefinition(_, y, expr) => + if (bot.right.size != 1) { + return proof.InvalidProofTactic("Right-hand side of bottom sequent should contain only 1 formula.") + } + + // Extract variable labels to instantiate them later in the proof + val LambdaTermFormula(vars, _) = expr + val instantiations: Seq[(SchematicTermLabel, LambdaTermTerm)] = vars.zip(xs.map(x => LambdaTermTerm(Seq(), x))) + + // Instantiate terms in the definition + val P = LambdaTermFormula(Seq(y), expr(xs)) + + val expected = P(Seq(f(xs))) + if (!isSame(expected, bot.right.head)) { + return proof.InvalidProofTactic("Right-hand side of bottom sequent should be of the form P(f(xs)).") + } + + TacticSubproof { + have(∀(y, (y === f(xs)) <=> P(Seq(y)))) by Tautology.from(uniqueness, definition.of(instantiations: _*)) + thenHave((y === f(xs)) <=> P(Seq(y))) by InstantiateForall(y) + thenHave((f(xs) === f(xs)) <=> P(Seq(f(xs)))) by InstFunSchema(Map(y -> f(xs))) + thenHave(P(Seq(f(xs)))) by Restate + } + + case SetTheoryLibrary.theory.Axiom(_, _) => proof.InvalidProofTactic("Axiomatic definitions are not supported.") + case _ => proof.InvalidProofTactic("Could not get definition of function.") + } + } + + /** + *
+     *
+     * -------------- if f(xs) = The(y, (φ ==> Q(y)) /\ (!φ ==> R(y)))
+     * φ |- Q(f(xs))
+     * 
+ */ + def conditional(using proof: SetTheoryLibrary.Proof, om: OutputManager) + (f: ConstantFunctionLabel, uniqueness: proof.Fact)(xs: Term*) + (bot: Sequent): proof.ProofTacticJudgement = { + f.definition match { + case definition@SetTheoryLibrary.theory.FunctionDefinition(_, y, expr) => + if (bot.right.size != 1) { + return proof.InvalidProofTactic("Right-hand side of bottom sequent should contain exactly 1 formula.") + } else if (bot.left.isEmpty) { + return proof.InvalidProofTactic("Left-hand side of bottom sequent should not be empty.") + } + + // Extract variable labels to instantiate them later in the proof + val LambdaTermFormula(vars, _) = expr + val instantiations: Seq[(SchematicTermLabel, LambdaTermTerm)] = vars.zip(xs.map(x => LambdaTermTerm(Seq(), x))) + + // Instantiate terms in the definition + val P = LambdaTermFormula(Seq(y), expr(xs)) + + // Unfold the conditional definition to find Q + val phi = ConnectorFormula(And, bot.left.toSeq) + val Q = P.body match { + case ConnectorFormula(And, Seq( + ConnectorFormula(Implies, Seq(a, f)), + ConnectorFormula(Implies, Seq(b, g)) + )) if isSame(neg(a), b) => + if (isSame(a, phi)) LambdaTermFormula(Seq(y), f) + else if (isSame(b, phi)) LambdaTermFormula(Seq(y), g) + else return proof.InvalidProofTactic("Condition of definition is not satisfied.") + + case _ => return proof.InvalidProofTactic("Definition is not conditional.") + } + + val expected = Q(Seq(f(xs))) + if (!isSame(expected, bot.right.head)) { + return proof.InvalidProofTactic("Right-hand side of bottom sequent should be of the form Q(f(xs)).") + } + + TacticSubproof { + have(∀(y, (y === f(xs)) <=> P(Seq(y)))) by Tautology.from(uniqueness, definition.of(instantiations: _*)) + thenHave((y === f(xs)) <=> P(Seq(y))) by InstantiateForall(y) + thenHave((f(xs) === f(xs)) <=> P(Seq(f(xs)))) by InstFunSchema(Map(y -> f(xs))) + thenHave(P(Seq(f(xs)))) by Restate + thenHave(phi ==> Q(Seq(f(xs)))) by Tautology + thenHave(phi |- Q(Seq(f(xs)))) by Restate + } + + case SetTheoryLibrary.theory.Axiom(_, _) => proof.InvalidProofTactic("Axiomatic definitions are not supported.") + case _ => proof.InvalidProofTactic("Could not get definition of function.") + } + } + } + /** *
    * Γ, φ |- Δ, Σ   Γ, ¬φ |- Δ, Σ'  

From 647389b1cfadced32499a2545cd3c290c8aa6f54 Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Mon, 24 Apr 2023 16:14:27 +0200
Subject: [PATCH 27/68] Fix merged FunctionDefinition unapply  in Definition
 tactic

---
 .../lisa/automation/grouptheory/GroupTheoryTactics.scala    | 6 +++---
 1 file changed, 3 insertions(+), 3 deletions(-)

diff --git a/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala b/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala
index 07e9f31b..4f8d06e3 100644
--- a/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala
+++ b/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala
@@ -135,7 +135,7 @@ object GroupTheoryTactics {
              (f: ConstantFunctionLabel, uniqueness: proof.Fact)(xs: Term*)
              (bot: Sequent): proof.ProofTacticJudgement = {
       f.definition match {
-        case SetTheoryLibrary.theory.FunctionDefinition(_, _, expr) =>
+        case SetTheoryLibrary.theory.FunctionDefinition(_, _, expr, _) =>
           // Check if the definition is conditional
           val method = expr(xs) match {
             case ConnectorFormula(And, Seq(
@@ -164,7 +164,7 @@ object GroupTheoryTactics {
                      (f: ConstantFunctionLabel, uniqueness: proof.Fact)(xs: Term*)
                      (bot: Sequent): proof.ProofTacticJudgement = {
       f.definition match {
-        case definition@SetTheoryLibrary.theory.FunctionDefinition(_, y, expr) =>
+        case definition@SetTheoryLibrary.theory.FunctionDefinition(_, y, expr, _) =>
           if (bot.right.size != 1) {
             return proof.InvalidProofTactic("Right-hand side of bottom sequent should contain only 1 formula.")
           }
@@ -204,7 +204,7 @@ object GroupTheoryTactics {
                    (f: ConstantFunctionLabel, uniqueness: proof.Fact)(xs: Term*)
                    (bot: Sequent): proof.ProofTacticJudgement = {
       f.definition match {
-        case definition@SetTheoryLibrary.theory.FunctionDefinition(_, y, expr) =>
+        case definition@SetTheoryLibrary.theory.FunctionDefinition(_, y, expr, _) =>
           if (bot.right.size != 1) {
             return proof.InvalidProofTactic("Right-hand side of bottom sequent should contain exactly 1 formula.")
           } else if (bot.left.isEmpty) {

From ac30619be0b1edd970d46b7de585dec1c7f40276 Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Mon, 24 Apr 2023 21:56:32 +0200
Subject: [PATCH 28/68] Prove restricted function domain

---
 .../lisa/mathematics/FirstOrderLogic.scala    |  77 ++++++++++
 .../scala/lisa/mathematics/SetTheory.scala    | 133 +++++++++++++++++-
 2 files changed, 208 insertions(+), 2 deletions(-)

diff --git a/src/main/scala/lisa/mathematics/FirstOrderLogic.scala b/src/main/scala/lisa/mathematics/FirstOrderLogic.scala
index 58d7d287..07c7fff0 100644
--- a/src/main/scala/lisa/mathematics/FirstOrderLogic.scala
+++ b/src/main/scala/lisa/mathematics/FirstOrderLogic.scala
@@ -153,6 +153,83 @@ object FirstOrderLogic extends lisa.Main {
     have(thesis) by RightIff(fwd, bwd)
   }
 
+  /**
+   * Theorem -- Existential quantification distributes conjunction.
+   */
+  val existentialConjunctionDistribution = Theorem(
+    exists(x, P(x) /\ Q(x)) |- exists(x, P(x)) /\ exists(x, Q(x))
+  ) {
+    have(P(x) |- P(x)) by Hypothesis
+    thenHave(P(x) /\ Q(x) |- P(x)) by Weakening
+    thenHave(P(x) /\ Q(x) |- exists(x, P(x))) by RightExists
+    val left = thenHave(exists(x, P(x) /\ Q(x)) |- exists(x, P(x))) by LeftExists
+
+    have(Q(x) |- Q(x)) by Hypothesis
+    thenHave(P(x) /\ Q(x) |- Q(x)) by Weakening
+    thenHave(P(x) /\ Q(x) |- exists(x, Q(x))) by RightExists
+    val right = thenHave(exists(x, P(x) /\ Q(x)) |- exists(x, Q(x))) by LeftExists
+
+    have(thesis) by RightAnd(left, right)
+  }
+
+  /**
+   * Theorem -- Existential quantification fully distributes when the conjunction involves one closed formula.
+   */
+  val existentialConjunctionWithClosedFormula = Theorem(
+    exists(x, P(x) /\ p()) <=> (exists(x, P(x)) /\ p())
+  ) {
+    val forward = have(exists(x, P(x) /\ p()) ==> (exists(x, P(x)) /\ p())) subproof {
+      have(exists(x, P(x) /\ p()) |- exists(x, P(x)) /\ exists(x, p())) by Restate.from(existentialConjunctionDistribution of (
+        Q -> lambda(x, p())
+      ))
+      val substitution = thenHave(
+        (exists(x, P(x) /\ p()), (exists(x, P(x)) /\ exists(x, p())) <=> (exists(x, P(x)) /\ p())) |- exists(x, P(x)) /\ p()
+      ) by RightSubstIff(List((exists(x, P(x)) /\ exists(x, p()), exists(x, P(x)) /\ p())), lambda(p, p()))
+
+      have(exists(x, p()) <=> p()) by Restate.from(closedFormulaExistential)
+      thenHave((exists(x, P(x)) /\ exists(x, p())) <=> (exists(x, P(x)) /\ p())) by Tautology
+
+      have(exists(x, P(x) /\ p()) |- exists(x, P(x)) /\ p()) by Cut(lastStep, substitution)
+      thenHave(thesis) by Restate
+    }
+
+    val backward = have((exists(x, P(x)) /\ p()) ==> exists(x, P(x) /\ p())) subproof {
+      have(P(x) /\ p() |- P(x) /\ p()) by Hypothesis
+      thenHave((P(x), p()) |- P(x) /\ p()) by Restate
+      thenHave((P(x), p()) |- exists(x, P(x) /\ p())) by RightExists
+      thenHave((exists(x, P(x)), p()) |- exists(x, P(x) /\ p())) by LeftExists
+      thenHave(thesis) by Restate
+    }
+
+    have(thesis) by RightIff(forward, backward)
+  }
+
+  /**
+   * Theorem -- If there is an equality on the existential quantifier's bound variable inside its body, then we can reduce
+   * the existential quantifier to the satisfaction of the remaining body.
+   */
+  val equalityInExistentialQuantifier = Theorem(
+    exists(x, P(x) /\ (y === x)) <=> P(y)
+  ) {
+    have(exists(x, P(x) /\ (y === x)) |- P(y)) subproof {
+      have(P(x) |- P(x)) by Hypothesis
+      thenHave((P(x), y === x) |- P(y)) by RightSubstEq(List((y, x)), lambda(y, P(y)))
+      thenHave(P(x) /\ (y === x) |- P(y)) by Restate
+      thenHave(thesis) by LeftExists
+    }
+    val forward = thenHave(exists(x, P(x) /\ (y === x)) ==> P(y)) by Restate
+
+    have(P(y) |- exists(x, P(x) /\ (y === x))) subproof {
+      have(P(x) /\ (y === x) |- P(x) /\ (y === x)) by Hypothesis
+      thenHave(P(x) /\ (y === x) |- exists(x, P(x) /\ (y === x))) by RightExists
+      thenHave(P(y) /\ (y === y) |- exists(x, P(x) /\ (y === x))) by InstFunSchema(Map(x -> y))
+      thenHave(thesis) by Restate
+    }
+    val backward = thenHave(P(y) ==> exists(x, P(x) /\ (y === x))) by Restate
+
+    have(thesis) by RightIff(forward, backward)
+  }
+
   /**
    * Theorem --- Disjunction and existential quantification commute.
    */
diff --git a/src/main/scala/lisa/mathematics/SetTheory.scala b/src/main/scala/lisa/mathematics/SetTheory.scala
index fc33f2dd..19e5e319 100644
--- a/src/main/scala/lisa/mathematics/SetTheory.scala
+++ b/src/main/scala/lisa/mathematics/SetTheory.scala
@@ -1,5 +1,6 @@
 package lisa.mathematics
 
+import lisa.automation.grouptheory.GroupTheoryTactics.Definition
 import lisa.automation.kernel.OLPropositionalSolver.Tautology
 import lisa.automation.kernel.SimpleSimplifier.*
 import lisa.automation.settheory.SetTheoryTactics.*
@@ -367,7 +368,7 @@ object SetTheory extends lisa.Main {
   /**
    * Theorem --- The empty set is a subset of every set.
    *
-   *    `(∀ x.) x ⊆ ∅`
+   *    `(∀ x.) ∅ ⊆ x`
    */
   val emptySetIsASubset = Theorem(
     subset(∅, x)
@@ -855,6 +856,10 @@ object SetTheory extends lisa.Main {
    */
   val setIntersection = DEF(x, y) --> The(z, ∀(t, in(t, z) <=> (in(t, x) /\ in(t, y))))(setIntersectionUniqueness)
 
+  extension(x: Term) {
+    infix def ∩(y: Term) = setIntersection(x, y)
+  }
+
   val unaryIntersectionUniqueness = Theorem(
     ∃!(z, ∀(t, in(t, z) <=> (exists(b, in(b, x)) /\ ∀(b, in(b, x) ==> in(t, b)))))
   ) {
@@ -2207,7 +2212,131 @@ object SetTheory extends lisa.Main {
    */
   val restrictedFunction = DEF(f, x) --> The(g, ∀(t, in(t, g) <=> (in(t, f) /\ ∃(y, ∃(z, in(y, x) /\ (t === pair(y, z)))))))(restrictedFunctionUniqueness)
 
-  // TODO: functional restricted over x has its domain as x ∈tersect dom f
+  /**
+   * Pair membership in a restricted function -- A pair `(t, a)` is in `f_x` iff `(t, a) ∈ f` and `t ∈ x`.
+   *
+   * This is a direct but painful corollary of the definition.
+   */
+  val restrictedFunctionPairMembership = Lemma(
+    in(pair(t, a), restrictedFunction(f, x)) <=> (in(pair(t, a), f) /\ in(t, x))
+  ) {
+    val g = restrictedFunction(f, x)
+
+    have(∀(t, in(t, g) <=> (in(t, f) /\ ∃(y, ∃(z, in(y, x) /\ (t === pair(y, z))))))) by Definition(
+      restrictedFunction, restrictedFunctionUniqueness
+    )(f, x)
+    val pairMembership = thenHave(
+      in(pair(t, a), g) <=> (in(pair(t, a), f) /\ ∃(y, ∃(z, in(y, x) /\ (pair(t, a) === pair(y, z)))))
+    ) by InstantiateForall(pair(t, a))
+
+    have((pair(t, a) === pair(y, z)) <=> ((t === y) /\ (a === z))) by Restate.from(pairExtensionality of (a -> t, b -> a, c -> y, d -> z))
+    thenHave((in(y, x) /\ (pair(t, a) === pair(y, z))) <=> (in(y, x) /\ (t === y) /\ (a === z))) by Tautology
+    thenHave(∀(z, (in(y, x) /\ (pair(t, a) === pair(y, z))) <=> (in(y, x) /\ (t === y) /\ (a === z)))) by RightForall
+
+    val existentialEquiv1 = have(∃(z, in(y, x) /\ (pair(t, a) === pair(y, z))) <=> ∃(z, in(y, x) /\ (t === y) /\ (a === z))) by Cut(
+      lastStep, existentialEquivalenceDistribution of (
+        P -> lambda(z, in(y, x) /\ (pair(t, a) === pair(y, z))),
+        Q -> lambda(z, in(y, x) /\ (t === y) /\ (a === z))
+      )
+    )
+    
+    have(∃(z, in(y, x) /\ (t === y) /\ (a === z)) <=> (in(y, x) /\ (t === y))) by Restate.from(
+      equalityInExistentialQuantifier of (
+        P -> lambda(z, in(y, x) /\ (t === y)),
+        y -> a 
+      )
+    )
+    
+    have(∃(z, in(y, x) /\ (pair(t, a) === pair(y, z))) <=> (in(y, x) /\ (t === y))) by Tautology.from(existentialEquiv1, lastStep)
+    thenHave(∀(y, ∃(z, in(y, x) /\ (pair(t, a) === pair(y, z))) <=> (in(y, x) /\ (t === y)))) by RightForall
+    
+    val existentialEquiv2 = have(∃(y, ∃(z, in(y, x) /\ (pair(t, a) === pair(y, z)))) <=> ∃(y, in(y, x) /\ (t === y))) by Cut(
+      lastStep, existentialEquivalenceDistribution of(
+        P -> lambda(y, ∃(z, in(y, x) /\ (pair(t, a) === pair(y, z)))),
+        Q -> lambda(y, in(y, x) /\ (t === y))
+      )
+    )
+    
+    have(∃(y, in(y, x) /\ (t === y)) <=> in(t, x)) by Restate.from(
+      equalityInExistentialQuantifier of(
+        P -> lambda(y, in(y, x)),
+        y -> t
+      )
+    )
+    
+    have(∃(y, ∃(z, in(y, x) /\ (pair(t, a) === pair(y, z)))) <=> in(t, x)) by Tautology.from(existentialEquiv2, lastStep)
+    thenHave((in(pair(t, a), f) /\ ∃(y, ∃(z, in(y, x) /\ (pair(t, a) === pair(y, z))))) <=> (in(pair(t, a), f) /\ in(t, x))) by Tautology
+    
+    have(thesis) by Tautology.from(lastStep, pairMembership)
+  }
+
+  /**
+   * Restricted function domain -- For a function `f`, the domain of `f_x` is `x ∩ relationDomain(f)`.
+   */
+  val restrictedFunctionDomain = Theorem(
+    relationDomain(restrictedFunction(f, x)) === (x ∩ relationDomain(f))
+  ) {
+    val D = variable
+    val dom = x ∩ relationDomain(f)
+    val g = restrictedFunction(f, x)
+
+    // Characterize x ∩ relationDomain(f)
+    val domCharacterization = have(∀(t, in(t, dom) <=> (∃(a, in(pair(t, a), f)) /\ in(t, x)))) subproof {
+      // Use the definition of the intersection
+      have(∀(t, in(t, dom) <=> (in(t, x) /\ in(t, relationDomain(f))))) by Definition(
+        setIntersection, setIntersectionUniqueness
+      )(x, relationDomain(f))
+      val intersectionDef = thenHave(in(t, dom) <=> (in(t, x) /\ in(t, relationDomain(f)))) by InstantiateForall(t)
+
+      // Use the definition of the relation domain
+      have(∀(t, in(t, relationDomain(f)) <=> ∃(a, in(pair(t, a), f)))) by Definition(
+        relationDomain, relationDomainUniqueness
+      )(f)
+      thenHave(in(t, relationDomain(f)) <=> ∃(a, in(pair(t, a), f))) by InstantiateForall(t)
+
+      // Conclude
+      have(in(t, dom) <=> (∃(a, in(pair(t, a), f)) /\ in(t, x))) by Tautology.from(intersectionDef, lastStep)
+      thenHave(thesis) by RightForall
+    }
+
+    // Characterize the domain of g
+    have(∀(D, (relationDomain(g) === D) <=> ∀(t, in(t, D) <=> ∃(a, in(pair(t, a), g))))) by Tautology.from(
+      relationDomain.definition of (r -> g), relationDomainUniqueness
+    )
+    val characterization = thenHave((relationDomain(g) === dom) <=> ∀(t, in(t, dom) <=> ∃(a, in(pair(t, a), g)))) by InstantiateForall(dom)
+
+    // Use the membership of a pair in the restricted function to derive a simpler characterization
+    have(∀(a, in(pair(t, a), g) <=> (in(pair(t, a), f) /\ in(t, x)))) by RightForall(restrictedFunctionPairMembership)
+    have(∃(a, in(pair(t, a), g)) <=> ∃(a, in(pair(t, a), f) /\ in(t, x))) by Tautology.from(
+      lastStep, existentialEquivalenceDistribution of (
+        P -> lambda(a, in(pair(t, a), g)),
+        Q -> lambda(a, in(pair(t, a), f) /\ in(t, x))
+      )
+    )
+
+    // Extract in(t, x) from the existential quantifier
+    val p = formulaVariable // local shadowing to correctly use the theorem
+    have(∃(a, in(pair(t, a), g)) <=> ∃(a, in(pair(t, a), f)) /\ in(t, x)) by Tautology.from(
+      lastStep, existentialConjunctionWithClosedFormula of (
+        P -> lambda(a, in(pair(t, a), f)),
+        p -> lambda(Seq(), in(t, x))
+      )
+    )
+
+    thenHave((in(t, dom) <=> ∃(a, in(pair(t, a), g))) <=> (in(t, dom) <=> ∃(a, in(pair(t, a), f)) /\ in(t, x))) by Tautology
+    thenHave(∀(t, (in(t, dom) <=> ∃(a, in(pair(t, a), g))) <=> (in(t, dom) <=> ∃(a, in(pair(t, a), f)) /\ in(t, x)))) by RightForall
+
+    have(∀(t, in(t, dom) <=> ∃(a, in(pair(t, a), g))) <=> ∀(t, in(t, dom) <=> ∃(a, in(pair(t, a), f)) /\ in(t, x))) by Cut(
+      lastStep, universalEquivalenceDistribution of (
+        P -> lambda(t, in(t, dom) <=> ∃(a, in(pair(t, a), g))),
+        Q -> lambda(t, in(t, dom) <=> ∃(a, in(pair(t, a), f)) /\ in(t, x))
+      )
+    )
+
+    val simplerCharacterization = have((relationDomain(g) === dom) <=> ∀(t, in(t, dom) <=> ∃(a, in(pair(t, a), f)) /\ in(t, x))) by Tautology.from(characterization, lastStep)
+
+    have(thesis) by Tautology.from(domCharacterization, simplerCharacterization)
+  }
 
   // TODO: any subset of a functional is functional
   // TODO: a functional over something restricted to x is still functional

From 652af81ef7f6ff73016281750a43429be31ee11b Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Thu, 27 Apr 2023 16:02:16 +0200
Subject: [PATCH 29/68] Rename binaryRelation to binaryFunction

---
 src/main/scala/lisa/mathematics/GroupTheory.scala | 6 +++---
 1 file changed, 3 insertions(+), 3 deletions(-)

diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala
index 05ee3166..4aa172ba 100644
--- a/src/main/scala/lisa/mathematics/GroupTheory.scala
+++ b/src/main/scala/lisa/mathematics/GroupTheory.scala
@@ -130,10 +130,10 @@ object GroupTheory extends lisa.Main {
   //
 
   /**
-   * Binary relation --- `*` is a binary relation on `G` if it associates to each pair of elements of `G`
+   * Binary function --- `*` is a binary function on `G` if it associates to each pair of elements of `G`
    * a unique element in `G`. In other words, `*` is a function `G × G -> G`.
    */
-  val binaryRelation = DEF(G, *) --> functionFrom(*, cartesianProduct(G, G), G)
+  val binaryFunction = DEF(G, *) --> functionFrom(*, cartesianProduct(G, G), G)
 
   /**
    * Shorthand for `x * y`.
@@ -170,7 +170,7 @@ object GroupTheory extends lisa.Main {
    * Group --- A group (G, *) is a set along with a law of composition `*`, satisfying [[associativity]], [[identityExistence]]
    * and [[inverseExistence]].
    */
-  val group = DEF(G, *) --> binaryRelation(G, *) /\ associativity(G, *) /\ identityExistence(G, *) /\ inverseExistence(G, *)
+  val group = DEF(G, *) --> binaryFunction(G, *) /\ associativity(G, *) /\ identityExistence(G, *) /\ inverseExistence(G, *)
 
   /**
    * Identity uniqueness --- In a group (G, *), an identity element is unique, i.e. if both `e * x = x * e = x` and

From 4565ec68980b485c9b84f2f22db6ccfe17a248c2 Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Sun, 30 Apr 2023 16:20:29 +0200
Subject: [PATCH 30/68] Fix inverseUniqueness proof assumptions bug

---
 .../scala/lisa/mathematics/GroupTheory.scala  | 19 +++++++------------
 1 file changed, 7 insertions(+), 12 deletions(-)

diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala
index 4aa172ba..0c45cd92 100644
--- a/src/main/scala/lisa/mathematics/GroupTheory.scala
+++ b/src/main/scala/lisa/mathematics/GroupTheory.scala
@@ -12,7 +12,7 @@ import lisa.mathematics.SetTheory.*
  */
 object GroupTheory extends lisa.Main {
   // Groups
-  private val G = variable
+  private val G, H = variable
 
   // Group laws
   private val * = variable
@@ -240,11 +240,6 @@ object GroupTheory extends lisa.Main {
     val uniqueness = have((group(G, *), x ∈ G, isInverse(y, x, G, *), isInverse(z, x, G, *)) |- (y === z)) subproof {
       val inverseMembership = have(isInverse(y, x, G, *) |- y ∈ G) by Tautology.from(isInverse.definition)
 
-      assume(group(G, *))
-      assume(x ∈ G)
-      assume(isInverse(y, x, G, *))
-      assume(isInverse(z, x, G, *))
-
       // 1. (yx)z = z
       val leftNeutrality = have((group(G, *), x ∈ G, isInverse(y, x, G, *), z ∈ G) |- (op(op(y, *, x), *, z) === z)) subproof {
         assume(group(G, *))
@@ -256,7 +251,7 @@ object GroupTheory extends lisa.Main {
         thenHave(z ∈ G ==> ((op(op(y, *, x), *, z) === z) /\ (op(z, *, op(y, *, x)) === z))) by InstantiateForall(z)
         thenHave(op(op(y, *, x), *, z) === z) by Tautology
       }
-      val firstEq = have(op(op(y, *, x), *, z) === z) by Cut(inverseMembership of (y -> z), leftNeutrality)
+      val firstEq = have((group(G, *), x ∈ G, isInverse(y, x, G, *), isInverse(z, x, G, *)) |- op(op(y, *, x), *, z) === z) by Cut(inverseMembership of (y -> z), leftNeutrality)
 
       // 2. (yx)z = y(xz)
       val permuteParentheses = have((group(G, *), x ∈ G, y ∈ G, z ∈ G) |- (op(op(y, *, x), *, z) === op(y, *, op(x, *, z)))) subproof {
@@ -274,8 +269,8 @@ object GroupTheory extends lisa.Main {
         thenHave(op(op(y, *, x), *, z) === op(y, *, op(x, *, z))) by Tautology
       }
       val associativityCut = have((group(G, *), x ∈ G /\ y ∈ G /\ z ∈ G) |- (op(op(y, *, x), *, z) === op(y, *, op(x, *, z)))) by Restate.from(permuteParentheses)
-      val memberships = have(x ∈ G /\ y ∈ G /\ z ∈ G) by Tautology.from(inverseMembership of (y -> y), inverseMembership of (y -> z))
-      val secondEq = have(op(op(y, *, x), *, z) === op(y, *, op(x, *, z))) by Cut(memberships, associativityCut)
+      val memberships = have((x ∈ G, isInverse(y, x, G, *), isInverse(z, x, G, *)) |- x ∈ G /\ y ∈ G /\ z ∈ G) by Tautology.from(inverseMembership of (y -> y), inverseMembership of (y -> z))
+      val secondEq = have((group(G, *), x ∈ G, isInverse(y, x, G, *), isInverse(z, x, G, *)) |- op(op(y, *, x), *, z) === op(y, *, op(x, *, z))) by Cut(memberships, associativityCut)
 
       // 3. y(xz) = y
       val rightNeutrality = have((group(G, *), x ∈ G, y ∈ G, isInverse(z, x, G, *)) |- (op(y, *, op(x, *, z)) === y)) subproof {
@@ -288,17 +283,17 @@ object GroupTheory extends lisa.Main {
         thenHave(y ∈ G ==> ((op(op(x, *, z), *, y) === y) /\ (op(y, *, op(x, *, z)) === y))) by InstantiateForall(y)
         thenHave(op(y, *, op(x, *, z)) === y) by Tautology
       }
-      val thirdEq = have(op(y, *, op(x, *, z)) === y) by Cut(inverseMembership of (y -> y), rightNeutrality)
+      val thirdEq = have((group(G, *), x ∈ G, isInverse(y, x, G, *), isInverse(z, x, G, *)) |- op(y, *, op(x, *, z)) === y) by Cut(inverseMembership of (y -> y), rightNeutrality)
 
       // Conclude by transitivity
 
       // 4. z = y(xz)
-      val fourthEq = have(z === op(y, *, op(x, *, z))) by Tautology.from(
+      val fourthEq = have((group(G, *), x ∈ G, isInverse(y, x, G, *), isInverse(z, x, G, *)) |- z === op(y, *, op(x, *, z))) by Tautology.from(
         firstEq, secondEq, equalityTransitivity of (x -> z, y -> op(op(y, *, x), *, z), z -> op(y, *, op(x, *, z)))
       )
 
       // 5. z = y
-      have(z === y) by Tautology.from(
+      have((group(G, *), x ∈ G, isInverse(y, x, G, *), isInverse(z, x, G, *)) |- z === y) by Tautology.from(
         thirdEq, fourthEq, equalityTransitivity of (x -> z, y -> op(y, *, op(x, *, z)), z -> y)
       )
     }

From a8c6dc636a7c391abd7f9d96e5b7df745a136f5d Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Sun, 30 Apr 2023 18:32:55 +0200
Subject: [PATCH 31/68] Add subgroup definitions and prove subgroup operation

---
 .../scala/lisa/mathematics/GroupTheory.scala  | 191 ++++++++++++++++--
 1 file changed, 179 insertions(+), 12 deletions(-)

diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala
index 0c45cd92..46dd8ab6 100644
--- a/src/main/scala/lisa/mathematics/GroupTheory.scala
+++ b/src/main/scala/lisa/mathematics/GroupTheory.scala
@@ -1,6 +1,6 @@
 package lisa.mathematics
 
-import lisa.automation.grouptheory.GroupTheoryTactics.{Cases, ExistenceAndUniqueness}
+import lisa.automation.grouptheory.GroupTheoryTactics.{Cases, Definition, ExistenceAndUniqueness}
 import lisa.automation.kernel.OLPropositionalSolver.Tautology
 import lisa.mathematics.FirstOrderLogic.{equalityTransitivity, existsOneImpliesExists, substitutionInUniquenessQuantifier}
 import lisa.mathematics.SetTheory.*
@@ -51,7 +51,7 @@ object GroupTheory extends lisa.Main {
       val substF = substituteVariables(completeDef, Map[VariableLabel, Term](u -> defaultValue))
       val substDef = substituteVariables(completeDef, Map[VariableLabel, Term](u -> v))
 
-      val completeUniquenessTheorem = Theorem(
+      val completeUniquenessTheorem = Lemma(
         ∃!(u, completeDef)
       ) {
         val case1 = have(prem |- ∃!(u, completeDef)) subproof {
@@ -136,9 +136,9 @@ object GroupTheory extends lisa.Main {
   val binaryFunction = DEF(G, *) --> functionFrom(*, cartesianProduct(G, G), G)
 
   /**
-   * Shorthand for `x * y`.
+   * Short-hand alias for `x * y`.
    */
-  val op = DEF(x, *, y) --> app(*, pair(x, y))
+  inline def op(x: Term, * : Term, y: Term) = app(*, pair(x, y))
 
   /**
    * Associativity --- `*` is associative (in `G`) if `(x * y) * z = x * (y * z)` for all `x, y, z` in `G`.
@@ -172,6 +172,58 @@ object GroupTheory extends lisa.Main {
    */
   val group = DEF(G, *) --> binaryFunction(G, *) /\ associativity(G, *) /\ identityExistence(G, *) /\ inverseExistence(G, *)
 
+  /**
+   * Group operation is functional -- The group operation `*` is functional.
+   *
+   * Similarly to the lemma below, follows more or less directly from the definitions, but still useful.
+   */
+  val groupOperationIsFunctional = Lemma(
+    group(G, *) |- functional(*)
+  ) {
+    assume(group(G, *))
+    have(binaryFunction(G, *)) by Tautology.from(group.definition)
+    have(functionFrom(*, cartesianProduct(G, G), G)) by Tautology.from(lastStep, binaryFunction.definition)
+    have(functional(*)) by Tautology.from(lastStep, functionFromImpliesFunctional of (f -> *, x -> cartesianProduct(G, G), y -> G))
+  }
+
+  /**
+   * Group operation domain -- The domain of a group law is the cartesian product of the group with itself.
+   *
+   * Follows directly from the definition of `binaryRelation`, but it is a useful lemma to have in some proofs.
+   */
+  val groupOperationDomain = Lemma(
+    group(G, *) |- relationDomain(*) === cartesianProduct(G, G)
+  ) {
+    assume(group(G, *))
+    have(binaryFunction(G, *)) by Tautology.from(group.definition)
+    have(functionFrom(*, cartesianProduct(G, G), G)) by Tautology.from(lastStep, binaryFunction.definition)
+    have(relationDomain(*) === cartesianProduct(G, G)) by Tautology.from(lastStep, functionFromImpliesDomainEq of(f -> *, x -> cartesianProduct(G, G), y -> G))
+  }
+
+  /**
+   * Lemma --- If `x` and `y` are two elements of the group, the pair `(x, y)` is in the relation domain of `*.
+   */
+  val groupPairInOperationDomain = Lemma(
+    (group(G, *), x ∈ G, y ∈ G) |- pair(x, y) ∈ relationDomain(*)
+  ) {
+    assume(group(G, *))
+    assume(x ∈ G)
+    assume(y ∈ G)
+
+    val a, b = variable
+
+    have(x ∈ G /\ y ∈ G) by Tautology
+    have(pair(x, y) ∈ cartesianProduct(G, G)) by Tautology.from(
+      lastStep, pairInCartesianProduct of (a -> x, b -> y, x -> G, y -> G)
+    )
+    thenHave((relationDomain(*) === cartesianProduct(G, G)) |- pair(x, y) ∈ relationDomain(*)) by RightSubstEq(
+      List((relationDomain(*), cartesianProduct(G, G))), lambda(z, pair(x, y) ∈ z)
+    )
+
+    have(thesis) by Cut(groupOperationDomain, lastStep)
+  }
+  show
+
   /**
    * Identity uniqueness --- In a group (G, *), an identity element is unique, i.e. if both `e * x = x * e = x` and
    * `f * x = x * f = x` for all `x`, then `e = f`.
@@ -194,22 +246,22 @@ object GroupTheory extends lisa.Main {
       // 1. ef = f
       have(∀(x, x ∈ G ==> ((op(e, *, x) === x) /\ (op(x, *, e) === x)))) by Tautology.from(isNeutral.definition)
       thenHave(f ∈ G ==> ((op(e, *, f) === f) /\ (op(f, *, e) === f))) by InstantiateForall(f)
-      val cut1 = thenHave(f ∈ G |- ((op(e, *, f) === f) /\ (op(f, *, e) === f))) by Restate
+      val neutrality = thenHave(f ∈ G |- ((op(e, *, f) === f) /\ (op(f, *, e) === f))) by Restate
 
-      have((op(e, *, f) === f) /\ (op(f, *, e) === f)) by Cut(membership of (e -> f), cut1)
+      have((op(e, *, f) === f) /\ (op(f, *, e) === f)) by Cut(membership of (e -> f), neutrality)
       val firstEq = thenHave(op(e, *, f) === f) by Tautology
 
       // 2. ef = e
-      val cut2 = have(e ∈ G |- ((op(f, *, e) === e) /\ (op(e, *, f) === e))) by InstFunSchema(Map(e -> f, f -> e))(cut1)
+      have(e ∈ G |- ((op(f, *, e) === e) /\ (op(e, *, f) === e))) by InstFunSchema(Map(e -> f, f -> e))(neutrality)
 
-      have((op(f, *, e) === e) /\ (op(e, *, f) === e)) by Cut(membership of (e -> e), cut2)
+      have((op(f, *, e) === e) /\ (op(e, *, f) === e)) by Cut(membership of (e -> e), lastStep)
       val secondEq = thenHave(e === op(e, *, f)) by Tautology
 
       // 3. Conclude by transitivity
       val eqs = have((e === op(e, *, f)) /\ (op(e, *, f) === f)) by RightAnd(secondEq, firstEq)
-      val cut3 = have(((e === op(e, *, f)) /\ (op(e, *, f) === f)) |- (e === f)) by Tautology.from(equalityTransitivity of (x -> e, y -> op(e, *, f), z -> f))
+      have(((e === op(e, *, f)) /\ (op(e, *, f) === f)) |- (e === f)) by Tautology.from(equalityTransitivity of (x -> e, y -> op(e, *, f), z -> f))
 
-      have(e === f) by Cut(eqs, cut3)
+      have(e === f) by Cut(eqs, lastStep)
     }
 
     have(group(G, *) |- ∃!(e, isNeutral(e, G, *))) by ExistenceAndUniqueness(isNeutral(e, G, *))(existence, uniqueness)
@@ -319,9 +371,9 @@ object GroupTheory extends lisa.Main {
 
       have(y ∈ G /\ isNeutral(op(x, *, y), G, *) /\ isNeutral(op(y, *, x), G, *)) by Tautology.from(isInverse.definition)
       thenHave(isNeutral(op(y, *, x), G, *) /\ isNeutral(op(x, *, y), G, *)) by Tautology
-      val definition = thenHave(x ∈ G /\ isNeutral(op(y, *, x), G, *) /\ isNeutral(op(x, *, y), G, *)) by Tautology
+      thenHave(x ∈ G /\ isNeutral(op(y, *, x), G, *) /\ isNeutral(op(x, *, y), G, *)) by Tautology
 
-      have(thesis) by Tautology.from(definition, isInverse.definition of (y -> x, x -> y))
+      have(thesis) by Tautology.from(lastStep, isInverse.definition of (y -> x, x -> y))
     }
     thenHave(group(G, *) |- x ∈ G ==> (isInverse(y, x, G, *) ==> isInverse(x, y, G, *))) by Restate
     thenHave(thesis) by RightForall
@@ -377,4 +429,119 @@ object GroupTheory extends lisa.Main {
     thenHave(x ∈ G ==> (x === u)) by Restate
     thenHave(thesis) by RightForall
   }
+
+  //
+  // 1.1 Subgroups
+  //
+
+  /**
+   * Subgroup --- `H` is a subgroup of `(G, *)` if `H` is a subset of `G`, such that the restriction of `*` to `H` is
+   * a group law for `H`, i.e. `(H, *_H)` is a group.
+   *
+   * We denote `H <= G` for `H` a subgroup of `G`.
+   */
+  val subgroup = DEF(H, G, *) --> group(G, *) /\ subset(H, G) /\ group(H, restrictedFunction(*, cartesianProduct(H, H)))
+
+  /**
+   * Lemma --- If `x` and `y` are two elements of the subgroup `H` of `(G, *)`, the pair belongs to the relation domain
+   * of the parent group's operation `*`.
+   *
+   * Analogous to [[groupPairInOperationDomain]], except that the considered relation is different.
+   */
+  val subgroupPairInParentOperationDomain = Lemma(
+    (subgroup(H, G, *), x ∈ H, y ∈ H) |- pair(x, y) ∈ relationDomain(*)
+  ) {
+    assume(subgroup(H, G, *))
+    assume(x ∈ H)
+    assume(y ∈ H)
+
+    val a, b = variable
+
+    have(subset(H, G)) by Tautology.from(subgroup.definition)
+    have(∀(x, x ∈ H ==> x ∈ G)) by Tautology.from(lastStep, subset.definition of(x -> H, y -> G))
+    val subsetDef = thenHave(x ∈ H ==> x ∈ G) by InstantiateForall(x)
+
+    val left = have(x ∈ G) by Tautology.from(subsetDef)
+    val right = have(y ∈ G) by Tautology.from(subsetDef of (x -> y))
+
+    have(group(G, *)) by Tautology.from(subgroup.definition)
+
+    have(thesis) by Tautology.from(lastStep, left, right, groupPairInOperationDomain)
+  }
+
+  /**
+   * Theorem --- The subgroup operation is exactly the same as in the above group, i.e. if `(G, *)` is a group, `H` a
+   * subgroup of `G`, then for elements `x, y ∈ H` we have `x ★ y = x * y`, where `★ = *_H`.
+   */
+  val subgroupOperation = Theorem(
+    (subgroup(H, G, *), x ∈ H, y ∈ H) |- (op(x, restrictedFunction(*, cartesianProduct(H, H)), y) === op(x, *, y))
+  ) {
+    val H2 = cartesianProduct(H, H)
+    val ★ = restrictedFunction(*, H2)
+    val r = app(★, pair(x, y))
+
+    // We characterize op(x, *, y), and show that op(x, ★, y) satisfies all requirements
+    have(
+      ∀(z, (z === app(*, pair(x, y))) <=> (((functional(*) /\ in(pair(x, y), relationDomain(*))) ==> in(pair(pair(x, y), z), *)) /\
+                                          ((!functional(*) \/ !in(pair(x, y), relationDomain(*))) ==> (z === ∅))))
+    ) by Tautology.from(app.definition of (f -> *, x -> pair(x, y)), functionApplicationUniqueness)
+    val characterization = thenHave(
+      (r === app(*, pair(x, y))) <=> (((functional(*) /\ in(pair(x, y), relationDomain(*))) ==> in(pair(pair(x, y), r), *)) /\
+                                     ((!functional(*) \/ !in(pair(x, y), relationDomain(*))) ==> (r === ∅)))
+    ) by InstantiateForall(r)
+
+    // Prove that the premises of the first implication hold
+    val leftPremise = have(subgroup(H, G, *) |- functional(*)) subproof {
+      assume(subgroup(H, G, *))
+      have(group(G, *)) by Tautology.from(subgroup.definition)
+      have(functional(*)) by Tautology.from(lastStep, groupOperationIsFunctional)
+    }
+
+    val premises = have((subgroup(H, G, *), x ∈ H, y ∈ H) |- functional(*) /\ pair(x, y) ∈ relationDomain(*)) by RightAnd(
+      leftPremise, subgroupPairInParentOperationDomain
+    )
+
+    // We show that op(x, ★, y) satisfies the conclusion of the implication
+    val appDef = have(
+      (functional(★), pair(x, y) ∈ relationDomain(★)) |- pair(pair(x, y), r) ∈ ★
+    ) by Definition(app, functionApplicationUniqueness)(★, pair(x, y))
+
+    // Reduce the assumptions of the definition to our subgroup assumption
+    val reduction1 = have(subgroup(H, G, *) |- group(H, ★)) by Tautology.from(subgroup.definition)
+    val reduction2 = have(subgroup(H, G, *) |- functional(★)) by Tautology.from(lastStep, groupOperationIsFunctional of (G -> H, * -> ★))
+
+    val reduction3 = have((subgroup(H, G, *), pair(x, y) ∈ relationDomain(★)) |- pair(pair(x, y), r) ∈ ★) by Tautology.from(reduction2, appDef)
+
+    have((subgroup(H, G, *), x ∈ H, y ∈ H) |- pair(x, y) ∈ relationDomain(★)) by Cut(
+      reduction1, groupPairInOperationDomain of (G -> H, * -> ★)
+    )
+    val reducedDef = have((subgroup(H, G, *), x ∈ H, y ∈ H) |- pair(pair(x, y), r) ∈ ★) by Cut(lastStep, reduction3)
+
+    have(∀(u, (u ∈ ★) <=> (u ∈ * /\ ∃(y, ∃(z, y ∈ H2 /\ (u === pair(y, z))))))) by Definition(restrictedFunction, restrictedFunctionUniqueness)(*, H2)
+    thenHave((u ∈ ★) <=> (u ∈ * /\ ∃(y, ∃(z, y ∈ H2 /\ (u === pair(y, z)))))) by InstantiateForall(u)
+    thenHave(u ∈ ★ ==> u ∈ *) by Tautology
+
+    val satisfaction = have((subgroup(H, G, *), x ∈ H, y ∈ H) |- pair(pair(x, y), r) ∈ *) by Tautology.from(
+      lastStep of (u -> pair(pair(x, y), r)), reducedDef
+    )
+
+    // Reconstruct the whole definition
+    assume(subgroup(H, G, *))
+    assume(x ∈ H)
+    assume(y ∈ H)
+
+    val pos = have((functional(*) /\ pair(x, y) ∈ relationDomain(*)) ==> pair(pair(x, y), r) ∈ *) by Tautology.from(premises, satisfaction)
+
+    have(!(functional(*) /\ pair(x, y) ∈ relationDomain(*)) |- ()) by LeftNot(premises)
+    thenHave(!functional(*) \/ !(pair(x, y) ∈ relationDomain(*)) |- ()) by Restate
+    thenHave(!functional(*) \/ !(pair(x, y) ∈ relationDomain(*)) |- (r === emptySet())) by Weakening
+    val neg = thenHave((!functional(*) \/ !(pair(x, y) ∈ relationDomain(*))) ==> (r === emptySet())) by Restate
+
+    have(
+      ((functional(*) /\ pair(x, y) ∈ relationDomain(*)) ==> pair(pair(x, y), r) ∈ *) /\
+      ((!functional(*) \/ !(pair(x, y) ∈ relationDomain(*))) ==> (r === emptySet()))
+    ) by RightAnd(pos, neg)
+
+    have(thesis) by Tautology.from(lastStep, characterization)
+  }
 }

From 72489107446bcd53b8074de52e86c133bff79d79 Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Sun, 30 Apr 2023 18:41:03 +0200
Subject: [PATCH 32/68] Simply GroupTheory tactics

---
 .../grouptheory/GroupTheoryTactics.scala      | 24 ++++++++++---------
 1 file changed, 13 insertions(+), 11 deletions(-)

diff --git a/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala b/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala
index 4f8d06e3..4033ff66 100644
--- a/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala
+++ b/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala
@@ -18,8 +18,10 @@ object GroupTheoryTactics {
    * Γ, Γ' |- ∃!x. φ, Δ, Δ'
    * 
* - * This tactic separates the existence and the uniqueness proofs, which are often easier to prove independently, - * and thus is easier to use than [[RightExistsOne]]. + * This tactic separates the existence and the uniqueness proofs, which are often easier to prove independently, at + * the expense of brevity. + * + * @see [[RightExistsOne]]. */ object ExistenceAndUniqueness extends ProofTactic { def withParameters(using proof: SetTheoryLibrary.Proof, om: OutputManager)(phi: Formula, x: VariableLabel, y: VariableLabel) @@ -35,15 +37,15 @@ object GroupTheoryTactics { if (x == y) { proof.InvalidProofTactic("x and y can not be equal.") } else if (!contains(existenceSeq.right, existenceFormula)) { - proof.InvalidProofTactic(s"Existence sequent conclusion does not contain ∃$x. φ.") + proof.InvalidProofTactic(s"Existence sequent conclusion does not contain ∃x. φ.") } else if (!contains(uniquenessSeq.left, phi)) { proof.InvalidProofTactic("Uniqueness sequent premises do not contain φ.") } else if (!contains(uniquenessSeq.left, substPhi)) { - proof.InvalidProofTactic(s"Uniqueness sequent premises do not contain φ[$y/$x].") + proof.InvalidProofTactic(s"Uniqueness sequent premises do not contain φ[y/x].") } else if (!contains(uniquenessSeq.right, x === y) && !contains(uniquenessSeq.right, y === x)) { - proof.InvalidProofTactic(s"Uniqueness sequent conclusion does not contain $x = $y") + proof.InvalidProofTactic(s"Uniqueness sequent conclusion does not contain x = y") } else if (!contains(bot.right, uniqueExistenceFormula)) { - proof.InvalidProofTactic(s"Bottom sequent conclusion does not contain ∃!$x. φ") + proof.InvalidProofTactic(s"Bottom sequent conclusion does not contain ∃!x. φ") } // Checking pivots @@ -71,9 +73,9 @@ object GroupTheoryTactics { thenHave(phi |- ∀(y, (x === y) <=> substPhi)) by RightForall thenHave(phi |- ∃(x, ∀(y, (x === y) <=> substPhi))) by RightExists thenHave(∃(x, phi) |- ∃(x, ∀(y, (x === y) <=> substPhi))) by LeftExists - val cut = thenHave(∃(x, phi) |- ∃!(x, phi)) by RightExistsOne + thenHave(∃(x, phi) |- ∃!(x, phi)) by RightExistsOne - have(bot) by Cut(existence, cut) + have(bot) by Cut(existence, lastStep) } } } @@ -95,7 +97,7 @@ object GroupTheoryTactics { val commonVars = bot.right.collect { case BinderFormula(ExistsOne, x, f) if isSame(f, phi) && existsVars.contains(x) => x } - if (commonVars.isEmpty || commonVars.size > 1) { + if (commonVars.size != 1) { return proof.InvalidProofTactic("Could not infer correct variable x in quantifiers.") } @@ -139,8 +141,8 @@ object GroupTheoryTactics { // Check if the definition is conditional val method = expr(xs) match { case ConnectorFormula(And, Seq( - ConnectorFormula(Implies, Seq(a, _)), - ConnectorFormula(Implies, Seq(b, _)) + ConnectorFormula(Implies, Seq(a, _)), + ConnectorFormula(Implies, Seq(b, _)) )) if isSame(neg(a), b) => conditional case _ => unconditional From 5650fcda804c90b7bbcdac265388961c5f61ac53 Mon Sep 17 00:00:00 2001 From: dhalilov Date: Sun, 30 Apr 2023 22:54:18 +0200 Subject: [PATCH 33/68] Revert "Simply GroupTheory tactics" This reverts commit 72489107446bcd53b8074de52e86c133bff79d79. --- .../grouptheory/GroupTheoryTactics.scala | 24 +++++++++---------- 1 file changed, 11 insertions(+), 13 deletions(-) diff --git a/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala b/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala index 4033ff66..4f8d06e3 100644 --- a/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala +++ b/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala @@ -18,10 +18,8 @@ object GroupTheoryTactics { * Γ, Γ' |- ∃!x. φ, Δ, Δ' *
* - * This tactic separates the existence and the uniqueness proofs, which are often easier to prove independently, at - * the expense of brevity. - * - * @see [[RightExistsOne]]. + * This tactic separates the existence and the uniqueness proofs, which are often easier to prove independently, + * and thus is easier to use than [[RightExistsOne]]. */ object ExistenceAndUniqueness extends ProofTactic { def withParameters(using proof: SetTheoryLibrary.Proof, om: OutputManager)(phi: Formula, x: VariableLabel, y: VariableLabel) @@ -37,15 +35,15 @@ object GroupTheoryTactics { if (x == y) { proof.InvalidProofTactic("x and y can not be equal.") } else if (!contains(existenceSeq.right, existenceFormula)) { - proof.InvalidProofTactic(s"Existence sequent conclusion does not contain ∃x. φ.") + proof.InvalidProofTactic(s"Existence sequent conclusion does not contain ∃$x. φ.") } else if (!contains(uniquenessSeq.left, phi)) { proof.InvalidProofTactic("Uniqueness sequent premises do not contain φ.") } else if (!contains(uniquenessSeq.left, substPhi)) { - proof.InvalidProofTactic(s"Uniqueness sequent premises do not contain φ[y/x].") + proof.InvalidProofTactic(s"Uniqueness sequent premises do not contain φ[$y/$x].") } else if (!contains(uniquenessSeq.right, x === y) && !contains(uniquenessSeq.right, y === x)) { - proof.InvalidProofTactic(s"Uniqueness sequent conclusion does not contain x = y") + proof.InvalidProofTactic(s"Uniqueness sequent conclusion does not contain $x = $y") } else if (!contains(bot.right, uniqueExistenceFormula)) { - proof.InvalidProofTactic(s"Bottom sequent conclusion does not contain ∃!x. φ") + proof.InvalidProofTactic(s"Bottom sequent conclusion does not contain ∃!$x. φ") } // Checking pivots @@ -73,9 +71,9 @@ object GroupTheoryTactics { thenHave(phi |- ∀(y, (x === y) <=> substPhi)) by RightForall thenHave(phi |- ∃(x, ∀(y, (x === y) <=> substPhi))) by RightExists thenHave(∃(x, phi) |- ∃(x, ∀(y, (x === y) <=> substPhi))) by LeftExists - thenHave(∃(x, phi) |- ∃!(x, phi)) by RightExistsOne + val cut = thenHave(∃(x, phi) |- ∃!(x, phi)) by RightExistsOne - have(bot) by Cut(existence, lastStep) + have(bot) by Cut(existence, cut) } } } @@ -97,7 +95,7 @@ object GroupTheoryTactics { val commonVars = bot.right.collect { case BinderFormula(ExistsOne, x, f) if isSame(f, phi) && existsVars.contains(x) => x } - if (commonVars.size != 1) { + if (commonVars.isEmpty || commonVars.size > 1) { return proof.InvalidProofTactic("Could not infer correct variable x in quantifiers.") } @@ -141,8 +139,8 @@ object GroupTheoryTactics { // Check if the definition is conditional val method = expr(xs) match { case ConnectorFormula(And, Seq( - ConnectorFormula(Implies, Seq(a, _)), - ConnectorFormula(Implies, Seq(b, _)) + ConnectorFormula(Implies, Seq(a, _)), + ConnectorFormula(Implies, Seq(b, _)) )) if isSame(neg(a), b) => conditional case _ => unconditional From 2e023c3d224a42254b093be84cc7669e1550cbc8 Mon Sep 17 00:00:00 2001 From: dhalilov Date: Sun, 30 Apr 2023 23:22:05 +0200 Subject: [PATCH 34/68] Fix merge substituteVariable bug --- .../lisa/automation/grouptheory/GroupTheoryTactics.scala | 6 +++--- src/main/scala/lisa/mathematics/GroupTheory.scala | 4 ++-- 2 files changed, 5 insertions(+), 5 deletions(-) diff --git a/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala b/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala index 4f8d06e3..f5bdaaa4 100644 --- a/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala +++ b/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala @@ -27,7 +27,7 @@ object GroupTheoryTactics { val existenceSeq = proof.getSequent(existence) val uniquenessSeq = proof.getSequent(uniqueness) - lazy val substPhi = substituteVariables(phi, Map[VariableLabel, Term](x -> y)) + lazy val substPhi = substituteVariables(phi, Map[VariableLabel, Term](x -> y), Seq()) lazy val existenceFormula = ∃(x, phi) lazy val uniqueExistenceFormula = ∃!(x, phi) @@ -104,10 +104,10 @@ object GroupTheoryTactics { // Infer y from the equalities in the uniqueness sequent uniquenessSeq.right.collectFirst { case PredicateFormula(`equality`, List(Term(`x`, _), Term(y: VariableLabel, _))) - if x != y && contains(uniquenessSeq.left, substituteVariables(phi, Map[VariableLabel, Term](x -> y))) => y + if x != y && contains(uniquenessSeq.left, substituteVariables(phi, Map[VariableLabel, Term](x -> y), Seq())) => y case PredicateFormula(`equality`, List(Term(y: VariableLabel, _), Term(`x`, _))) - if x != y && contains(uniquenessSeq.left, substituteVariables(phi, Map[VariableLabel, Term](x -> y))) => y + if x != y && contains(uniquenessSeq.left, substituteVariables(phi, Map[VariableLabel, Term](x -> y), Seq())) => y } match { case Some(y) => ExistenceAndUniqueness.withParameters(phi, x, y)(existence, uniqueness)(bot) case None => proof.InvalidProofTactic("Could not infer correct variable y in uniqueness sequent.") diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala index 46dd8ab6..73fb25e5 100644 --- a/src/main/scala/lisa/mathematics/GroupTheory.scala +++ b/src/main/scala/lisa/mathematics/GroupTheory.scala @@ -48,8 +48,8 @@ object GroupTheory extends lisa.Main { } else { val prem = if (seq.left.size == 1) seq.left.head else And(seq.left.toSeq: _*) val completeDef = (prem ==> f) /\ (!prem ==> (u === defaultValue)) - val substF = substituteVariables(completeDef, Map[VariableLabel, Term](u -> defaultValue)) - val substDef = substituteVariables(completeDef, Map[VariableLabel, Term](u -> v)) + val substF = substituteVariables(completeDef, Map[VariableLabel, Term](u -> defaultValue), Seq()) + val substDef = substituteVariables(completeDef, Map[VariableLabel, Term](u -> v), Seq()) val completeUniquenessTheorem = Lemma( ∃!(u, completeDef) From 52e8f62b7e753041e7c31f82f08eb55174edb7d4 Mon Sep 17 00:00:00 2001 From: dhalilov Date: Thu, 4 May 2023 23:45:35 +0200 Subject: [PATCH 35/68] Prove group is non empty --- .../scala/lisa/mathematics/GroupTheory.scala | 26 +++++++++++++++++++ 1 file changed, 26 insertions(+) diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala index 73fb25e5..9300ee56 100644 --- a/src/main/scala/lisa/mathematics/GroupTheory.scala +++ b/src/main/scala/lisa/mathematics/GroupTheory.scala @@ -272,6 +272,32 @@ object GroupTheory extends lisa.Main { */ val identity = DEF(G, *) --> TheConditional(e, isNeutral(e, G, *))(identityUniqueness) + /** + * Theorem --- The identity element belongs to the group. + * + * This might seem like a silly theorem, but it is useful in proving that groups are non-empty by definition. + */ + val identityInGroup = Theorem( + group(G, *) |- identity(G, *) ∈ G + ) { + assume(group(G, *)) + have(isNeutral(identity(G, *), G, *)) by Definition(identity, identityUniqueness)(G, *) + have(thesis) by Tautology.from( + lastStep, isNeutral.definition of (e -> identity(G, *)) + ) + } + + /** + * Theorem --- A group is non-empty. + * + * Direct corollary of [[identityInGroup]]. + */ + val groupNonEmpty = Theorem( + group(G, *) |- (G !== ∅) + ) { + have(thesis) by Cut(identityInGroup, setWithElementNonEmpty of (x -> G, y -> identity(G, *))) + } + /** * Theorem --- The inverse of an element `x` (i.e. `y` such that `x * y = y * x = e`) in `G` is unique. */ From c645f5706004799b3824881f57246d9940048e82 Mon Sep 17 00:00:00 2001 From: dhalilov Date: Fri, 5 May 2023 09:34:28 +0200 Subject: [PATCH 36/68] Prove group is subgroup of itself --- .../scala/lisa/mathematics/GroupTheory.scala | 62 ++++++++++++++++++- .../scala/lisa/mathematics/SetTheory.scala | 17 +++++ .../lisa/settheory/SetTheoryLibrary.scala | 1 + 3 files changed, 78 insertions(+), 2 deletions(-) diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala index 9300ee56..13ac791b 100644 --- a/src/main/scala/lisa/mathematics/GroupTheory.scala +++ b/src/main/scala/lisa/mathematics/GroupTheory.scala @@ -6,7 +6,7 @@ import lisa.mathematics.FirstOrderLogic.{equalityTransitivity, existsOneImpliesE import lisa.mathematics.SetTheory.* /** - * Group theory, following Chapter 2 of S. Lang "Undergraduate Algebra". + * Group theory, developed following Chapter 2 of S. Lang "Undergraduate Algebra". * * Book : [[https://link.springer.com/book/10.1007/978-1-4684-9234-7]] */ @@ -222,7 +222,6 @@ object GroupTheory extends lisa.Main { have(thesis) by Cut(groupOperationDomain, lastStep) } - show /** * Identity uniqueness --- In a group (G, *), an identity element is unique, i.e. if both `e * x = x * e = x` and @@ -468,6 +467,44 @@ object GroupTheory extends lisa.Main { */ val subgroup = DEF(H, G, *) --> group(G, *) /\ subset(H, G) /\ group(H, restrictedFunction(*, cartesianProduct(H, H))) + /** + * Lemma --- A group is a subgroup of itself, i.e. the subgroup relationship is reflexive. + */ + val groupIsSubgroupOfItself = Theorem( + group(G, *) |- subgroup(G, G, *) + ) { + val condition1 = have(group(G, *) |- group(G, *)) by Hypothesis + val condition2 = have(subset(G, G)) by Restate.from(subsetReflexivity of (x -> G)) + + // For condition 3, we need to substitute everything using the 3 following equalities: + // 1. restrictedFunction(*, relationDomain(*)) === * (restrictedFunctionCancellation) + // 2. relationDomain(*) === cartesianProduct(G, G) (groupOperationDomain) + // 3. restrictedFunction(*, cartesianProduct(G, G)) === * (derived from 1. and 2.) + + val substitution = have((group(G, *), restrictedFunction(*, cartesianProduct(G, G)) === *) |- group(G, restrictedFunction(*, cartesianProduct(G, G)))) by RightSubstEq( + List((restrictedFunction(*, cartesianProduct(G, G)), *)), lambda(z, group(G, z)) + )(condition1) + + val eq3 = have(group(G, *) |- restrictedFunction(*, cartesianProduct(G, G)) === *) subproof { + val eq1 = have(restrictedFunction(*, relationDomain(*)) === *) by Restate.from(restrictedFunctionCancellation of (f -> *)) + thenHave((relationDomain(*) === cartesianProduct(G, G)) |- restrictedFunction(*, cartesianProduct(G, G)) === *) by RightSubstEq( + List((relationDomain(*), cartesianProduct(G, G))), lambda(z, restrictedFunction(*, z) === *) + ) + + have(thesis) by Cut(groupOperationDomain, lastStep) + } + + val condition3 = have(group(G, *) |- group(G, restrictedFunction(*, cartesianProduct(G, G)))) by Cut(eq3, substitution) + + have(group(G, *) |- group(G, *) /\ subset(G, G) /\ group(G, restrictedFunction(*, cartesianProduct(G, G)))) by RightAnd(condition1, condition2, condition3) + have(thesis) by Tautology.from(lastStep, subgroup.definition of (G -> G, H -> G)) + } + + /** + * Proper subgroup --- `H` is a proper subgroup of `(G, *)` if `H` is a subgroup of `G` and `H != G`. + */ + val properSubgroup = DEF(H, G, *) --> subgroup(H, G, *) /\ (H !== G) + /** * Lemma --- If `x` and `y` are two elements of the subgroup `H` of `(G, *)`, the pair belongs to the relation domain * of the parent group's operation `*`. @@ -570,4 +607,25 @@ object GroupTheory extends lisa.Main { have(thesis) by Tautology.from(lastStep, characterization) } + + /** + * Identity in subgroup --- The identity element of the subgroup is exactly the identity element of the parent group. + */ + val subgroupIdentity = Theorem( + subgroup(H, G, *) |- identity(H, restrictedFunction(*, cartesianProduct(H, H))) === identity(G, *) + ) { + /*// The crux of the proof relies on proving that identity(G, *) ∈ H + // From there we can derive quite easily the theorem using [[subgroupOperation]] + have(subgroup(H, G, *) |- identity(G, *) ∈ H) subproof { + // The main idea is to notice that for H = G, identity(G, *) must be in H per [[identityInGroup]] + // hence identity(G, *) ∉ H would be contradictory + have((subgroup(H, G, *) ==> identity(G, *) ∉ H) |- (subgroup(H, G, *) ==> identity(G, *) ∉ H)) by Hypothesis + thenHave((subgroup(H, G, *) ==> identity(G, *) ∉ H) |- ∀(H, subgroup(H, G, *) ==> identity(G, *) ∉ H)) by RightForall + } + + have(group(G, *) |- isNeutral(identity(G, *), G, *)) by Definition(identity, identityUniqueness)(G, *) + have(group(G, *) |- identity(G, *) ∈ G /\ ∀(x, (x ∈ G) ==> ((op(identity(G, *), *, x) === x) /\ (op(identity(G, *), *, e) === x)))) by Tautology.from(isNeutral.definition of (e -> identity(G, *)))*/ + + sorry + } } diff --git a/src/main/scala/lisa/mathematics/SetTheory.scala b/src/main/scala/lisa/mathematics/SetTheory.scala index 19e5e319..30999da5 100644 --- a/src/main/scala/lisa/mathematics/SetTheory.scala +++ b/src/main/scala/lisa/mathematics/SetTheory.scala @@ -2338,6 +2338,23 @@ object SetTheory extends lisa.Main { have(thesis) by Tautology.from(domCharacterization, simplerCharacterization) } + /** + * Restricted function cancellation --- Restricting a function to its relation domain does nothing. + */ + val restrictedFunctionCancellation = Theorem( + functional(f) |- restrictedFunction(f, relationDomain(f)) === f + ) { + /*val g = restrictedFunction(f, relationDomain(f)) + + have(∀(t, in(t, relationDomain(f)) <=> ∃(a, in(pair(t, a), f)))) by Definition(relationDomain, relationDomainUniqueness)(f) + thenHave(in(y, relationDomain(f)) <=> ∃(a, in(pair(y, a), f))) by InstantiateForall(y) + + have(∀(t, in(t, g) <=> (in(t, f) /\ ∃(y, ∃(z, in(y, relationDomain(f)) /\ (t === pair(y, z))))))) by Definition( + restrictedFunction, restrictedFunctionUniqueness + )(f, relationDomain(f))*/ + sorry + } + // TODO: any subset of a functional is functional // TODO: a functional over something restricted to x is still functional diff --git a/src/main/scala/lisa/settheory/SetTheoryLibrary.scala b/src/main/scala/lisa/settheory/SetTheoryLibrary.scala index 4dc84632..b4a81444 100644 --- a/src/main/scala/lisa/settheory/SetTheoryLibrary.scala +++ b/src/main/scala/lisa/settheory/SetTheoryLibrary.scala @@ -15,6 +15,7 @@ object SetTheoryLibrary extends Library with SetTheoryTGAxioms { val ∅ : Term = emptySet() extension (t: Term) { infix def ∈(u: Term): Formula = PredicateFormula(in, Seq(t, u)) + infix def ∉(u: Term): Formula = !(t ∈ u) } } From 188d1c05f9fb2164206dfec1bf5c1c8e284c6a1f Mon Sep 17 00:00:00 2001 From: dhalilov Date: Mon, 15 May 2023 11:22:17 +0200 Subject: [PATCH 37/68] Prove restrictedFunctionCancellation --- .../scala/lisa/mathematics/SetTheory.scala | 149 +++++++++++++++++- 1 file changed, 146 insertions(+), 3 deletions(-) diff --git a/src/main/scala/lisa/mathematics/SetTheory.scala b/src/main/scala/lisa/mathematics/SetTheory.scala index 30999da5..550ee637 100644 --- a/src/main/scala/lisa/mathematics/SetTheory.scala +++ b/src/main/scala/lisa/mathematics/SetTheory.scala @@ -1908,6 +1908,92 @@ object SetTheory extends lisa.Main { */ val relationRange = DEF(r) --> The(z, ∀(t, in(t, z) <=> ∃(a, in(pair(a, t), r))))(relationRangeUniqueness) + /** + * Theorem --- If `r` is a relation, then `r` is a relation between its domain and its range. + */ + val relationImpliesRelationBetweenDomainAndRange = Theorem( + relation(r) |- relationBetween(r, relationDomain(r), relationRange(r)) + ) { + // Lay out the definitions to apply them later + have(∀(t, in(t, relationDomain(r)) <=> ∃(b, in(pair(t, b), r)))) by Definition(relationDomain, relationDomainUniqueness)(r) + val relationDomainDef = thenHave(in(t, relationDomain(r)) <=> ∃(b, in(pair(t, b), r))) by InstantiateForall(t) + + have(∀(t, in(t, relationRange(r)) <=> ∃(a, in(pair(a, t), r)))) by Definition(relationRange, relationRangeUniqueness)(r) + val relationRangeDef = thenHave(in(t, relationRange(r)) <=> ∃(a, in(pair(a, t), r))) by InstantiateForall(t) + + // Start the proof + have(relation(r) |- ∃(x, ∃(y, relationBetween(r, x, y)))) by Tautology.from(relation.definition) + + have(relationBetween(r, x, y) |- subset(r, cartesianProduct(x, y))) by Tautology.from(relationBetween.definition of (a -> x, b -> y)) + have(relationBetween(r, x, y) |- ∀(t, in(t, r) ==> in(t, cartesianProduct(x, y)))) by Tautology.from( + lastStep, + subset.definition of (x -> r, y -> cartesianProduct(x, y)) + ) + thenHave(relationBetween(r, x, y) |- in(t, r) ==> in(t, cartesianProduct(x, y))) by InstantiateForall(t) + thenHave((relationBetween(r, x, y), in(t, r)) |- in(t, cartesianProduct(x, y))) by Restate + + // Apply the definition of the cartesian product + val relationDef = have((relationBetween(r, x, y), in(t, r)) |- ∃(a, ∃(b, (t === pair(a, b)) /\ in(a, x) /\ in(b, y)))) by Tautology.from( + lastStep, + elemOfCartesianProduct + ) + + // Show that x ⊇ relationDomain(r) and y ⊇ relationRange(r) + val memberships = have((in(t, r), (t === pair(a, b))) |- in(a, relationDomain(r)) /\ in(b, relationRange(r))) subproof { + have(in(t, r) |- in(t, r)) by Hypothesis + val membership = thenHave((in(t, r), (t === pair(a, b))) |- in(pair(a, b), r)) by Substitution + + assume(in(t, r)) + assume(t === pair(a, b)) + have(∃(b, in(pair(a, b), r))) by RightExists(membership) + val left = have(in(a, relationDomain(r))) by Tautology.from(lastStep, relationDomainDef of (t -> a)) + + have(∃(a, in(pair(a, b), r))) by RightExists(membership) + val right = have(in(b, relationRange(r))) by Tautology.from(lastStep, relationRangeDef of (t -> b)) + + have(thesis) by RightAnd(left, right) + } + + // We can now reconstruct the definition of relationBetween(r, relationDomain(r), relationRange(r)) + have((t === pair(a, b)) |- (t === pair(a, b))) by Hypothesis + val toCut = have((in(t, r), (t === pair(a, b))) |- (t === pair(a, b)) /\ in(a, relationDomain(r)) /\ in(b, relationRange(r))) by RightAnd(lastStep, memberships) + + have((t === pair(a, b)) /\ in(a, x) /\ in(b, y) |- (t === pair(a, b))) by Tautology + have((in(t, r), (t === pair(a, b)) /\ in(a, x) /\ in(b, y)) |- (t === pair(a, b)) /\ in(a, relationDomain(r)) /\ in(b, relationRange(r))) by Cut(lastStep, toCut) + + // Re-add the existential quantifiers + thenHave((in(t, r), (t === pair(a, b)) /\ in(a, x) /\ in(b, y)) |- ∃(b, (t === pair(a, b)) /\ in(a, relationDomain(r)) /\ in(b, relationRange(r)))) by RightExists + thenHave((in(t, r), (t === pair(a, b)) /\ in(a, x) /\ in(b, y)) |- ∃(a, ∃(b, (t === pair(a, b)) /\ in(a, relationDomain(r)) /\ in(b, relationRange(r))))) by RightExists + thenHave((in(t, r), ∃(b, (t === pair(a, b)) /\ in(a, x) /\ in(b, y))) |- ∃(a, ∃(b, (t === pair(a, b)) /\ in(a, relationDomain(r)) /\ in(b, relationRange(r))))) by LeftExists + thenHave((in(t, r), ∃(a, ∃(b, (t === pair(a, b)) /\ in(a, x) /\ in(b, y)))) |- ∃(a, ∃(b, (t === pair(a, b)) /\ in(a, relationDomain(r)) /\ in(b, relationRange(r))))) by LeftExists + + // Cut and rewrap the definition + have((in(t, r), relationBetween(r, x, y)) |- ∃(a, ∃(b, (t === pair(a, b)) /\ in(a, relationDomain(r)) /\ in(b, relationRange(r))))) by Cut( + relationDef, + lastStep + ) + have((in(t, r), relationBetween(r, x, y)) |- in(t, cartesianProduct(relationDomain(r), relationRange(r)))) by Tautology.from( + lastStep, + elemOfCartesianProduct of (x -> relationDomain(r), y -> relationRange(r)) + ) + thenHave(relationBetween(r, x, y) |- in(t, r) ==> in(t, cartesianProduct(relationDomain(r), relationRange(r)))) by Restate + thenHave(relationBetween(r, x, y) |- ∀(t, in(t, r) ==> in(t, cartesianProduct(relationDomain(r), relationRange(r))))) by RightForall + have(relationBetween(r, x, y) |- subset(r, cartesianProduct(relationDomain(r), relationRange(r)))) by Tautology.from( + lastStep, + subset.definition of (x -> r, y -> cartesianProduct(relationDomain(r), relationRange(r))) + ) + have(relationBetween(r, x, y) |- relationBetween(r, relationDomain(r), relationRange(r))) by Tautology.from( + lastStep, + relationBetween.definition of (a -> relationDomain(r), b -> relationRange(r)) + ) + + // Add the existential quantifier to finish the proofs + thenHave(∃(y, relationBetween(r, x, y)) |- relationBetween(r, relationDomain(r), relationRange(r))) by LeftExists + thenHave(∃(x, ∃(y, relationBetween(r, x, y))) |- relationBetween(r, relationDomain(r), relationRange(r))) by LeftExists + + have(thesis) by Tautology.from(lastStep, relation.definition) + } + /** * (Binary) Relation Field --- The union of the domain and range of a * relation, or the set of all elements related by `r`. @@ -1984,6 +2070,45 @@ object SetTheory extends lisa.Main { */ val functionalOver = DEF(f, x) --> functional(f) /\ (relationDomain(f) === x) + /** + * Lemma --- If `f` is a function, then `t ∈ f` implies `t = (x, y)` such that `x ∈ relationDomain(f)`. + */ + val functionalMembership = Lemma( + functional(f) |- ∀(t, in(t, f) ==> ∃(x, ∃(y, in(x, relationDomain(f)) /\ (t === pair(x, y))))) + ) { + assume(functional(f)) + + have((functional(f), in(t, f)) |- ∃(x, ∃(y, in(x, relationDomain(f)) /\ (t === pair(x, y))))) subproof { + val isRelation = have(relation(f)) by Tautology.from(functional.definition) + + // Use the definitions + have(relationBetween(f, relationDomain(f), relationRange(f)) |- ∀(x, in(x, f) ==> in(x, cartesianProduct(relationDomain(f), relationRange(f))))) by Tautology.from( + relationBetween.definition of (r -> f, a -> relationDomain(f), b -> relationRange(f)), + subset.definition of (x -> f, y -> cartesianProduct(relationDomain(f), relationRange(f))) + ) + thenHave(relationBetween(f, relationDomain(f), relationRange(f)) |- in(t, f) ==> in(t, cartesianProduct(relationDomain(f), relationRange(f)))) by InstantiateForall(t) + thenHave((relationBetween(f, relationDomain(f), relationRange(f)), in(t, f)) |- in(t, cartesianProduct(relationDomain(f), relationRange(f)))) by Restate + + val almostThere = have((relationBetween(f, relationDomain(f), relationRange(f)), in(t, f)) |- ∃(x, ∃(y, (t === pair(x, y)) /\ in(x, relationDomain(f)) /\ in(y, relationRange(f))))) by Tautology.from( + lastStep, + elemOfCartesianProduct of (x -> relationDomain(f), y -> relationRange(f)) + ) + + // Remove the extraneous term in the conjunction + have((t === pair(x, y)) /\ in(x, relationDomain(f)) /\ in(y, relationRange(f)) |- in(x, relationDomain(f)) /\ (t === pair(x, y))) by Tautology + thenHave((t === pair(x, y)) /\ in(x, relationDomain(f)) /\ in(y, relationRange(f)) |- ∃(y, in(x, relationDomain(f)) /\ (t === pair(x, y)))) by RightExists + thenHave((t === pair(x, y)) /\ in(x, relationDomain(f)) /\ in(y, relationRange(f)) |- ∃(x, ∃(y, in(x, relationDomain(f)) /\ (t === pair(x, y))))) by RightExists + thenHave(∃(y, (t === pair(x, y)) /\ in(x, relationDomain(f)) /\ in(y, relationRange(f))) |- ∃(x, ∃(y, in(x, relationDomain(f)) /\ (t === pair(x, y))))) by LeftExists + thenHave(∃(x, ∃(y, (t === pair(x, y)) /\ in(x, relationDomain(f)) /\ in(y, relationRange(f)))) |- ∃(x, ∃(y, in(x, relationDomain(f)) /\ (t === pair(x, y))))) by LeftExists + + have((relationBetween(f, relationDomain(f), relationRange(f)), in(t, f)) |- ∃(x, ∃(y, in(x, relationDomain(f)) /\ (t === pair(x, y))))) by Cut(almostThere, lastStep) + have((relation(f), in(t, f)) |- ∃(x, ∃(y, in(x, relationDomain(f)) /\ (t === pair(x, y))))) by Cut(relationImpliesRelationBetweenDomainAndRange of (r -> f), lastStep) + have(in(t, f) |- ∃(x, ∃(y, in(x, relationDomain(f)) /\ (t === pair(x, y))))) by Cut(isRelation, lastStep) + } + thenHave(in(t, f) ==> ∃(x, ∃(y, in(x, relationDomain(f)) /\ (t === pair(x, y))))) by Restate + thenHave(thesis) by RightForall + } + val setOfFunctionsUniqueness = Theorem( ∃!(z, ∀(t, in(t, z) <=> (in(t, powerSet(cartesianProduct(x, y))) /\ functionalOver(t, x)))) ) { @@ -2344,15 +2469,33 @@ object SetTheory extends lisa.Main { val restrictedFunctionCancellation = Theorem( functional(f) |- restrictedFunction(f, relationDomain(f)) === f ) { - /*val g = restrictedFunction(f, relationDomain(f)) + val g = restrictedFunction(f, relationDomain(f)) + + assume(functional(f)) have(∀(t, in(t, relationDomain(f)) <=> ∃(a, in(pair(t, a), f)))) by Definition(relationDomain, relationDomainUniqueness)(f) thenHave(in(y, relationDomain(f)) <=> ∃(a, in(pair(y, a), f))) by InstantiateForall(y) have(∀(t, in(t, g) <=> (in(t, f) /\ ∃(y, ∃(z, in(y, relationDomain(f)) /\ (t === pair(y, z))))))) by Definition( restrictedFunction, restrictedFunctionUniqueness - )(f, relationDomain(f))*/ - sorry + )(f, relationDomain(f)) + val equiv = thenHave(in(t, g) <=> (in(t, f) /\ ∃(y, ∃(z, in(y, relationDomain(f)) /\ (t === pair(y, z)))))) by InstantiateForall(t) + + // Prove that the second part of the conjunction is extraneous + val hypo = have(in(t, f) |- in(t, f)) by Hypothesis + have(in(t, f) |- ∃(y, ∃(z, in(y, relationDomain(f)) /\ (t === pair(y, z))))) by InstantiateForall(t)(functionalMembership) + have(in(t, f) |- in(t, f) /\ ∃(y, ∃(z, in(y, relationDomain(f)) /\ (t === pair(y, z))))) by RightAnd(hypo, lastStep) + val forward = thenHave(in(t, f) ==> (in(t, f) /\ ∃(y, ∃(z, in(y, relationDomain(f)) /\ (t === pair(y, z)))))) by Restate + + val backward = have(in(t, f) /\ ∃(y, ∃(z, in(y, relationDomain(f)) /\ (t === pair(y, z)))) ==> in(t, f)) by Tautology + + have(in(t, f) <=> (in(t, f) /\ ∃(y, ∃(z, in(y, relationDomain(f)) /\ (t === pair(y, z)))))) by RightIff(forward, backward) + + // Conclude by extensionnality + have(in(t, g) <=> in(t, f)) by Tautology.from(equiv, lastStep) + thenHave(∀(t, in(t, g) <=> in(t, f))) by RightForall + + have(g === f) by Tautology.from(extensionalityAxiom of (x -> g, y -> f), lastStep) } // TODO: any subset of a functional is functional From 80e14496e1447dfdcdf47ebcd41f69d1fe518abe Mon Sep 17 00:00:00 2001 From: dhalilov Date: Mon, 15 May 2023 11:22:30 +0200 Subject: [PATCH 38/68] Fix groupIsSubgroupOfItself proof --- src/main/scala/lisa/mathematics/GroupTheory.scala | 6 +++++- 1 file changed, 5 insertions(+), 1 deletion(-) diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala index 13ac791b..f258631e 100644 --- a/src/main/scala/lisa/mathematics/GroupTheory.scala +++ b/src/main/scala/lisa/mathematics/GroupTheory.scala @@ -486,7 +486,11 @@ object GroupTheory extends lisa.Main { )(condition1) val eq3 = have(group(G, *) |- restrictedFunction(*, cartesianProduct(G, G)) === *) subproof { - val eq1 = have(restrictedFunction(*, relationDomain(*)) === *) by Restate.from(restrictedFunctionCancellation of (f -> *)) + assume(group(G, *)) + val eq1 = have(restrictedFunction(*, relationDomain(*)) === *) by Cut( + groupOperationIsFunctional, + restrictedFunctionCancellation of (f -> *) + ) thenHave((relationDomain(*) === cartesianProduct(G, G)) |- restrictedFunction(*, cartesianProduct(G, G)) === *) by RightSubstEq( List((relationDomain(*), cartesianProduct(G, G))), lambda(z, restrictedFunction(*, z) === *) ) From c37431b193f761f35ecd2c27ccee756eaa951445 Mon Sep 17 00:00:00 2001 From: dhalilov Date: Mon, 15 May 2023 11:25:11 +0200 Subject: [PATCH 39/68] scala{fix, fmt} --- .../lisa/mathematics/FirstOrderLogic.scala | 11 +-- .../scala/lisa/mathematics/GroupTheory.scala | 70 ++++++++++++------- .../scala/lisa/mathematics/SetTheory.scala | 57 +++++++++------ 3 files changed, 86 insertions(+), 52 deletions(-) diff --git a/src/main/scala/lisa/mathematics/FirstOrderLogic.scala b/src/main/scala/lisa/mathematics/FirstOrderLogic.scala index 07c7fff0..6159f98c 100644 --- a/src/main/scala/lisa/mathematics/FirstOrderLogic.scala +++ b/src/main/scala/lisa/mathematics/FirstOrderLogic.scala @@ -96,7 +96,8 @@ object FirstOrderLogic extends lisa.Main { have((x === y) <=> P(y) |- (x === y) <=> P(y)) by Hypothesis thenHave(∀(y, (x === y) <=> P(y)) |- (x === y) <=> P(y)) by LeftForall val substitution = thenHave((∀(y, (x === y) <=> P(y)), P(y) <=> Q(y)) |- (x === y) <=> Q(y)) by RightSubstIff( - List((P(y), Q(y))), lambda(p, (x === y) <=> p) + List((P(y), Q(y))), + lambda(p, (x === y) <=> p) ) have(∀(y, P(y) <=> Q(y)) |- ∀(y, P(y) <=> Q(y))) by Hypothesis @@ -179,9 +180,11 @@ object FirstOrderLogic extends lisa.Main { exists(x, P(x) /\ p()) <=> (exists(x, P(x)) /\ p()) ) { val forward = have(exists(x, P(x) /\ p()) ==> (exists(x, P(x)) /\ p())) subproof { - have(exists(x, P(x) /\ p()) |- exists(x, P(x)) /\ exists(x, p())) by Restate.from(existentialConjunctionDistribution of ( - Q -> lambda(x, p()) - )) + have(exists(x, P(x) /\ p()) |- exists(x, P(x)) /\ exists(x, p())) by Restate.from( + existentialConjunctionDistribution of ( + Q -> lambda(x, p()) + ) + ) val substitution = thenHave( (exists(x, P(x) /\ p()), (exists(x, P(x)) /\ exists(x, p())) <=> (exists(x, P(x)) /\ p())) |- exists(x, P(x)) /\ p() ) by RightSubstIff(List((exists(x, P(x)) /\ exists(x, p()), exists(x, P(x)) /\ p())), lambda(p, p())) diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala index f258631e..3d60740b 100644 --- a/src/main/scala/lisa/mathematics/GroupTheory.scala +++ b/src/main/scala/lisa/mathematics/GroupTheory.scala @@ -1,8 +1,12 @@ package lisa.mathematics -import lisa.automation.grouptheory.GroupTheoryTactics.{Cases, Definition, ExistenceAndUniqueness} +import lisa.automation.grouptheory.GroupTheoryTactics.Cases +import lisa.automation.grouptheory.GroupTheoryTactics.Definition +import lisa.automation.grouptheory.GroupTheoryTactics.ExistenceAndUniqueness import lisa.automation.kernel.OLPropositionalSolver.Tautology -import lisa.mathematics.FirstOrderLogic.{equalityTransitivity, existsOneImpliesExists, substitutionInUniquenessQuantifier} +import lisa.mathematics.FirstOrderLogic.equalityTransitivity +import lisa.mathematics.FirstOrderLogic.existsOneImpliesExists +import lisa.mathematics.FirstOrderLogic.substitutionInUniquenessQuantifier import lisa.mathematics.SetTheory.* /** @@ -79,7 +83,7 @@ object GroupTheory extends lisa.Main { } val substitution = have((∃!(u, f), ∀(u, f <=> completeDef)) |- ∃!(u, completeDef)) by Restate.from( - substitutionInUniquenessQuantifier of(P -> lambda(u, f), Q -> lambda(u, completeDef)) + substitutionInUniquenessQuantifier of (P -> lambda(u, f), Q -> lambda(u, completeDef)) ) val implication = have((prem, ∃!(u, f)) |- ∃!(u, completeDef)) by Cut(equiv, substitution) @@ -197,7 +201,7 @@ object GroupTheory extends lisa.Main { assume(group(G, *)) have(binaryFunction(G, *)) by Tautology.from(group.definition) have(functionFrom(*, cartesianProduct(G, G), G)) by Tautology.from(lastStep, binaryFunction.definition) - have(relationDomain(*) === cartesianProduct(G, G)) by Tautology.from(lastStep, functionFromImpliesDomainEq of(f -> *, x -> cartesianProduct(G, G), y -> G)) + have(relationDomain(*) === cartesianProduct(G, G)) by Tautology.from(lastStep, functionFromImpliesDomainEq of (f -> *, x -> cartesianProduct(G, G), y -> G)) } /** @@ -214,10 +218,12 @@ object GroupTheory extends lisa.Main { have(x ∈ G /\ y ∈ G) by Tautology have(pair(x, y) ∈ cartesianProduct(G, G)) by Tautology.from( - lastStep, pairInCartesianProduct of (a -> x, b -> y, x -> G, y -> G) + lastStep, + pairInCartesianProduct of (a -> x, b -> y, x -> G, y -> G) ) thenHave((relationDomain(*) === cartesianProduct(G, G)) |- pair(x, y) ∈ relationDomain(*)) by RightSubstEq( - List((relationDomain(*), cartesianProduct(G, G))), lambda(z, pair(x, y) ∈ z) + List((relationDomain(*), cartesianProduct(G, G))), + lambda(z, pair(x, y) ∈ z) ) have(thesis) by Cut(groupOperationDomain, lastStep) @@ -273,7 +279,7 @@ object GroupTheory extends lisa.Main { /** * Theorem --- The identity element belongs to the group. - * + * * This might seem like a silly theorem, but it is useful in proving that groups are non-empty by definition. */ val identityInGroup = Theorem( @@ -282,13 +288,14 @@ object GroupTheory extends lisa.Main { assume(group(G, *)) have(isNeutral(identity(G, *), G, *)) by Definition(identity, identityUniqueness)(G, *) have(thesis) by Tautology.from( - lastStep, isNeutral.definition of (e -> identity(G, *)) + lastStep, + isNeutral.definition of (e -> identity(G, *)) ) } /** * Theorem --- A group is non-empty. - * + * * Direct corollary of [[identityInGroup]]. */ val groupNonEmpty = Theorem( @@ -366,15 +373,19 @@ object GroupTheory extends lisa.Main { // 4. z = y(xz) val fourthEq = have((group(G, *), x ∈ G, isInverse(y, x, G, *), isInverse(z, x, G, *)) |- z === op(y, *, op(x, *, z))) by Tautology.from( - firstEq, secondEq, equalityTransitivity of (x -> z, y -> op(op(y, *, x), *, z), z -> op(y, *, op(x, *, z))) + firstEq, + secondEq, + equalityTransitivity of (x -> z, y -> op(op(y, *, x), *, z), z -> op(y, *, op(x, *, z))) ) // 5. z = y have((group(G, *), x ∈ G, isInverse(y, x, G, *), isInverse(z, x, G, *)) |- z === y) by Tautology.from( - thirdEq, fourthEq, equalityTransitivity of (x -> z, y -> op(y, *, op(x, *, z)), z -> y) + thirdEq, + fourthEq, + equalityTransitivity of (x -> z, y -> op(y, *, op(x, *, z)), z -> y) ) } - + have(thesis) by ExistenceAndUniqueness(isInverse(y, x, G, *))(existence, uniqueness) } @@ -387,7 +398,7 @@ object GroupTheory extends lisa.Main { * Theorem --- `y` is the inverse of `x` iff `x` is the inverse of `y` */ val inverseSymmetry = Theorem( - group(G, *) |- ∀(x, x ∈ G ==> (isInverse(y, x, G, *) ==> isInverse(x, y, G, *))) + group(G, *) |- ∀(x, x ∈ G ==> (isInverse(y, x, G, *) ==> isInverse(x, y, G, *))) ) { have((group(G, *), x ∈ G, isInverse(y, x, G, *)) |- isInverse(x, y, G, *)) subproof { assume(group(G, *)) @@ -430,7 +441,8 @@ object GroupTheory extends lisa.Main { } val membership = have((group(G, *), x ∈ G) |- in(inverse(x, G, *), G)) by Tautology.from( - inverseIsInverse, isInverse.definition of (y -> inverse(x, G, *)) + inverseIsInverse, + isInverse.definition of (y -> inverse(x, G, *)) ) // Use the symmetry of the inverse relation to prove that x is an inverse of inverse(x) @@ -482,9 +494,10 @@ object GroupTheory extends lisa.Main { // 3. restrictedFunction(*, cartesianProduct(G, G)) === * (derived from 1. and 2.) val substitution = have((group(G, *), restrictedFunction(*, cartesianProduct(G, G)) === *) |- group(G, restrictedFunction(*, cartesianProduct(G, G)))) by RightSubstEq( - List((restrictedFunction(*, cartesianProduct(G, G)), *)), lambda(z, group(G, z)) + List((restrictedFunction(*, cartesianProduct(G, G)), *)), + lambda(z, group(G, z)) )(condition1) - + val eq3 = have(group(G, *) |- restrictedFunction(*, cartesianProduct(G, G)) === *) subproof { assume(group(G, *)) val eq1 = have(restrictedFunction(*, relationDomain(*)) === *) by Cut( @@ -492,7 +505,8 @@ object GroupTheory extends lisa.Main { restrictedFunctionCancellation of (f -> *) ) thenHave((relationDomain(*) === cartesianProduct(G, G)) |- restrictedFunction(*, cartesianProduct(G, G)) === *) by RightSubstEq( - List((relationDomain(*), cartesianProduct(G, G))), lambda(z, restrictedFunction(*, z) === *) + List((relationDomain(*), cartesianProduct(G, G))), + lambda(z, restrictedFunction(*, z) === *) ) have(thesis) by Cut(groupOperationDomain, lastStep) @@ -525,7 +539,7 @@ object GroupTheory extends lisa.Main { val a, b = variable have(subset(H, G)) by Tautology.from(subgroup.definition) - have(∀(x, x ∈ H ==> x ∈ G)) by Tautology.from(lastStep, subset.definition of(x -> H, y -> G)) + have(∀(x, x ∈ H ==> x ∈ G)) by Tautology.from(lastStep, subset.definition of (x -> H, y -> G)) val subsetDef = thenHave(x ∈ H ==> x ∈ G) by InstantiateForall(x) val left = have(x ∈ G) by Tautology.from(subsetDef) @@ -549,12 +563,15 @@ object GroupTheory extends lisa.Main { // We characterize op(x, *, y), and show that op(x, ★, y) satisfies all requirements have( - ∀(z, (z === app(*, pair(x, y))) <=> (((functional(*) /\ in(pair(x, y), relationDomain(*))) ==> in(pair(pair(x, y), z), *)) /\ - ((!functional(*) \/ !in(pair(x, y), relationDomain(*))) ==> (z === ∅)))) + ∀( + z, + (z === app(*, pair(x, y))) <=> (((functional(*) /\ in(pair(x, y), relationDomain(*))) ==> in(pair(pair(x, y), z), *)) /\ + ((!functional(*) \/ !in(pair(x, y), relationDomain(*))) ==> (z === ∅))) + ) ) by Tautology.from(app.definition of (f -> *, x -> pair(x, y)), functionApplicationUniqueness) val characterization = thenHave( (r === app(*, pair(x, y))) <=> (((functional(*) /\ in(pair(x, y), relationDomain(*))) ==> in(pair(pair(x, y), r), *)) /\ - ((!functional(*) \/ !in(pair(x, y), relationDomain(*))) ==> (r === ∅))) + ((!functional(*) \/ !in(pair(x, y), relationDomain(*))) ==> (r === ∅))) ) by InstantiateForall(r) // Prove that the premises of the first implication hold @@ -565,7 +582,8 @@ object GroupTheory extends lisa.Main { } val premises = have((subgroup(H, G, *), x ∈ H, y ∈ H) |- functional(*) /\ pair(x, y) ∈ relationDomain(*)) by RightAnd( - leftPremise, subgroupPairInParentOperationDomain + leftPremise, + subgroupPairInParentOperationDomain ) // We show that op(x, ★, y) satisfies the conclusion of the implication @@ -580,7 +598,8 @@ object GroupTheory extends lisa.Main { val reduction3 = have((subgroup(H, G, *), pair(x, y) ∈ relationDomain(★)) |- pair(pair(x, y), r) ∈ ★) by Tautology.from(reduction2, appDef) have((subgroup(H, G, *), x ∈ H, y ∈ H) |- pair(x, y) ∈ relationDomain(★)) by Cut( - reduction1, groupPairInOperationDomain of (G -> H, * -> ★) + reduction1, + groupPairInOperationDomain of (G -> H, * -> ★) ) val reducedDef = have((subgroup(H, G, *), x ∈ H, y ∈ H) |- pair(pair(x, y), r) ∈ ★) by Cut(lastStep, reduction3) @@ -589,7 +608,8 @@ object GroupTheory extends lisa.Main { thenHave(u ∈ ★ ==> u ∈ *) by Tautology val satisfaction = have((subgroup(H, G, *), x ∈ H, y ∈ H) |- pair(pair(x, y), r) ∈ *) by Tautology.from( - lastStep of (u -> pair(pair(x, y), r)), reducedDef + lastStep of (u -> pair(pair(x, y), r)), + reducedDef ) // Reconstruct the whole definition @@ -606,7 +626,7 @@ object GroupTheory extends lisa.Main { have( ((functional(*) /\ pair(x, y) ∈ relationDomain(*)) ==> pair(pair(x, y), r) ∈ *) /\ - ((!functional(*) \/ !(pair(x, y) ∈ relationDomain(*))) ==> (r === emptySet())) + ((!functional(*) \/ !(pair(x, y) ∈ relationDomain(*))) ==> (r === emptySet())) ) by RightAnd(pos, neg) have(thesis) by Tautology.from(lastStep, characterization) diff --git a/src/main/scala/lisa/mathematics/SetTheory.scala b/src/main/scala/lisa/mathematics/SetTheory.scala index 550ee637..a046ede1 100644 --- a/src/main/scala/lisa/mathematics/SetTheory.scala +++ b/src/main/scala/lisa/mathematics/SetTheory.scala @@ -856,7 +856,7 @@ object SetTheory extends lisa.Main { */ val setIntersection = DEF(x, y) --> The(z, ∀(t, in(t, z) <=> (in(t, x) /\ in(t, y))))(setIntersectionUniqueness) - extension(x: Term) { + extension (x: Term) { infix def ∩(y: Term) = setIntersection(x, y) } @@ -2089,10 +2089,11 @@ object SetTheory extends lisa.Main { thenHave(relationBetween(f, relationDomain(f), relationRange(f)) |- in(t, f) ==> in(t, cartesianProduct(relationDomain(f), relationRange(f)))) by InstantiateForall(t) thenHave((relationBetween(f, relationDomain(f), relationRange(f)), in(t, f)) |- in(t, cartesianProduct(relationDomain(f), relationRange(f)))) by Restate - val almostThere = have((relationBetween(f, relationDomain(f), relationRange(f)), in(t, f)) |- ∃(x, ∃(y, (t === pair(x, y)) /\ in(x, relationDomain(f)) /\ in(y, relationRange(f))))) by Tautology.from( - lastStep, - elemOfCartesianProduct of (x -> relationDomain(f), y -> relationRange(f)) - ) + val almostThere = + have((relationBetween(f, relationDomain(f), relationRange(f)), in(t, f)) |- ∃(x, ∃(y, (t === pair(x, y)) /\ in(x, relationDomain(f)) /\ in(y, relationRange(f))))) by Tautology.from( + lastStep, + elemOfCartesianProduct of (x -> relationDomain(f), y -> relationRange(f)) + ) // Remove the extraneous term in the conjunction have((t === pair(x, y)) /\ in(x, relationDomain(f)) /\ in(y, relationRange(f)) |- in(x, relationDomain(f)) /\ (t === pair(x, y))) by Tautology @@ -2348,7 +2349,8 @@ object SetTheory extends lisa.Main { val g = restrictedFunction(f, x) have(∀(t, in(t, g) <=> (in(t, f) /\ ∃(y, ∃(z, in(y, x) /\ (t === pair(y, z))))))) by Definition( - restrictedFunction, restrictedFunctionUniqueness + restrictedFunction, + restrictedFunctionUniqueness )(f, x) val pairMembership = thenHave( in(pair(t, a), g) <=> (in(pair(t, a), f) /\ ∃(y, ∃(z, in(y, x) /\ (pair(t, a) === pair(y, z))))) @@ -2359,39 +2361,41 @@ object SetTheory extends lisa.Main { thenHave(∀(z, (in(y, x) /\ (pair(t, a) === pair(y, z))) <=> (in(y, x) /\ (t === y) /\ (a === z)))) by RightForall val existentialEquiv1 = have(∃(z, in(y, x) /\ (pair(t, a) === pair(y, z))) <=> ∃(z, in(y, x) /\ (t === y) /\ (a === z))) by Cut( - lastStep, existentialEquivalenceDistribution of ( + lastStep, + existentialEquivalenceDistribution of ( P -> lambda(z, in(y, x) /\ (pair(t, a) === pair(y, z))), Q -> lambda(z, in(y, x) /\ (t === y) /\ (a === z)) ) ) - + have(∃(z, in(y, x) /\ (t === y) /\ (a === z)) <=> (in(y, x) /\ (t === y))) by Restate.from( equalityInExistentialQuantifier of ( P -> lambda(z, in(y, x) /\ (t === y)), - y -> a + y -> a ) ) - + have(∃(z, in(y, x) /\ (pair(t, a) === pair(y, z))) <=> (in(y, x) /\ (t === y))) by Tautology.from(existentialEquiv1, lastStep) thenHave(∀(y, ∃(z, in(y, x) /\ (pair(t, a) === pair(y, z))) <=> (in(y, x) /\ (t === y)))) by RightForall - + val existentialEquiv2 = have(∃(y, ∃(z, in(y, x) /\ (pair(t, a) === pair(y, z)))) <=> ∃(y, in(y, x) /\ (t === y))) by Cut( - lastStep, existentialEquivalenceDistribution of( + lastStep, + existentialEquivalenceDistribution of ( P -> lambda(y, ∃(z, in(y, x) /\ (pair(t, a) === pair(y, z)))), Q -> lambda(y, in(y, x) /\ (t === y)) ) ) - + have(∃(y, in(y, x) /\ (t === y)) <=> in(t, x)) by Restate.from( - equalityInExistentialQuantifier of( + equalityInExistentialQuantifier of ( P -> lambda(y, in(y, x)), y -> t ) ) - + have(∃(y, ∃(z, in(y, x) /\ (pair(t, a) === pair(y, z)))) <=> in(t, x)) by Tautology.from(existentialEquiv2, lastStep) thenHave((in(pair(t, a), f) /\ ∃(y, ∃(z, in(y, x) /\ (pair(t, a) === pair(y, z))))) <=> (in(pair(t, a), f) /\ in(t, x))) by Tautology - + have(thesis) by Tautology.from(lastStep, pairMembership) } @@ -2409,13 +2413,15 @@ object SetTheory extends lisa.Main { val domCharacterization = have(∀(t, in(t, dom) <=> (∃(a, in(pair(t, a), f)) /\ in(t, x)))) subproof { // Use the definition of the intersection have(∀(t, in(t, dom) <=> (in(t, x) /\ in(t, relationDomain(f))))) by Definition( - setIntersection, setIntersectionUniqueness + setIntersection, + setIntersectionUniqueness )(x, relationDomain(f)) val intersectionDef = thenHave(in(t, dom) <=> (in(t, x) /\ in(t, relationDomain(f)))) by InstantiateForall(t) // Use the definition of the relation domain have(∀(t, in(t, relationDomain(f)) <=> ∃(a, in(pair(t, a), f)))) by Definition( - relationDomain, relationDomainUniqueness + relationDomain, + relationDomainUniqueness )(f) thenHave(in(t, relationDomain(f)) <=> ∃(a, in(pair(t, a), f))) by InstantiateForall(t) @@ -2426,14 +2432,16 @@ object SetTheory extends lisa.Main { // Characterize the domain of g have(∀(D, (relationDomain(g) === D) <=> ∀(t, in(t, D) <=> ∃(a, in(pair(t, a), g))))) by Tautology.from( - relationDomain.definition of (r -> g), relationDomainUniqueness + relationDomain.definition of (r -> g), + relationDomainUniqueness ) val characterization = thenHave((relationDomain(g) === dom) <=> ∀(t, in(t, dom) <=> ∃(a, in(pair(t, a), g)))) by InstantiateForall(dom) // Use the membership of a pair in the restricted function to derive a simpler characterization have(∀(a, in(pair(t, a), g) <=> (in(pair(t, a), f) /\ in(t, x)))) by RightForall(restrictedFunctionPairMembership) have(∃(a, in(pair(t, a), g)) <=> ∃(a, in(pair(t, a), f) /\ in(t, x))) by Tautology.from( - lastStep, existentialEquivalenceDistribution of ( + lastStep, + existentialEquivalenceDistribution of ( P -> lambda(a, in(pair(t, a), g)), Q -> lambda(a, in(pair(t, a), f) /\ in(t, x)) ) @@ -2442,7 +2450,8 @@ object SetTheory extends lisa.Main { // Extract in(t, x) from the existential quantifier val p = formulaVariable // local shadowing to correctly use the theorem have(∃(a, in(pair(t, a), g)) <=> ∃(a, in(pair(t, a), f)) /\ in(t, x)) by Tautology.from( - lastStep, existentialConjunctionWithClosedFormula of ( + lastStep, + existentialConjunctionWithClosedFormula of ( P -> lambda(a, in(pair(t, a), f)), p -> lambda(Seq(), in(t, x)) ) @@ -2452,7 +2461,8 @@ object SetTheory extends lisa.Main { thenHave(∀(t, (in(t, dom) <=> ∃(a, in(pair(t, a), g))) <=> (in(t, dom) <=> ∃(a, in(pair(t, a), f)) /\ in(t, x)))) by RightForall have(∀(t, in(t, dom) <=> ∃(a, in(pair(t, a), g))) <=> ∀(t, in(t, dom) <=> ∃(a, in(pair(t, a), f)) /\ in(t, x))) by Cut( - lastStep, universalEquivalenceDistribution of ( + lastStep, + universalEquivalenceDistribution of ( P -> lambda(t, in(t, dom) <=> ∃(a, in(pair(t, a), g))), Q -> lambda(t, in(t, dom) <=> ∃(a, in(pair(t, a), f)) /\ in(t, x)) ) @@ -2477,7 +2487,8 @@ object SetTheory extends lisa.Main { thenHave(in(y, relationDomain(f)) <=> ∃(a, in(pair(y, a), f))) by InstantiateForall(y) have(∀(t, in(t, g) <=> (in(t, f) /\ ∃(y, ∃(z, in(y, relationDomain(f)) /\ (t === pair(y, z))))))) by Definition( - restrictedFunction, restrictedFunctionUniqueness + restrictedFunction, + restrictedFunctionUniqueness )(f, relationDomain(f)) val equiv = thenHave(in(t, g) <=> (in(t, f) /\ ∃(y, ∃(z, in(y, relationDomain(f)) /\ (t === pair(y, z)))))) by InstantiateForall(t) From 2ffe0d760c6d7bd3da017ddb1c25d6b46937ac02 Mon Sep 17 00:00:00 2001 From: dhalilov Date: Mon, 15 May 2023 17:26:18 +0200 Subject: [PATCH 40/68] Create associativity lemma --- .../scala/lisa/mathematics/GroupTheory.scala | 37 +++++++++++-------- 1 file changed, 22 insertions(+), 15 deletions(-) diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala index 3d60740b..c5c39fe3 100644 --- a/src/main/scala/lisa/mathematics/GroupTheory.scala +++ b/src/main/scala/lisa/mathematics/GroupTheory.scala @@ -229,6 +229,26 @@ object GroupTheory extends lisa.Main { have(thesis) by Cut(groupOperationDomain, lastStep) } + /** + * Simple reformulation of the associativity. + */ + val associativityApplication = Lemma( + (group(G, *), x ∈ G, y ∈ G, z ∈ G) |- op(op(x, *, y), *, z) === op(x, *, op(y, *, z)) + ) { + assume(group(G, *)) + + have(∀(x, x ∈ G ==> ∀(y, y ∈ G ==> ∀(z, z ∈ G ==> (op(op(x, *, y), *, z) === op(x, *, op(y, *, z))))))) by Tautology.from( + group.definition, + associativity.definition + ) + thenHave(x ∈ G ==> ∀(y, y ∈ G ==> ∀(z, z ∈ G ==> (op(op(x, *, y), *, z) === op(x, *, op(y, *, z)))))) by InstantiateForall(x) + thenHave(x ∈ G |- ∀(y, y ∈ G ==> ∀(z, z ∈ G ==> (op(op(x, *, y), *, z) === op(x, *, op(y, *, z)))))) by Restate + thenHave(x ∈ G |- y ∈ G ==> ∀(z, z ∈ G ==> (op(op(x, *, y), *, z) === op(x, *, op(y, *, z))))) by InstantiateForall(y) + thenHave((x ∈ G, y ∈ G) |- ∀(z, z ∈ G ==> (op(op(x, *, y), *, z) === op(x, *, op(y, *, z))))) by Restate + thenHave((x ∈ G, y ∈ G) |- z ∈ G ==> (op(op(x, *, y), *, z) === op(x, *, op(y, *, z)))) by InstantiateForall(z) + thenHave((x ∈ G, y ∈ G, z ∈ G) |- (op(op(x, *, y), *, z) === op(x, *, op(y, *, z)))) by Restate + } + /** * Identity uniqueness --- In a group (G, *), an identity element is unique, i.e. if both `e * x = x * e = x` and * `f * x = x * f = x` for all `x`, then `e = f`. @@ -338,21 +358,8 @@ object GroupTheory extends lisa.Main { val firstEq = have((group(G, *), x ∈ G, isInverse(y, x, G, *), isInverse(z, x, G, *)) |- op(op(y, *, x), *, z) === z) by Cut(inverseMembership of (y -> z), leftNeutrality) // 2. (yx)z = y(xz) - val permuteParentheses = have((group(G, *), x ∈ G, y ∈ G, z ∈ G) |- (op(op(y, *, x), *, z) === op(y, *, op(x, *, z)))) subproof { - assume(group(G, *)) - assume(x ∈ G) - assume(y ∈ G) - assume(z ∈ G) - - have(∀(u, u ∈ G ==> ∀(v, v ∈ G ==> ∀(w, w ∈ G ==> (op(op(u, *, v), *, w) === op(u, *, op(v, *, w))))))) by Tautology.from(group.definition, associativity.definition) - thenHave(y ∈ G ==> ∀(v, v ∈ G ==> ∀(w, w ∈ G ==> (op(op(y, *, v), *, w) === op(y, *, op(v, *, w)))))) by InstantiateForall(y) - thenHave(∀(v, v ∈ G ==> ∀(w, w ∈ G ==> (op(op(y, *, v), *, w) === op(y, *, op(v, *, w)))))) by Tautology - thenHave(x ∈ G ==> ∀(w, w ∈ G ==> (op(op(y, *, x), *, w) === op(y, *, op(x, *, w))))) by InstantiateForall(x) - thenHave(∀(w, w ∈ G ==> (op(op(y, *, x), *, w) === op(y, *, op(x, *, w))))) by Tautology - thenHave(z ∈ G ==> (op(op(y, *, x), *, z) === op(y, *, op(x, *, z)))) by InstantiateForall(z) - thenHave(op(op(y, *, x), *, z) === op(y, *, op(x, *, z))) by Tautology - } - val associativityCut = have((group(G, *), x ∈ G /\ y ∈ G /\ z ∈ G) |- (op(op(y, *, x), *, z) === op(y, *, op(x, *, z)))) by Restate.from(permuteParentheses) + have((group(G, *), x ∈ G, y ∈ G, z ∈ G) |- (op(op(y, *, x), *, z) === op(y, *, op(x, *, z)))) by Restate.from(associativityApplication of (x -> y, y -> x)) + val associativityCut = thenHave((group(G, *), x ∈ G /\ y ∈ G /\ z ∈ G) |- (op(op(y, *, x), *, z) === op(y, *, op(x, *, z)))) by Restate val memberships = have((x ∈ G, isInverse(y, x, G, *), isInverse(z, x, G, *)) |- x ∈ G /\ y ∈ G /\ z ∈ G) by Tautology.from(inverseMembership of (y -> y), inverseMembership of (y -> z)) val secondEq = have((group(G, *), x ∈ G, isInverse(y, x, G, *), isInverse(z, x, G, *)) |- op(op(y, *, x), *, z) === op(y, *, op(x, *, z))) by Cut(memberships, associativityCut) From ae454add1384385501052de6e7ced527c74df2e8 Mon Sep 17 00:00:00 2001 From: dhalilov Date: Tue, 16 May 2023 16:12:42 +0200 Subject: [PATCH 41/68] Simplify proofs --- .../scala/lisa/mathematics/GroupTheory.scala | 58 +++++++++---------- 1 file changed, 27 insertions(+), 31 deletions(-) diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala index c5c39fe3..5f39117d 100644 --- a/src/main/scala/lisa/mathematics/GroupTheory.scala +++ b/src/main/scala/lisa/mathematics/GroupTheory.scala @@ -8,6 +8,7 @@ import lisa.mathematics.FirstOrderLogic.equalityTransitivity import lisa.mathematics.FirstOrderLogic.existsOneImpliesExists import lisa.mathematics.FirstOrderLogic.substitutionInUniquenessQuantifier import lisa.mathematics.SetTheory.* +import lisa.kernel.proof.SequentCalculus.Hypothesis /** * Group theory, developed following Chapter 2 of S. Lang "Undergraduate Algebra". @@ -22,6 +23,7 @@ object GroupTheory extends lisa.Main { private val * = variable // Group elements + private val a, b = variable private val x, y, z = variable private val u, v, w = variable @@ -147,7 +149,7 @@ object GroupTheory extends lisa.Main { /** * Associativity --- `*` is associative (in `G`) if `(x * y) * z = x * (y * z)` for all `x, y, z` in `G`. */ - val associativity = DEF(G, *) --> + val associativityAxiom = DEF(G, *) --> ∀(x, x ∈ G ==> ∀(y, y ∈ G ==> ∀(z, z ∈ G ==> (op(op(x, *, y), *, z) === op(x, *, op(y, *, z)))))) /** @@ -171,10 +173,10 @@ object GroupTheory extends lisa.Main { val inverseExistence = DEF(G, *) --> ∀(x, (x ∈ G) ==> ∃(y, isInverse(y, x, G, *))) /** - * Group --- A group (G, *) is a set along with a law of composition `*`, satisfying [[associativity]], [[identityExistence]] + * Group --- A group (G, *) is a set along with a law of composition `*`, satisfying [[associativityAxiom]], [[identityExistence]] * and [[inverseExistence]]. */ - val group = DEF(G, *) --> binaryFunction(G, *) /\ associativity(G, *) /\ identityExistence(G, *) /\ inverseExistence(G, *) + val group = DEF(G, *) --> binaryFunction(G, *) /\ associativityAxiom(G, *) /\ identityExistence(G, *) /\ inverseExistence(G, *) /** * Group operation is functional -- The group operation `*` is functional. @@ -184,10 +186,11 @@ object GroupTheory extends lisa.Main { val groupOperationIsFunctional = Lemma( group(G, *) |- functional(*) ) { - assume(group(G, *)) - have(binaryFunction(G, *)) by Tautology.from(group.definition) - have(functionFrom(*, cartesianProduct(G, G), G)) by Tautology.from(lastStep, binaryFunction.definition) - have(functional(*)) by Tautology.from(lastStep, functionFromImpliesFunctional of (f -> *, x -> cartesianProduct(G, G), y -> G)) + have(thesis) by Tautology.from( + group.definition, + binaryFunction.definition, + functionFromImpliesFunctional of (f -> *, x -> cartesianProduct(G, G), y -> G) + ) } /** @@ -198,10 +201,11 @@ object GroupTheory extends lisa.Main { val groupOperationDomain = Lemma( group(G, *) |- relationDomain(*) === cartesianProduct(G, G) ) { - assume(group(G, *)) - have(binaryFunction(G, *)) by Tautology.from(group.definition) - have(functionFrom(*, cartesianProduct(G, G), G)) by Tautology.from(lastStep, binaryFunction.definition) - have(relationDomain(*) === cartesianProduct(G, G)) by Tautology.from(lastStep, functionFromImpliesDomainEq of (f -> *, x -> cartesianProduct(G, G), y -> G)) + have(thesis) by Tautology.from( + group.definition, + binaryFunction.definition, + functionFromImpliesDomainEq of (f -> *, x -> cartesianProduct(G, G), y -> G) + ) } /** @@ -214,8 +218,6 @@ object GroupTheory extends lisa.Main { assume(x ∈ G) assume(y ∈ G) - val a, b = variable - have(x ∈ G /\ y ∈ G) by Tautology have(pair(x, y) ∈ cartesianProduct(G, G)) by Tautology.from( lastStep, @@ -230,16 +232,18 @@ object GroupTheory extends lisa.Main { } /** - * Simple reformulation of the associativity. + * Lemma --- For elements `x, y, z` in a group `(G, *)`, we have `(xy)z = x(yz)`. + * + * Simpler reformulation of the [[associativityAxiom]]. */ - val associativityApplication = Lemma( + val associativity = Lemma( (group(G, *), x ∈ G, y ∈ G, z ∈ G) |- op(op(x, *, y), *, z) === op(x, *, op(y, *, z)) ) { assume(group(G, *)) have(∀(x, x ∈ G ==> ∀(y, y ∈ G ==> ∀(z, z ∈ G ==> (op(op(x, *, y), *, z) === op(x, *, op(y, *, z))))))) by Tautology.from( group.definition, - associativity.definition + associativityAxiom.definition ) thenHave(x ∈ G ==> ∀(y, y ∈ G ==> ∀(z, z ∈ G ==> (op(op(x, *, y), *, z) === op(x, *, op(y, *, z)))))) by InstantiateForall(x) thenHave(x ∈ G |- ∀(y, y ∈ G ==> ∀(z, z ∈ G ==> (op(op(x, *, y), *, z) === op(x, *, op(y, *, z)))))) by Restate @@ -252,6 +256,7 @@ object GroupTheory extends lisa.Main { /** * Identity uniqueness --- In a group (G, *), an identity element is unique, i.e. if both `e * x = x * e = x` and * `f * x = x * f = x` for all `x`, then `e = f`. + * * This justifies calling `e` the identity element. */ val identityUniqueness = Theorem( @@ -300,13 +305,12 @@ object GroupTheory extends lisa.Main { /** * Theorem --- The identity element belongs to the group. * - * This might seem like a silly theorem, but it is useful in proving that groups are non-empty by definition. + * This might seem like a silly theorem, but it is useful in [[groupNonEmpty]]. */ val identityInGroup = Theorem( group(G, *) |- identity(G, *) ∈ G ) { - assume(group(G, *)) - have(isNeutral(identity(G, *), G, *)) by Definition(identity, identityUniqueness)(G, *) + have(group(G, *) |- isNeutral(identity(G, *), G, *)) by Definition(identity, identityUniqueness)(G, *) have(thesis) by Tautology.from( lastStep, isNeutral.definition of (e -> identity(G, *)) @@ -358,8 +362,9 @@ object GroupTheory extends lisa.Main { val firstEq = have((group(G, *), x ∈ G, isInverse(y, x, G, *), isInverse(z, x, G, *)) |- op(op(y, *, x), *, z) === z) by Cut(inverseMembership of (y -> z), leftNeutrality) // 2. (yx)z = y(xz) - have((group(G, *), x ∈ G, y ∈ G, z ∈ G) |- (op(op(y, *, x), *, z) === op(y, *, op(x, *, z)))) by Restate.from(associativityApplication of (x -> y, y -> x)) - val associativityCut = thenHave((group(G, *), x ∈ G /\ y ∈ G /\ z ∈ G) |- (op(op(y, *, x), *, z) === op(y, *, op(x, *, z)))) by Restate + val associativityCut = have((group(G, *), x ∈ G /\ y ∈ G /\ z ∈ G) |- (op(op(y, *, x), *, z) === op(y, *, op(x, *, z)))) by Restate.from( + associativity of (x -> y, y -> x) + ) val memberships = have((x ∈ G, isInverse(y, x, G, *), isInverse(z, x, G, *)) |- x ∈ G /\ y ∈ G /\ z ∈ G) by Tautology.from(inverseMembership of (y -> y), inverseMembership of (y -> z)) val secondEq = have((group(G, *), x ∈ G, isInverse(y, x, G, *), isInverse(z, x, G, *)) |- op(op(y, *, x), *, z) === op(y, *, op(x, *, z))) by Cut(memberships, associativityCut) @@ -438,14 +443,7 @@ object GroupTheory extends lisa.Main { } // Prove that inverse(x) is an inverse of x - val inverseIsInverse = have((group(G, *), x ∈ G) |- isInverse(inverse(x, G, *), x, G, *)) subproof { - assume(group(G, *)) - - have(x ∈ G |- ((inverse(x, G, *) === inverse(x, G, *)) <=> isInverse(inverse(x, G, *), x, G, *))) by Restate.from( - inverseCharacterization of (y -> inverse(x, G, *)) - ) - thenHave(thesis) by Restate - } + val inverseIsInverse = have((group(G, *), x ∈ G) |- isInverse(inverse(x, G, *), x, G, *)) by Definition(inverse, inverseUniqueness)(x, G, *) val membership = have((group(G, *), x ∈ G) |- in(inverse(x, G, *), G)) by Tautology.from( inverseIsInverse, @@ -543,8 +541,6 @@ object GroupTheory extends lisa.Main { assume(x ∈ H) assume(y ∈ H) - val a, b = variable - have(subset(H, G)) by Tautology.from(subgroup.definition) have(∀(x, x ∈ H ==> x ∈ G)) by Tautology.from(lastStep, subset.definition of (x -> H, y -> G)) val subsetDef = thenHave(x ∈ H ==> x ∈ G) by InstantiateForall(x) From e3ee43d37b466e4fc21eb12752b59bbd314b00a5 Mon Sep 17 00:00:00 2001 From: dhalilov Date: Sat, 27 May 2023 21:34:44 +0200 Subject: [PATCH 42/68] Migrate tactics to CommonTactics --- .../CommonTactics.scala} | 187 +++++++++--------- .../scala/lisa/mathematics/GroupTheory.scala | 6 +- .../scala/lisa/mathematics/SetTheory.scala | 2 +- 3 files changed, 101 insertions(+), 94 deletions(-) rename src/main/scala/lisa/automation/{grouptheory/GroupTheoryTactics.scala => kernel/CommonTactics.scala} (52%) diff --git a/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala b/src/main/scala/lisa/automation/kernel/CommonTactics.scala similarity index 52% rename from src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala rename to src/main/scala/lisa/automation/kernel/CommonTactics.scala index f5bdaaa4..5d5db1d1 100644 --- a/src/main/scala/lisa/automation/grouptheory/GroupTheoryTactics.scala +++ b/src/main/scala/lisa/automation/kernel/CommonTactics.scala @@ -1,15 +1,16 @@ -package lisa.automation.grouptheory +package lisa.automation.kernel import lisa.automation.kernel.OLPropositionalSolver.Tautology -import lisa.automation.kernel.SimpleSimplifier.Substitution -import lisa.mathematics.SetTheory.* -import lisa.settheory.SetTheoryLibrary -import lisa.prooflib.* +import lisa.kernel.fol.FOL +import lisa.kernel.fol.FOL.* +import lisa.kernel.proof.SequentCalculus.Sequent import lisa.prooflib.BasicStepTactic.* -import lisa.prooflib.ProofTacticLib.{*, given} +import lisa.prooflib.ProofTacticLib.{_, given} +import lisa.prooflib.SimpleDeducedSteps.* +import lisa.prooflib.* +import lisa.utils.KernelHelpers.{_, given} -object GroupTheoryTactics { - import lisa.settheory.SetTheoryLibrary.{*, given} +object CommonTactics { /** *
@@ -18,16 +19,19 @@ object GroupTheoryTactics {
    * Γ, Γ' |- ∃!x. φ, Δ, Δ'
    * 
* - * This tactic separates the existence and the uniqueness proofs, which are often easier to prove independently, - * and thus is easier to use than [[RightExistsOne]]. + * This tactic separates the existence and the uniqueness proofs, which are often easier to prove independently, at + * the expense of brevity. + * + * @see [[RightExistsOne]]. */ object ExistenceAndUniqueness extends ProofTactic { - def withParameters(using proof: SetTheoryLibrary.Proof, om: OutputManager)(phi: Formula, x: VariableLabel, y: VariableLabel) - (existence: proof.Fact, uniqueness: proof.Fact)(bot: Sequent): proof.ProofTacticJudgement = { + def withParameters(using lib: Library, proof: lib.Proof, om: OutputManager)(phi: FOL.Formula, x: FOL.VariableLabel, y: FOL.VariableLabel)(existence: proof.Fact, uniqueness: proof.Fact)( + bot: Sequent + ): proof.ProofTacticJudgement = { val existenceSeq = proof.getSequent(existence) val uniquenessSeq = proof.getSequent(uniqueness) - lazy val substPhi = substituteVariables(phi, Map[VariableLabel, Term](x -> y), Seq()) + lazy val substPhi = substituteVariables(phi, Map[FOL.VariableLabel, FOL.Term](x -> y)) lazy val existenceFormula = ∃(x, phi) lazy val uniqueExistenceFormula = ∃!(x, phi) @@ -35,15 +39,15 @@ object GroupTheoryTactics { if (x == y) { proof.InvalidProofTactic("x and y can not be equal.") } else if (!contains(existenceSeq.right, existenceFormula)) { - proof.InvalidProofTactic(s"Existence sequent conclusion does not contain ∃$x. φ.") + proof.InvalidProofTactic(s"Existence sequent conclusion does not contain ∃x. φ.") } else if (!contains(uniquenessSeq.left, phi)) { proof.InvalidProofTactic("Uniqueness sequent premises do not contain φ.") } else if (!contains(uniquenessSeq.left, substPhi)) { - proof.InvalidProofTactic(s"Uniqueness sequent premises do not contain φ[$y/$x].") + proof.InvalidProofTactic(s"Uniqueness sequent premises do not contain φ[y/x].") } else if (!contains(uniquenessSeq.right, x === y) && !contains(uniquenessSeq.right, y === x)) { - proof.InvalidProofTactic(s"Uniqueness sequent conclusion does not contain $x = $y") + proof.InvalidProofTactic(s"Uniqueness sequent conclusion does not contain x = y") } else if (!contains(bot.right, uniqueExistenceFormula)) { - proof.InvalidProofTactic(s"Bottom sequent conclusion does not contain ∃!$x. φ") + proof.InvalidProofTactic(s"Bottom sequent conclusion does not contain ∃!x. φ") } // Checking pivots @@ -53,49 +57,51 @@ object GroupTheoryTactics { proof.InvalidProofTactic("Could not infer correct right pivots.") } else { val gammaPrime = uniquenessSeq.left.filter(f => !isSame(f, phi) && !isSame(f, substPhi)) - + TacticSubproof { - val forward = have(phi |- ((x === y) ==> substPhi)) subproof { - assume(phi) - thenHave((x === y) |- substPhi) by Substitution - thenHave((x === y) ==> substPhi) by Restate + // There's got to be a better way of importing have/thenHave/assume methods + // but I did not find one + + val forward = lib.have(phi |- ((x === y) ==> substPhi)) subproof { + lib.assume(phi) + lib.thenHave((x === y) |- substPhi) by RightSubstEq(List((x, y)), lambda(x, phi)) + lib.thenHave((x === y) ==> substPhi) by Restate } for (f <- gammaPrime) { - assume(f) + lib.assume(f) } - - val backward = have(phi |- (substPhi ==> (x === y))) by Restate.from(uniqueness) - have(phi |- ((x === y) <=> substPhi)) by RightIff(forward, backward) - thenHave(phi |- ∀(y, (x === y) <=> substPhi)) by RightForall - thenHave(phi |- ∃(x, ∀(y, (x === y) <=> substPhi))) by RightExists - thenHave(∃(x, phi) |- ∃(x, ∀(y, (x === y) <=> substPhi))) by LeftExists - val cut = thenHave(∃(x, phi) |- ∃!(x, phi)) by RightExistsOne + val backward = lib.have(phi |- (substPhi ==> (x === y))) by Restate.from(uniqueness) + + lib.have(phi |- ((x === y) <=> substPhi)) by RightIff(forward, backward) + lib.thenHave(phi |- ∀(y, (x === y) <=> substPhi)) by RightForall + lib.thenHave(phi |- ∃(x, ∀(y, (x === y) <=> substPhi))) by RightExists + lib.thenHave(∃(x, phi) |- ∃(x, ∀(y, (x === y) <=> substPhi))) by LeftExists + lib.thenHave(∃(x, phi) |- ∃!(x, phi)) by RightExistsOne - have(bot) by Cut(existence, cut) + lib.have(bot) by Cut(existence, lib.lastStep) } } } - def apply(using proof: SetTheoryLibrary.Proof, om: OutputManager)(phi: Formula)(existence: proof.Fact, uniqueness: proof.Fact) - (bot: Sequent): proof.ProofTacticJudgement = { + def apply(using lib: Library, proof: lib.Proof, om: OutputManager)(phi: FOL.Formula)(existence: proof.Fact, uniqueness: proof.Fact)(bot: Sequent): proof.ProofTacticJudgement = { val existenceSeq = proof.getSequent(existence) val uniquenessSeq = proof.getSequent(uniqueness) // Try to infer x from the premises // Specifically, find variables in the correct quantifiers, common to all three sequents val existsVars = existenceSeq.right.collect { - case BinderFormula(Exists, x, f) if isSame(f, phi) => x + case FOL.BinderFormula(FOL.Exists, x, f) if isSame(f, phi) => x } if (existsVars.isEmpty) { return proof.InvalidProofTactic("Missing existential quantifier in the existence sequent.") } val commonVars = bot.right.collect { - case BinderFormula(ExistsOne, x, f) if isSame(f, phi) && existsVars.contains(x) => x + case FOL.BinderFormula(FOL.ExistsOne, x, f) if isSame(f, phi) && existsVars.contains(x) => x } - if (commonVars.isEmpty || commonVars.size > 1) { + if (commonVars.size != 1) { return proof.InvalidProofTactic("Could not infer correct variable x in quantifiers.") } @@ -103,11 +109,13 @@ object GroupTheoryTactics { // Infer y from the equalities in the uniqueness sequent uniquenessSeq.right.collectFirst { - case PredicateFormula(`equality`, List(Term(`x`, _), Term(y: VariableLabel, _))) - if x != y && contains(uniquenessSeq.left, substituteVariables(phi, Map[VariableLabel, Term](x -> y), Seq())) => y + case FOL.PredicateFormula(FOL.`equality`, List(FOL.Term(`x`, _), FOL.Term(y: FOL.VariableLabel, _))) + if x != y && contains(uniquenessSeq.left, substituteVariables(phi, Map[FOL.VariableLabel, FOL.Term](x -> y))) => + y - case PredicateFormula(`equality`, List(Term(y: VariableLabel, _), Term(`x`, _))) - if x != y && contains(uniquenessSeq.left, substituteVariables(phi, Map[VariableLabel, Term](x -> y), Seq())) => y + case FOL.PredicateFormula(FOL.`equality`, List(FOL.Term(y: FOL.VariableLabel, _), FOL.Term(`x`, _))) + if x != y && contains(uniquenessSeq.left, substituteVariables(phi, Map[FOL.VariableLabel, FOL.Term](x -> y))) => + y } match { case Some(y) => ExistenceAndUniqueness.withParameters(phi, x, y)(existence, uniqueness)(bot) case None => proof.InvalidProofTactic("Could not infer correct variable y in uniqueness sequent.") @@ -131,24 +139,25 @@ object GroupTheoryTactics { * */ object Definition extends ProofTactic { - def apply(using proof: SetTheoryLibrary.Proof, om: OutputManager) - (f: ConstantFunctionLabel, uniqueness: proof.Fact)(xs: Term*) - (bot: Sequent): proof.ProofTacticJudgement = { - f.definition match { - case SetTheoryLibrary.theory.FunctionDefinition(_, _, expr, _) => + def apply(using lib: Library, proof: lib.Proof)(f: FOL.ConstantFunctionLabel, uniqueness: proof.Fact)(xs: FOL.Term*)(bot: Sequent): proof.ProofTacticJudgement = { + lib.theory.getDefinition(f) match { + case Some(lib.theory.FunctionDefinition(_, _, expr, _)) => // Check if the definition is conditional val method = expr(xs) match { - case ConnectorFormula(And, Seq( - ConnectorFormula(Implies, Seq(a, _)), - ConnectorFormula(Implies, Seq(b, _)) - )) if isSame(neg(a), b) => conditional + case FOL.ConnectorFormula( + FOL.And, + Seq( + FOL.ConnectorFormula(FOL.Implies, Seq(a, _)), + FOL.ConnectorFormula(FOL.Implies, Seq(b, _)) + ) + ) if isSame(FOL.Neg(a), b) => + conditional case _ => unconditional } method(f, uniqueness)(xs)(bot) - case SetTheoryLibrary.theory.Axiom(_, _) => proof.InvalidProofTactic("Axiomatic definitions are not supported.") case _ => proof.InvalidProofTactic("Could not get definition of function.") } } @@ -160,21 +169,19 @@ object GroupTheoryTactics { * |- P(f(xs)) * */ - def unconditional(using proof: SetTheoryLibrary.Proof, om: OutputManager) - (f: ConstantFunctionLabel, uniqueness: proof.Fact)(xs: Term*) - (bot: Sequent): proof.ProofTacticJudgement = { - f.definition match { - case definition@SetTheoryLibrary.theory.FunctionDefinition(_, y, expr, _) => + def unconditional(using lib: Library, proof: lib.Proof)(f: FOL.ConstantFunctionLabel, uniqueness: proof.Fact)(xs: FOL.Term*)(bot: Sequent): proof.ProofTacticJudgement = { + lib.theory.getDefinition(f) match { + case Some(definition @ lib.theory.FunctionDefinition(_, y, expr, _)) => if (bot.right.size != 1) { return proof.InvalidProofTactic("Right-hand side of bottom sequent should contain only 1 formula.") } // Extract variable labels to instantiate them later in the proof - val LambdaTermFormula(vars, _) = expr - val instantiations: Seq[(SchematicTermLabel, LambdaTermTerm)] = vars.zip(xs.map(x => LambdaTermTerm(Seq(), x))) + val FOL.LambdaTermFormula(vars, _) = expr + val instantiations: Seq[(FOL.SchematicTermLabel, FOL.LambdaTermTerm)] = vars.zip(xs.map(x => FOL.LambdaTermTerm(Seq(), x))) // Instantiate terms in the definition - val P = LambdaTermFormula(Seq(y), expr(xs)) + val P = FOL.LambdaTermFormula(Seq(y), expr(xs)) val expected = P(Seq(f(xs))) if (!isSame(expected, bot.right.head)) { @@ -182,13 +189,12 @@ object GroupTheoryTactics { } TacticSubproof { - have(∀(y, (y === f(xs)) <=> P(Seq(y)))) by Tautology.from(uniqueness, definition.of(instantiations: _*)) - thenHave((y === f(xs)) <=> P(Seq(y))) by InstantiateForall(y) - thenHave((f(xs) === f(xs)) <=> P(Seq(f(xs)))) by InstFunSchema(Map(y -> f(xs))) - thenHave(P(Seq(f(xs)))) by Restate + lib.have(∀(y, (y === f(xs)) <=> P(Seq(y)))) by Tautology.from(uniqueness, definition.of(instantiations: _*)) + lib.thenHave((y === f(xs)) <=> P(Seq(y))) by InstantiateForall(y) + lib.thenHave((f(xs) === f(xs)) <=> P(Seq(f(xs)))) by InstFunSchema(Map(y -> f(xs))) + lib.thenHave(P(Seq(f(xs)))) by Restate } - case SetTheoryLibrary.theory.Axiom(_, _) => proof.InvalidProofTactic("Axiomatic definitions are not supported.") case _ => proof.InvalidProofTactic("Could not get definition of function.") } } @@ -200,11 +206,9 @@ object GroupTheoryTactics { * φ |- Q(f(xs)) * */ - def conditional(using proof: SetTheoryLibrary.Proof, om: OutputManager) - (f: ConstantFunctionLabel, uniqueness: proof.Fact)(xs: Term*) - (bot: Sequent): proof.ProofTacticJudgement = { - f.definition match { - case definition@SetTheoryLibrary.theory.FunctionDefinition(_, y, expr, _) => + def conditional(using lib: Library, proof: lib.Proof)(f: FOL.ConstantFunctionLabel, uniqueness: proof.Fact)(xs: FOL.Term*)(bot: Sequent): proof.ProofTacticJudgement = { + lib.theory.getDefinition(f) match { + case Some(definition @ lib.theory.FunctionDefinition(_, y, expr, _)) => if (bot.right.size != 1) { return proof.InvalidProofTactic("Right-hand side of bottom sequent should contain exactly 1 formula.") } else if (bot.left.isEmpty) { @@ -212,22 +216,26 @@ object GroupTheoryTactics { } // Extract variable labels to instantiate them later in the proof - val LambdaTermFormula(vars, _) = expr - val instantiations: Seq[(SchematicTermLabel, LambdaTermTerm)] = vars.zip(xs.map(x => LambdaTermTerm(Seq(), x))) + val FOL.LambdaTermFormula(vars, _) = expr + val instantiations: Seq[(FOL.SchematicTermLabel, FOL.LambdaTermTerm)] = vars.zip(xs.map(x => FOL.LambdaTermTerm(Seq(), x))) // Instantiate terms in the definition - val P = LambdaTermFormula(Seq(y), expr(xs)) + val P = FOL.LambdaTermFormula(Seq(y), expr(xs)) // Unfold the conditional definition to find Q - val phi = ConnectorFormula(And, bot.left.toSeq) + val phi = FOL.ConnectorFormula(FOL.And, bot.left.toSeq) val Q = P.body match { - case ConnectorFormula(And, Seq( - ConnectorFormula(Implies, Seq(a, f)), - ConnectorFormula(Implies, Seq(b, g)) - )) if isSame(neg(a), b) => - if (isSame(a, phi)) LambdaTermFormula(Seq(y), f) - else if (isSame(b, phi)) LambdaTermFormula(Seq(y), g) + case FOL.ConnectorFormula( + FOL.And, + Seq( + FOL.ConnectorFormula(FOL.Implies, Seq(a, f)), + FOL.ConnectorFormula(FOL.Implies, Seq(b, g)) + ) + ) if isSame(FOL.Neg(a), b) => { + if (isSame(a, phi)) FOL.LambdaTermFormula(Seq(y), f) + else if (isSame(b, phi)) FOL.LambdaTermFormula(Seq(y), g) else return proof.InvalidProofTactic("Condition of definition is not satisfied.") + } case _ => return proof.InvalidProofTactic("Definition is not conditional.") } @@ -238,15 +246,14 @@ object GroupTheoryTactics { } TacticSubproof { - have(∀(y, (y === f(xs)) <=> P(Seq(y)))) by Tautology.from(uniqueness, definition.of(instantiations: _*)) - thenHave((y === f(xs)) <=> P(Seq(y))) by InstantiateForall(y) - thenHave((f(xs) === f(xs)) <=> P(Seq(f(xs)))) by InstFunSchema(Map(y -> f(xs))) - thenHave(P(Seq(f(xs)))) by Restate - thenHave(phi ==> Q(Seq(f(xs)))) by Tautology - thenHave(phi |- Q(Seq(f(xs)))) by Restate + lib.have(∀(y, (y === f(xs)) <=> P(Seq(y)))) by Tautology.from(uniqueness, definition.of(instantiations: _*)) + lib.thenHave((y === f(xs)) <=> P(Seq(y))) by InstantiateForall(y) + lib.thenHave((f(xs) === f(xs)) <=> P(Seq(f(xs)))) by InstFunSchema(Map(y -> f(xs))) + lib.thenHave(P(Seq(f(xs)))) by Restate + lib.thenHave(phi ==> Q(Seq(f(xs)))) by Tautology + lib.thenHave(bot) by Restate } - case SetTheoryLibrary.theory.Axiom(_, _) => proof.InvalidProofTactic("Axiomatic definitions are not supported.") case _ => proof.InvalidProofTactic("Could not get definition of function.") } } @@ -263,7 +270,7 @@ object GroupTheoryTactics { * TODO: Extending the tactic to more general pivots */ object Cases extends ProofTactic { - def withParameters(using proof: SetTheoryLibrary.Proof, om: OutputManager)(phi: Formula)(pos: proof.Fact, neg: proof.Fact)(bot: Sequent): proof.ProofTacticJudgement = { + def withParameters(using lib: Library, proof: lib.Proof, om: OutputManager)(phi: Formula)(pos: proof.Fact, neg: proof.Fact)(bot: Sequent): proof.ProofTacticJudgement = { val seqPos = proof.getSequent(pos) val seqNeg = proof.getSequent(neg) @@ -273,15 +280,15 @@ object GroupTheoryTactics { } else { val gamma = bot.left TacticSubproof { - val excludedMiddle = have(phi \/ !phi) by Tautology - val toCut = have((gamma + (phi \/ !phi)) |- bot.right) by LeftOr(pos, neg) + val excludedMiddle = lib.have(phi \/ !phi) by Tautology + val toCut = lib.have((gamma + (phi \/ !phi)) |- bot.right) by LeftOr(pos, neg) - have(bot) by Cut(excludedMiddle, toCut) + lib.have(bot) by Cut(excludedMiddle, toCut) } } } - def apply(using proof: SetTheoryLibrary.Proof, om: OutputManager)(pos: proof.Fact, neg: proof.Fact)(bot: Sequent): proof.ProofTacticJudgement = { + def apply(using lib: Library, proof: lib.Proof, om: OutputManager)(pos: proof.Fact, neg: proof.Fact)(bot: Sequent): proof.ProofTacticJudgement = { val seqPos = proof.getSequent(pos) val pivot = seqPos.left -- bot.left diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala index 5f39117d..2a02f687 100644 --- a/src/main/scala/lisa/mathematics/GroupTheory.scala +++ b/src/main/scala/lisa/mathematics/GroupTheory.scala @@ -1,8 +1,8 @@ package lisa.mathematics -import lisa.automation.grouptheory.GroupTheoryTactics.Cases -import lisa.automation.grouptheory.GroupTheoryTactics.Definition -import lisa.automation.grouptheory.GroupTheoryTactics.ExistenceAndUniqueness +import lisa.automation.kernel.CommonTactics.Cases +import lisa.automation.kernel.CommonTactics.Definition +import lisa.automation.kernel.CommonTactics.ExistenceAndUniqueness import lisa.automation.kernel.OLPropositionalSolver.Tautology import lisa.mathematics.FirstOrderLogic.equalityTransitivity import lisa.mathematics.FirstOrderLogic.existsOneImpliesExists diff --git a/src/main/scala/lisa/mathematics/SetTheory.scala b/src/main/scala/lisa/mathematics/SetTheory.scala index a046ede1..11257644 100644 --- a/src/main/scala/lisa/mathematics/SetTheory.scala +++ b/src/main/scala/lisa/mathematics/SetTheory.scala @@ -1,6 +1,6 @@ package lisa.mathematics -import lisa.automation.grouptheory.GroupTheoryTactics.Definition +import lisa.automation.kernel.CommonTactics.Definition import lisa.automation.kernel.OLPropositionalSolver.Tautology import lisa.automation.kernel.SimpleSimplifier.* import lisa.automation.settheory.SetTheoryTactics.* From 59534af7695c0d7020e9fea01a686e8fb6dd20cf Mon Sep 17 00:00:00 2001 From: dhalilov Date: Tue, 30 May 2023 23:59:00 +0200 Subject: [PATCH 43/68] Create Equalities tactic --- .../automation/kernel/CommonTactics.scala | 113 ++++++++++++++++++ 1 file changed, 113 insertions(+) diff --git a/src/main/scala/lisa/automation/kernel/CommonTactics.scala b/src/main/scala/lisa/automation/kernel/CommonTactics.scala index 5d5db1d1..1e90c5a4 100644 --- a/src/main/scala/lisa/automation/kernel/CommonTactics.scala +++ b/src/main/scala/lisa/automation/kernel/CommonTactics.scala @@ -9,6 +9,9 @@ import lisa.prooflib.ProofTacticLib.{_, given} import lisa.prooflib.SimpleDeducedSteps.* import lisa.prooflib.* import lisa.utils.KernelHelpers.{_, given} +import lisa.utils.FOLPrinter +import scala.collection.immutable.Queue +import scala.collection.mutable.{Map => MutableMap, Set => MutableSet} object CommonTactics { @@ -259,6 +262,116 @@ object CommonTactics { } } + /** + *
+   * Γ |- a(0) === a(1)   Γ' |- a(2) === a(3) ...
+   * --------------------------------------------
+   * Γ, Γ', ... |- a(i) === a(j)
+   * 
+ * Proves any equality induced transitively by the equalities of the premises. + * + * Internally, we construct an undirected graph, where sides of an equality are represented by vertices, and + * an edge between two terms `a` and `b` means that some premise proves `a === b` (or equivalently `b === a`). + * We also keep the premise from which the equality stems as a label of the edge to construct the final antecedent of + * the bottom sequent. + * + * We can see that an equality `a === b` is provable from the premises if and only if `a` is reachable from `b`. + * We thus run Breadth-First Search (BFS) on the graph starting from `a` to find the smallest solution (in terms of + * sequent calculus steps), if any. + */ + object Equalities extends ProofTactic { + def apply(using lib: Library, proof: lib.Proof)(equalities: proof.Fact*)(bot: Sequent): proof.ProofTacticJudgement = { + // Construct the graph as an adjacency list for O(1) equality checks + val graph = MutableMap[FOL.Term, List[FOL.Term]]() + val labels = MutableMap[(FOL.Term, FOL.Term), proof.Fact]() + + // Use a variable to avoid non-local returns + // This is because the below loop is rewritten using maps and filters under the hood + // I have to say, this is pretty ugly :( + var error: Option[proof.InvalidProofTactic] = None + for (premise <- equalities; f <- proof.getSequent(premise).right) { + f match { + case FOL.PredicateFormula(FOL.`equality`, Seq(x: FOL.Term, y: FOL.Term)) => + if (error.isEmpty) { + // In case of conflicts, it would be too costly in the general case to find which premise is appropriate + // We simply throw an error to indicate that something is wrong with the premises + if (labels.contains((x, y)) && labels((x, y)) != premise) { + // TODO Indicate which premises lead to the error + error = Some(proof.InvalidProofTactic(s"Equality ${FOLPrinter.prettyTerm(x)} === ${FOLPrinter.prettyTerm(y)} was proven in two different premises.")) + } else { + graph(x) = y :: graph.getOrElse(x, Nil) + graph(y) = x :: graph.getOrElse(y, Nil) + labels += ((x, y) -> premise) + labels += ((y, x) -> premise) + } + } + case _ => + if (error.isEmpty) { + error = Some(proof.InvalidProofTactic("Right-hand side of premises should only contain equalities.")) + } + } + } + + if (error.nonEmpty) { + return error.get + } + + if (bot.right.size != 1) { + return proof.InvalidProofTactic(s"Right-hand side of bottom sequent expected exactly 1 formula, got ${bot.right.size}") + } + + bot.right.head match { + case FOL.PredicateFormula(FOL.`equality`, Seq(x: FOL.Term, y: FOL.Term)) => + // Optimization in the trivial case x === x + if (isSameTerm(x, y)) { + return TacticSubproof { + lib.have(x === y) by RightRefl + if (bot.left.nonEmpty) { + lib.thenHave(bot) by Weakening + } + } + } + + // Run BFS on the graph + var Q = Queue[FOL.Term](x) + val explored = MutableSet[FOL.Term](x) + val parent = MutableMap[FOL.Term, FOL.Term]() + while (Q.nonEmpty) { + val (v, newQ) = Q.dequeue + Q = newQ + if (v == y) { + lazy val traversal: LazyList[FOL.Term] = y #:: traversal.map(parent) + val path = (traversal.tail.takeWhile(_ != x) :+ x).toList // Path from y (excluded) to x (included) + return TacticSubproof { + var label = labels((y, path.head)) + def labelSeq = proof.getSequent(label) + var leftEq = lib.have(labelSeq.left |- y === path.head) by Restate.from(label) + def leftEqSeq = proof.getSequent(label) + + for ((a, b) <- path.zip(path.tail)) { + label = labels((a, b)) + + val z = variable + val cut1 = lib.have(labelSeq.left |- a === b) by Restate.from(label) + val cut2 = lib.have((leftEqSeq.left + (a === b)) |- y === b) by RightSubstEq(List((a, b)), lambda(z, y === z))(leftEq) + leftEq = lib.have((leftEqSeq.left ++ labelSeq.left) |- y === b) by Cut(cut1, cut2) + } + } + } + + for (w <- graph(v) if !explored.contains(w)) { + explored += w + parent(w) = v + Q = Q.enqueue(w) + } + } + proof.InvalidProofTactic("Equality is not provable from the premises.") + + case _ => proof.InvalidProofTactic("Right-hand side of bottom sequent should be of the form x === y.") + } + } + } + /** *
    * Γ, φ |- Δ, Σ   Γ, ¬φ |- Δ, Σ'  

From 1479852fe57883c2538674c11a1d21e0c9af5897 Mon Sep 17 00:00:00 2001
From: Sankalp Gambhir 
Date: Tue, 2 May 2023 10:26:16 +0200
Subject: [PATCH 44/68] Add assumptions to subproof checking (#160)

Co-authored-by: SimonGuilloud 
---
 lisa-utils/src/main/scala/lisa/prooflib/ProofsHelpers.scala | 3 ++-
 1 file changed, 2 insertions(+), 1 deletion(-)

diff --git a/lisa-utils/src/main/scala/lisa/prooflib/ProofsHelpers.scala b/lisa-utils/src/main/scala/lisa/prooflib/ProofsHelpers.scala
index 92f6f209..d89d0e14 100644
--- a/lisa-utils/src/main/scala/lisa/prooflib/ProofsHelpers.scala
+++ b/lisa-utils/src/main/scala/lisa/prooflib/ProofsHelpers.scala
@@ -33,7 +33,8 @@ trait ProofsHelpers {
     }
 
     inline infix def subproof(using proof: Library#Proof, om: OutputManager, line: sourcecode.Line, file: sourcecode.File)(tactic: proof.InnerProof ?=> Unit): proof.ProofStep = {
-      (new BasicStepTactic.SUBPROOF(using proof)(Some(bot))(tactic)).judgement.validate(line, file).asInstanceOf[proof.ProofStep]
+      val botWithAssumptions = HaveSequent.this.bot ++ (proof.getAssumptions |- ())
+      (new BasicStepTactic.SUBPROOF(using proof)(Some(botWithAssumptions))(tactic)).judgement.validate(line, file).asInstanceOf[proof.ProofStep]
     }
 
   }

From f1da07ea8f81a0c13762f667336b832454ad9468 Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Wed, 31 May 2023 18:15:30 +0200
Subject: [PATCH 45/68] Fix Equalities assumption bug

---
 .../automation/kernel/CommonTactics.scala     | 63 +++++++++++++------
 1 file changed, 45 insertions(+), 18 deletions(-)

diff --git a/src/main/scala/lisa/automation/kernel/CommonTactics.scala b/src/main/scala/lisa/automation/kernel/CommonTactics.scala
index 1e90c5a4..609580d4 100644
--- a/src/main/scala/lisa/automation/kernel/CommonTactics.scala
+++ b/src/main/scala/lisa/automation/kernel/CommonTactics.scala
@@ -283,11 +283,10 @@ object CommonTactics {
     def apply(using lib: Library, proof: lib.Proof)(equalities: proof.Fact*)(bot: Sequent): proof.ProofTacticJudgement = {      
       // Construct the graph as an adjacency list for O(1) equality checks
       val graph = MutableMap[FOL.Term, List[FOL.Term]]()
-      val labels = MutableMap[(FOL.Term, FOL.Term), proof.Fact]()
+      val premises = MutableMap[(FOL.Term, FOL.Term), proof.Fact]()
 
       // Use a variable to avoid non-local returns
       // This is because the below loop is rewritten using maps and filters under the hood
-      // I have to say, this is pretty ugly :(
       var error: Option[proof.InvalidProofTactic] = None
       for (premise <- equalities; f <- proof.getSequent(premise).right) {
         f match {
@@ -295,14 +294,14 @@ object CommonTactics {
             if (error.isEmpty) {
               // In case of conflicts, it would be too costly in the general case to find which premise is appropriate
               // We simply throw an error to indicate that something is wrong with the premises
-              if (labels.contains((x, y)) && labels((x, y)) != premise) {
+              if (premises.contains((x, y)) && premises((x, y)) != premise) {
                 // TODO Indicate which premises lead to the error
                 error = Some(proof.InvalidProofTactic(s"Equality ${FOLPrinter.prettyTerm(x)} === ${FOLPrinter.prettyTerm(y)} was proven in two different premises."))
               } else {
                 graph(x) = y :: graph.getOrElse(x, Nil)
                 graph(y) = x :: graph.getOrElse(y, Nil)
-                labels += ((x, y) -> premise)
-                labels += ((y, x) -> premise)
+                premises += ((x, y) -> premise)
+                premises += ((y, x) -> premise)
               }
             }
           case _ => 
@@ -342,21 +341,49 @@ object CommonTactics {
             if (v == y) {
               lazy val traversal: LazyList[FOL.Term] = y #:: traversal.map(parent)
               val path = (traversal.tail.takeWhile(_ != x) :+ x).toList // Path from y (excluded) to x (included)
-              return TacticSubproof {
-                var label = labels((y, path.head))
-                def labelSeq = proof.getSequent(label)
-                var leftEq = lib.have(labelSeq.left |- y === path.head) by Restate.from(label)
-                def leftEqSeq = proof.getSequent(label)
-
-                for ((a, b) <- path.zip(path.tail)) {
-                  label = labels((a, b))
-
-                  val z = variable
-                  val cut1 = lib.have(labelSeq.left |- a === b) by Restate.from(label)
-                  val cut2 = lib.have((leftEqSeq.left + (a === b)) |- y === b) by RightSubstEq(List((a, b)), lambda(z, y === z))(leftEq)
-                  leftEq = lib.have((leftEqSeq.left ++ labelSeq.left) |- y === b) by Cut(cut1, cut2)
+
+              def getEqTerms(using lib: Library, proof: lib.Proof)(eq: proof.Fact): (FOL.Term, FOL.Term) =
+                proof.getSequent(eq).right.head match {
+                  case FOL.PredicateFormula(`equality`, Seq(a, b)) => (a, b)
+                  case _ => throw IllegalArgumentException("Not an equality.")
+                }
+
+              def order(using lib: Library, proof: lib.Proof)(eq: proof.Fact, first: FOL.Term, second: FOL.Term): proof.Fact = {
+                val seq = proof.getSequent(eq)
+                seq.right.head match {
+                  case FOL.PredicateFormula(`equality`, Seq(`first`, `second`)) => eq
+                  case FOL.PredicateFormula(`equality`, Seq(`second`, `first`)) => {
+                    val u = variable
+                    lib.have(first === first) by RightRefl
+                    lib.thenHave(second === first |- first === second) by RightSubstEq(List((second, first)), lambda(u, first === u))
+                    lib.have(seq.left |- first === second) by Cut(eq, lib.lastStep)
+                  }
+                  case _ => throw IllegalArgumentException("First or last is not present in the given equality.")
                 }
               }
+
+              return TacticSubproof { (innerProof: proof.InnerProof) ?=>
+                val initialStep = order(premises(y -> path.head), path.head, y)
+                val u = variable
+
+                path.zip(path.tail).foldLeft(initialStep)((leftEq, vars) => {
+                  val leftSeq = innerProof.getSequent(leftEq)
+                  val (a, b) = vars
+                  val premiseEq = premises(vars)
+                  val premiseSeq = proof.getSequent(premiseEq)
+                  
+                  // Apply equality transitivity on (a === y) /\ (a === b) to get (b === y)
+                  // TODO Watch for issue #161, as assumptions will break this step
+                  lib.have((leftSeq.left + premiseSeq.right.head) |- b === y) by RightSubstEq(
+                    List(getEqTerms(premiseEq)),
+                    lambda(u, u === y)
+                  )(leftEq)
+                  val seq = (leftSeq.left ++ premiseSeq.left) |- b === y
+                  val eq = lib.have(seq) by Cut(premiseEq, lib.lastStep)
+
+                  eq
+                })
+              }
             }
 
             for (w <- graph(v) if !explored.contains(w)) {

From 4110189f9b050401e61c0524d596938ee45e0a43 Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Wed, 31 May 2023 18:25:59 +0200
Subject: [PATCH 46/68] Simplify equality proofs

---
 .../scala/lisa/mathematics/GroupTheory.scala  | 23 ++++---------------
 1 file changed, 4 insertions(+), 19 deletions(-)

diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala
index 2a02f687..859e4f1c 100644
--- a/src/main/scala/lisa/mathematics/GroupTheory.scala
+++ b/src/main/scala/lisa/mathematics/GroupTheory.scala
@@ -2,6 +2,7 @@ package lisa.mathematics
 
 import lisa.automation.kernel.CommonTactics.Cases
 import lisa.automation.kernel.CommonTactics.Definition
+import lisa.automation.kernel.CommonTactics.Equalities
 import lisa.automation.kernel.CommonTactics.ExistenceAndUniqueness
 import lisa.automation.kernel.OLPropositionalSolver.Tautology
 import lisa.mathematics.FirstOrderLogic.equalityTransitivity
@@ -288,10 +289,7 @@ object GroupTheory extends lisa.Main {
       val secondEq = thenHave(e === op(e, *, f)) by Tautology
 
       // 3. Conclude by transitivity
-      val eqs = have((e === op(e, *, f)) /\ (op(e, *, f) === f)) by RightAnd(secondEq, firstEq)
-      have(((e === op(e, *, f)) /\ (op(e, *, f) === f)) |- (e === f)) by Tautology.from(equalityTransitivity of (x -> e, y -> op(e, *, f), z -> f))
-
-      have(e === f) by Cut(eqs, lastStep)
+      have(e === f) by Equalities(firstEq, secondEq)
     }
 
     have(group(G, *) |- ∃!(e, isNeutral(e, G, *))) by ExistenceAndUniqueness(isNeutral(e, G, *))(existence, uniqueness)
@@ -381,21 +379,8 @@ object GroupTheory extends lisa.Main {
       }
       val thirdEq = have((group(G, *), x ∈ G, isInverse(y, x, G, *), isInverse(z, x, G, *)) |- op(y, *, op(x, *, z)) === y) by Cut(inverseMembership of (y -> y), rightNeutrality)
 
-      // Conclude by transitivity
-
-      // 4. z = y(xz)
-      val fourthEq = have((group(G, *), x ∈ G, isInverse(y, x, G, *), isInverse(z, x, G, *)) |- z === op(y, *, op(x, *, z))) by Tautology.from(
-        firstEq,
-        secondEq,
-        equalityTransitivity of (x -> z, y -> op(op(y, *, x), *, z), z -> op(y, *, op(x, *, z)))
-      )
-
-      // 5. z = y
-      have((group(G, *), x ∈ G, isInverse(y, x, G, *), isInverse(z, x, G, *)) |- z === y) by Tautology.from(
-        thirdEq,
-        fourthEq,
-        equalityTransitivity of (x -> z, y -> op(y, *, op(x, *, z)), z -> y)
-      )
+      // 4. z = y
+      have((group(G, *), x ∈ G, isInverse(y, x, G, *), isInverse(z, x, G, *)) |- z === y) by Equalities(firstEq, secondEq, thirdEq)
     }
 
     have(thesis) by ExistenceAndUniqueness(isInverse(y, x, G, *))(existence, uniqueness)

From 5fa02b630644ef28e16e2a2489ac0c44fc747774 Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Wed, 31 May 2023 18:26:15 +0200
Subject: [PATCH 47/68] Prove left cancellation law

---
 .../scala/lisa/mathematics/GroupTheory.scala  | 43 +++++++++++++++++++
 1 file changed, 43 insertions(+)

diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala
index 859e4f1c..5dcd4842 100644
--- a/src/main/scala/lisa/mathematics/GroupTheory.scala
+++ b/src/main/scala/lisa/mathematics/GroupTheory.scala
@@ -457,6 +457,49 @@ object GroupTheory extends lisa.Main {
     thenHave(thesis) by RightForall
   }
 
+  /**
+   * Theorem --- In a group `(G, *)`, we have `xy = xz ==> y = z`.
+   */
+  val leftCancellation = Theorem(
+    (group(G, *), x ∈ G, y ∈ G, z ∈ G) |- (op(x, *, y) === op(x, *, z)) ==> (y === z)
+  ) {
+    val i = inverse(x, G, *)
+    val inverseDef = have((group(G, *), x ∈ G) |- isInverse(i, x, G, *)) by Definition(inverse, inverseUniqueness)(x, G, *)
+    val inverseMembership = have((group(G, *), x ∈ G) |- i ∈ G) by Tautology.from(inverseDef, isInverse.definition of (y -> i))
+
+    // 1. Prove that i * (xy) = y and i * (xz) = z
+    val cancellation = have((group(G, *), x ∈ G, y ∈ G) |- op(i, *, op(x, *, y)) === y) subproof {
+      // (ix)y = i(xy)
+      val eq1 = have((group(G, *), x ∈ G, y ∈ G) |- op(op(i, *, x), *, y) === op(i, *, op(x, *, y))) by Cut(
+        inverseMembership,
+        associativity of (x -> i, y -> x, z -> y)
+      )
+      
+      // (ix)y = y
+      have((group(G, *), x ∈ G) |- ∀(y, (y ∈ G) ==> ((op(op(i, *, x), *, y) === y) /\ (op(y, *, op(i, *, x)) === y)))) by Tautology.from(
+        inverseDef,
+        isInverse.definition of (y -> i),
+        isNeutral.definition of (e -> op(i, *, x))
+      )
+      thenHave((group(G, *), x ∈ G) |- (y ∈ G) ==> ((op(op(i, *, x), *, y) === y) /\ (op(y, *, op(i, *, x)) === y))) by InstantiateForall(y)
+      val eq2 = thenHave((group(G, *), x ∈ G, y ∈ G) |- op(op(i, *, x), *, y) === y) by Tautology
+
+      // i(xy) = y
+      have(thesis) by Equalities(eq1, eq2)
+    }
+
+    // 2. By substitution, xy = xz implies i(xy) = i(xz)
+    have(op(i, *, op(x, *, y)) === op(i, *, op(x, *, y))) by RightRefl
+    val substitution = thenHave(op(x, *, y) === op(x, *, z) |- op(i, *, op(x, *, y)) === op(i, *, op(x, *, z))) by RightSubstEq(
+      List((op(x, *, y), op(x, *, z))),
+      lambda(z, op(i, *, op(x, *, y)) === op(i, *, z))
+    )
+
+    // 3. Conclude that xy = xz ==> y === z
+    have((group(G, *), x ∈ G, y ∈ G, z ∈ G, op(x, *, y) === op(x, *, z)) |- y === z) by Equalities(cancellation, cancellation of (y -> z), substitution)
+    thenHave(thesis) by Restate
+  }
+
   //
   // 1.1 Subgroups
   //

From 69872902f49e0a56bf867569dc39a6ad6a192fe0 Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Wed, 31 May 2023 18:39:02 +0200
Subject: [PATCH 48/68] Create inverse lemmas

---
 .../scala/lisa/mathematics/GroupTheory.scala  | 37 ++++++++++++-------
 1 file changed, 24 insertions(+), 13 deletions(-)

diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala
index 5dcd4842..0d9cb0cb 100644
--- a/src/main/scala/lisa/mathematics/GroupTheory.scala
+++ b/src/main/scala/lisa/mathematics/GroupTheory.scala
@@ -391,6 +391,27 @@ object GroupTheory extends lisa.Main {
    */
   val inverse = DEF(x, G, *) --> TheConditional(y, isInverse(y, x, G, *))(inverseUniqueness)
 
+  /**
+   * Lemma --- The inverse of `x` is an inverse of `x` (by definition).
+   */
+  val inverseIsInverse = Lemma(
+    (group(G, *), x ∈ G) |- isInverse(inverse(x, G, *), x, G, *)
+  ) {
+    have(thesis) by Definition(inverse, inverseUniqueness)(x, G, *)
+  }
+
+  /**
+   * Lemma --- The inverse element `y` of `x` is in `G`.
+   */
+  val inverseIsInGroup = Lemma(
+    (group(G, *), x ∈ G) |- inverse(x, G, *) ∈ G
+  ) {
+    have(thesis) by Tautology.from(
+      inverseIsInverse,
+      isInverse.definition of (y -> inverse(x, G, *))
+    )
+  }
+
   /**
    * Theorem --- `y` is the inverse of `x` iff `x` is the inverse of `y`
    */
@@ -427,14 +448,6 @@ object GroupTheory extends lisa.Main {
       thenHave(thesis) by InstantiateForall(y)
     }
 
-    // Prove that inverse(x) is an inverse of x
-    val inverseIsInverse = have((group(G, *), x ∈ G) |- isInverse(inverse(x, G, *), x, G, *)) by Definition(inverse, inverseUniqueness)(x, G, *)
-
-    val membership = have((group(G, *), x ∈ G) |- in(inverse(x, G, *), G)) by Tautology.from(
-      inverseIsInverse,
-      isInverse.definition of (y -> inverse(x, G, *))
-    )
-
     // Use the symmetry of the inverse relation to prove that x is an inverse of inverse(x)
     val satisfaction = have((group(G, *), x ∈ G) |- isInverse(x, inverse(x, G, *), G, *)) subproof {
       assume(group(G, *))
@@ -452,7 +465,7 @@ object GroupTheory extends lisa.Main {
     val characterization = have(in(inverse(x, G, *), G) |- ((x === u) <=> isInverse(x, inverse(x, G, *), G, *))) by Restate.from(inverseCharacterization of (x -> inverse(x, G, *), y -> x))
     val eq = have((x ∈ G, in(inverse(x, G, *), G)) |- (x === u)) by Tautology.from(satisfaction, characterization)
 
-    have(x ∈ G |- (x === u)) by Cut(membership, eq)
+    have(x ∈ G |- (x === u)) by Cut(inverseIsInGroup, eq)
     thenHave(x ∈ G ==> (x === u)) by Restate
     thenHave(thesis) by RightForall
   }
@@ -464,20 +477,18 @@ object GroupTheory extends lisa.Main {
     (group(G, *), x ∈ G, y ∈ G, z ∈ G) |- (op(x, *, y) === op(x, *, z)) ==> (y === z)
   ) {
     val i = inverse(x, G, *)
-    val inverseDef = have((group(G, *), x ∈ G) |- isInverse(i, x, G, *)) by Definition(inverse, inverseUniqueness)(x, G, *)
-    val inverseMembership = have((group(G, *), x ∈ G) |- i ∈ G) by Tautology.from(inverseDef, isInverse.definition of (y -> i))
 
     // 1. Prove that i * (xy) = y and i * (xz) = z
     val cancellation = have((group(G, *), x ∈ G, y ∈ G) |- op(i, *, op(x, *, y)) === y) subproof {
       // (ix)y = i(xy)
       val eq1 = have((group(G, *), x ∈ G, y ∈ G) |- op(op(i, *, x), *, y) === op(i, *, op(x, *, y))) by Cut(
-        inverseMembership,
+        inverseIsInGroup,
         associativity of (x -> i, y -> x, z -> y)
       )
       
       // (ix)y = y
       have((group(G, *), x ∈ G) |- ∀(y, (y ∈ G) ==> ((op(op(i, *, x), *, y) === y) /\ (op(y, *, op(i, *, x)) === y)))) by Tautology.from(
-        inverseDef,
+        inverseIsInverse,
         isInverse.definition of (y -> i),
         isNeutral.definition of (e -> op(i, *, x))
       )

From ca4492de98de438db705afaaf03acf44f7984b56 Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Wed, 31 May 2023 19:06:14 +0200
Subject: [PATCH 49/68] Prove right cancellation law

---
 .../scala/lisa/mathematics/GroupTheory.scala  | 63 ++++++++++++++++---
 1 file changed, 55 insertions(+), 8 deletions(-)

diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala
index 0d9cb0cb..87c85cd1 100644
--- a/src/main/scala/lisa/mathematics/GroupTheory.scala
+++ b/src/main/scala/lisa/mathematics/GroupTheory.scala
@@ -511,6 +511,49 @@ object GroupTheory extends lisa.Main {
     thenHave(thesis) by Restate
   }
 
+  /**
+   * Theorem --- In a group `(G, *)`, we have `yx = zx ==> y = z`.
+   * 
+   * Analoguous to [[leftCancellation]].
+   */
+  val rightCancellation = Theorem(
+    (group(G, *), x ∈ G, y ∈ G, z ∈ G) |- (op(y, *, x) === op(z, *, x)) ==> (y === z)
+  ) {
+    val i = inverse(x, G, *)
+
+    // 1. Prove that (yx)i = y and (zx)i = z
+    val cancellation = have((group(G, *), x ∈ G, y ∈ G) |- op(op(y, *, x), *, i) === y) subproof {
+      // (xy)i = y(xi)
+      val eq1 = have((group(G, *), x ∈ G, y ∈ G) |- op(op(y, *, x), *, i) === op(y, *, op(x, *, i))) by Cut(
+        inverseIsInGroup,
+        associativity of (x -> y, y -> x, z -> i)
+      )
+      
+      // y(xi) = y
+      have((group(G, *), x ∈ G) |- ∀(y, (y ∈ G) ==> ((op(op(x, *, i), *, y) === y) /\ (op(y, *, op(x, *, i)) === y)))) by Tautology.from(
+        inverseIsInverse,
+        isInverse.definition of (y -> i),
+        isNeutral.definition of (e -> op(x, *, i))
+      )
+      thenHave((group(G, *), x ∈ G) |- (y ∈ G) ==> ((op(op(x, *, i), *, y) === y) /\ (op(y, *, op(x, *, i)) === y))) by InstantiateForall(y)
+      val eq2 = thenHave((group(G, *), x ∈ G, y ∈ G) |- op(y, *, op(x, *, i)) === y) by Tautology
+
+      // (yx)i = y
+      have(thesis) by Equalities(eq1, eq2)
+    }
+
+    // 2. By substitution, yx = zx implies (yx)i = (zx)i
+    have(op(op(y, *, x), *, i) === op(op(y, *, x), *, i)) by RightRefl
+    val substitution = thenHave(op(y, *, x) === op(z, *, x) |- op(op(y, *, x), *, i) === op(op(z, *, x), *, i)) by RightSubstEq(
+      List((op(y, *, x), op(z, *, x))),
+      lambda(z, op(op(y, *, x), *, i) === op(z, *, i))
+    )
+
+    // 3. Conclude that yx = zx ==> y === z
+    have((group(G, *), x ∈ G, y ∈ G, z ∈ G, op(y, *, x) === op(z, *, x)) |- y === z) by Equalities(cancellation, cancellation of (y -> z), substitution)
+    thenHave(thesis) by Restate
+  }
+
   //
   // 1.1 Subgroups
   //
@@ -677,21 +720,25 @@ object GroupTheory extends lisa.Main {
   /**
    * Identity in subgroup --- The identity element of the subgroup is exactly the identity element of the parent group.
    */
-  val subgroupIdentity = Theorem(
+  /*val subgroupIdentity = Theorem(
     subgroup(H, G, *) |- identity(H, restrictedFunction(*, cartesianProduct(H, H))) === identity(G, *)
   ) {
-    /*// The crux of the proof relies on proving that identity(G, *) ∈ H
+    // The crux of the proof relies on proving that identity(G, *) ∈ H
     // From there we can derive quite easily the theorem using [[subgroupOperation]]
     have(subgroup(H, G, *) |- identity(G, *) ∈ H) subproof {
+      assume(subgroup(H, G, *))
+      have(group(G, *)) by Tautology.from(subgroup.definition)
+      val isSubgroup = have(subgroup(G, G, *)) by Cut(lastStep, groupIsSubgroupOfItself)
+
       // The main idea is to notice that for H = G, identity(G, *) must be in H per [[identityInGroup]]
       // hence identity(G, *) ∉ H would be contradictory
-      have((subgroup(H, G, *) ==> identity(G, *) ∉ H) |- (subgroup(H, G, *) ==> identity(G, *) ∉ H)) by Hypothesis
-      thenHave((subgroup(H, G, *) ==> identity(G, *) ∉ H) |- ∀(H, subgroup(H, G, *) ==> identity(G, *) ∉ H)) by RightForall
+      have(subgroup(H, G, *) /\ identity(G, *) ∉ H |- ()) by Hypothesis
+      thenHave(∀(H, subgroup(H, G, *) ==> identity(G, *) ∉ H) |- subgroup(G, G, *) ==> identity(G, *) ∉ G) by InstantiateForall(G)
+
+      have(∀(H, subgroup(H, G, *) ==> identity(G, *) ∉ H) |- identity(G, *) ∉ G) by Tautology.from(lastStep, isSubgroup)
     }
 
     have(group(G, *) |- isNeutral(identity(G, *), G, *)) by Definition(identity, identityUniqueness)(G, *)
-    have(group(G, *) |- identity(G, *) ∈ G /\ ∀(x, (x ∈ G) ==> ((op(identity(G, *), *, x) === x) /\ (op(identity(G, *), *, e) === x)))) by Tautology.from(isNeutral.definition of (e -> identity(G, *)))*/
-
-    sorry
-  }
+    have(group(G, *) |- identity(G, *) ∈ G /\ ∀(x, (x ∈ G) ==> ((op(identity(G, *), *, x) === x) /\ (op(identity(G, *), *, e) === x)))) by Tautology.from(isNeutral.definition of (e -> identity(G, *)))
+  }*/
 }

From cfa74c883fc70058be8f2ee417164b30b094e063 Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Sun, 4 Jun 2023 00:01:48 +0200
Subject: [PATCH 50/68] Prove subgroup identity

---
 .../scala/lisa/mathematics/GroupTheory.scala  | 91 +++++++++++++++++--
 1 file changed, 83 insertions(+), 8 deletions(-)

diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala
index 87c85cd1..1921c600 100644
--- a/src/main/scala/lisa/mathematics/GroupTheory.scala
+++ b/src/main/scala/lisa/mathematics/GroupTheory.scala
@@ -718,17 +718,92 @@ object GroupTheory extends lisa.Main {
   }
 
   /**
-   * Identity in subgroup --- The identity element of the subgroup is exactly the identity element of the parent group.
+   * Lemma --- If `H` is a subgroup of `G`, then `e_H ∈ G`.
    */
-  /*val subgroupIdentity = Theorem(
+  val subgroupIdentityInParent = Lemma(
+    subgroup(H, G, *) |- identity(H, restrictedFunction(*, cartesianProduct(H, H))) ∈ G 
+  ) {
+    val ★ = restrictedFunction(*, cartesianProduct(H, H))
+    
+    val identityInH = have(subgroup(H, G, *) |- identity(H, ★) ∈ H) by Tautology.from(
+      subgroup.definition,
+      identityInGroup of (G -> H, * -> ★)
+    )
+
+    have(subgroup(H, G, *) |- ∀(x, x ∈ H ==> x ∈ G)) by Tautology.from(
+      subgroup.definition,
+      subset.definition of (x -> H, y -> G)
+    )
+    thenHave(subgroup(H, G, *) |- identity(H, ★) ∈ H ==> identity(H, ★) ∈ G) by InstantiateForall(identity(H, ★))
+    thenHave((subgroup(H, G, *), identity(H, ★) ∈ H) |- identity(H, ★) ∈ G) by Restate
+
+    have(thesis) by Cut(identityInH, lastStep)
+  }
+
+  /**
+   * Identity in subgroup --- The identity element `e_H` of a subgroup `H` of `G` is exactly the identity element `e_G` of
+   * the parent group `(G, *)`.
+   */
+  val subgroupIdentity = Theorem(
     subgroup(H, G, *) |- identity(H, restrictedFunction(*, cartesianProduct(H, H))) === identity(G, *)
   ) {
-    // The crux of the proof relies on proving that identity(G, *) ∈ H
-    // From there we can derive quite easily the theorem using [[subgroupOperation]]
-    have(subgroup(H, G, *) |- identity(G, *) ∈ H) subproof {
-      assume(subgroup(H, G, *))
-      have(group(G, *)) by Tautology.from(subgroup.definition)
-      val isSubgroup = have(subgroup(G, G, *)) by Cut(lastStep, groupIsSubgroupOfItself)
+    val ★ = restrictedFunction(*, cartesianProduct(H, H))
+    val e_G = identity(G, *)
+    val e_H = identity(H, ★)
+
+    val groupG = have(subgroup(H, G, *) |- group(G, *)) by Tautology.from(subgroup.definition)
+    val groupH = have(subgroup(H, G, *) |- group(H, ★)) by Tautology.from(subgroup.definition)
+
+    val subgroupIdentityInH = have(subgroup(H, G, *) |- identity(H, ★) ∈ H) by Tautology.from(
+      subgroup.definition,
+      identityInGroup of (G -> H, * -> ★)
+    )
+
+    have(group(G, *) |- isNeutral(e_G, G, *)) by Definition(identity, identityUniqueness)(G, *)
+    have(group(G, *) |- ∀(x, (x ∈ G) ==> ((op(e_G, *, x) === x) /\ (op(x, *, e_G) === x)))) by Tautology.from(
+      lastStep,
+      isNeutral.definition of (e -> e_G)
+    )
+    thenHave(group(G, *) |- (x ∈ G) ==> ((op(e_G, *, x) === x) /\ (op(x, *, e_G) === x))) by InstantiateForall(x)
+    val neutralDef = thenHave((group(G, *), x ∈ G) |- (op(e_G, *, x) === x) /\ (op(x, *, e_G) === x)) by Restate
+
+    // 1. e_H ★ e_H = e_H
+    val eq1 = have(subgroup(H, G, *) |- op(e_H, ★, e_H) === e_H) subproof {
+      have(group(H, ★) |- (op(e_H, ★, e_H) === e_H)) by Cut(
+        identityInGroup of (G -> H, * -> ★),
+        neutralDef of (G -> H, * -> ★, x -> e_H)
+      )
+
+      have(thesis) by Cut(groupH, lastStep)
+    }
+
+    // 2. e_H * e_H = e_H
+    have(subgroup(H, G, *) |- op(e_H, ★, e_H) === op(e_H, *, e_H)) by Cut(
+      subgroupIdentityInH,
+      subgroupOperation of (x -> e_H, y -> e_H)
+    )
+    val eq2 = have(subgroup(H, G, *) |- op(e_H, *, e_H) === e_H) by Equalities(eq1, lastStep)
+
+    // 3. e_G * e_H = e_H
+    val eq3 = have(subgroup(H, G, *) |- op(e_G, *, e_H) === e_H) subproof {
+      have((group(G, *), e_H ∈ G) |- op(e_G, *, e_H) === e_H) by Tautology.from(neutralDef of (x -> e_H))
+      have((subgroup(H, G, *), e_H ∈ G) |- op(e_G, *, e_H) === e_H) by Cut(groupG, lastStep)
+      have(thesis) by Cut(subgroupIdentityInParent, lastStep)
+    }
+
+    // Conclude by right cancellation
+    val eq4 = have(subgroup(H, G, *) |- op(e_H, *, e_H) === op(e_G, *, e_H)) by Equalities(eq2, eq3)
+    have((group(G, *), e_H ∈ G, e_G ∈ G, op(e_H, *, e_H) === op(e_G, *, e_H)) |- e_H === e_G) by Restate.from(
+      rightCancellation of (x -> e_H, y -> e_H, z -> e_G)
+    )
+    have((subgroup(H, G, *), e_H ∈ G, e_G ∈ G, op(e_H, *, e_H) === op(e_G, *, e_H)) |- e_H === e_G) by Cut(groupG, lastStep)
+    have((subgroup(H, G, *), e_H ∈ G, e_G ∈ G) |- e_H === e_G) by Cut(eq4, lastStep)
+
+    val finalStep = have((subgroup(H, G, *), e_G ∈ G) |- e_H === e_G) by Cut(subgroupIdentityInParent, lastStep)
+
+    have(subgroup(H, G, *) |- e_G ∈ G) by Cut(groupG, identityInGroup)
+    have(thesis) by Cut(lastStep, finalStep)
+  }
 
       // The main idea is to notice that for H = G, identity(G, *) must be in H per [[identityInGroup]]
       // hence identity(G, *) ∉ H would be contradictory

From 90c297276e7b2e7596b0071ce42168abf6bd1766 Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Sun, 4 Jun 2023 00:08:10 +0200
Subject: [PATCH 51/68] Rename binaryFunction to binaryOperation

---
 src/main/scala/lisa/mathematics/GroupTheory.scala | 10 +++++-----
 1 file changed, 5 insertions(+), 5 deletions(-)

diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala
index 1921c600..58095095 100644
--- a/src/main/scala/lisa/mathematics/GroupTheory.scala
+++ b/src/main/scala/lisa/mathematics/GroupTheory.scala
@@ -137,10 +137,10 @@ object GroupTheory extends lisa.Main {
   //
 
   /**
-   * Binary function --- `*` is a binary function on `G` if it associates to each pair of elements of `G`
+   * Binary operation --- `*` is a binary operation on `G` if it associates to each pair of elements of `G`
    * a unique element in `G`. In other words, `*` is a function `G × G -> G`.
    */
-  val binaryFunction = DEF(G, *) --> functionFrom(*, cartesianProduct(G, G), G)
+  val binaryOperation = DEF(G, *) --> functionFrom(*, cartesianProduct(G, G), G)
 
   /**
    * Short-hand alias for `x * y`.
@@ -177,7 +177,7 @@ object GroupTheory extends lisa.Main {
    * Group --- A group (G, *) is a set along with a law of composition `*`, satisfying [[associativityAxiom]], [[identityExistence]]
    * and [[inverseExistence]].
    */
-  val group = DEF(G, *) --> binaryFunction(G, *) /\ associativityAxiom(G, *) /\ identityExistence(G, *) /\ inverseExistence(G, *)
+  val group = DEF(G, *) --> binaryOperation(G, *) /\ associativityAxiom(G, *) /\ identityExistence(G, *) /\ inverseExistence(G, *)
 
   /**
    * Group operation is functional -- The group operation `*` is functional.
@@ -189,7 +189,7 @@ object GroupTheory extends lisa.Main {
   ) {
     have(thesis) by Tautology.from(
       group.definition,
-      binaryFunction.definition,
+      binaryOperation.definition,
       functionFromImpliesFunctional of (f -> *, x -> cartesianProduct(G, G), y -> G)
     )
   }
@@ -204,7 +204,7 @@ object GroupTheory extends lisa.Main {
   ) {
     have(thesis) by Tautology.from(
       group.definition,
-      binaryFunction.definition,
+      binaryOperation.definition,
       functionFromImpliesDomainEq of (f -> *, x -> cartesianProduct(G, G), y -> G)
     )
   }

From e28b3045e14292d026cd56e47f403f743a1d1c7b Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Sun, 4 Jun 2023 15:09:17 +0200
Subject: [PATCH 52/68] Add neutral element lemmas

---
 .../scala/lisa/mathematics/GroupTheory.scala  | 40 +++++++++++++------
 1 file changed, 28 insertions(+), 12 deletions(-)

diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala
index 58095095..783e4435 100644
--- a/src/main/scala/lisa/mathematics/GroupTheory.scala
+++ b/src/main/scala/lisa/mathematics/GroupTheory.scala
@@ -300,6 +300,31 @@ object GroupTheory extends lisa.Main {
    */
   val identity = DEF(G, *) --> TheConditional(e, isNeutral(e, G, *))(identityUniqueness)
 
+  /**
+   * Lemma --- The identity element is neutral by definition.
+   */
+  val identityIsNeutral = Lemma(
+    group(G, *) |- isNeutral(identity(G, *), G, *)
+  ) {
+    have(thesis) by Definition(identity, identityUniqueness)(G, *)
+  }
+
+  /**
+   * Lemma --- For any element `x` in a group `(G, *)`, we have `x * e = e * x = x`.
+   * 
+   * Simpler reformulation of [[identityIsNeutral]].
+   */
+  val identityNeutrality = Lemma(
+    (group(G, *), x ∈ G) |- (op(identity(G, *), *, x) === x) /\ (op(x, *, identity(G, *)) === x)
+  ) {
+    have(group(G, *) |- ∀(x, (x ∈ G) ==> ((op(identity(G, *), *, x) === x) /\ (op(x, *, identity(G, *)) === x)))) by Tautology.from(
+      identityIsNeutral,
+      isNeutral.definition of (e -> identity(G, *))
+    )
+    thenHave(group(G, *) |- (x ∈ G) ==> ((op(identity(G, *), *, x) === x) /\ (op(x, *, identity(G, *)) === x))) by InstantiateForall(x)
+    thenHave(thesis) by Restate
+  }
+
   /**
    * Theorem --- The identity element belongs to the group.
    *
@@ -308,9 +333,8 @@ object GroupTheory extends lisa.Main {
   val identityInGroup = Theorem(
     group(G, *) |- identity(G, *) ∈ G
   ) {
-    have(group(G, *) |- isNeutral(identity(G, *), G, *)) by Definition(identity, identityUniqueness)(G, *)
     have(thesis) by Tautology.from(
-      lastStep,
+      identityIsNeutral,
       isNeutral.definition of (e -> identity(G, *))
     )
   }
@@ -759,19 +783,11 @@ object GroupTheory extends lisa.Main {
       identityInGroup of (G -> H, * -> ★)
     )
 
-    have(group(G, *) |- isNeutral(e_G, G, *)) by Definition(identity, identityUniqueness)(G, *)
-    have(group(G, *) |- ∀(x, (x ∈ G) ==> ((op(e_G, *, x) === x) /\ (op(x, *, e_G) === x)))) by Tautology.from(
-      lastStep,
-      isNeutral.definition of (e -> e_G)
-    )
-    thenHave(group(G, *) |- (x ∈ G) ==> ((op(e_G, *, x) === x) /\ (op(x, *, e_G) === x))) by InstantiateForall(x)
-    val neutralDef = thenHave((group(G, *), x ∈ G) |- (op(e_G, *, x) === x) /\ (op(x, *, e_G) === x)) by Restate
-
     // 1. e_H ★ e_H = e_H
     val eq1 = have(subgroup(H, G, *) |- op(e_H, ★, e_H) === e_H) subproof {
       have(group(H, ★) |- (op(e_H, ★, e_H) === e_H)) by Cut(
         identityInGroup of (G -> H, * -> ★),
-        neutralDef of (G -> H, * -> ★, x -> e_H)
+        identityNeutrality of (G -> H, * -> ★, x -> e_H)
       )
 
       have(thesis) by Cut(groupH, lastStep)
@@ -786,7 +802,7 @@ object GroupTheory extends lisa.Main {
 
     // 3. e_G * e_H = e_H
     val eq3 = have(subgroup(H, G, *) |- op(e_G, *, e_H) === e_H) subproof {
-      have((group(G, *), e_H ∈ G) |- op(e_G, *, e_H) === e_H) by Tautology.from(neutralDef of (x -> e_H))
+      have((group(G, *), e_H ∈ G) |- op(e_G, *, e_H) === e_H) by Tautology.from(identityNeutrality of (x -> e_H))
       have((subgroup(H, G, *), e_H ∈ G) |- op(e_G, *, e_H) === e_H) by Cut(groupG, lastStep)
       have(thesis) by Cut(subgroupIdentityInParent, lastStep)
     }

From 23508ad65531981a697855e50bb6d1ff60a79bc3 Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Sun, 4 Jun 2023 15:10:01 +0200
Subject: [PATCH 53/68] Prove neutral element idempotence

---
 .../scala/lisa/mathematics/GroupTheory.scala  | 33 +++++++++++++++++++
 1 file changed, 33 insertions(+)

diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala
index 783e4435..1e95f727 100644
--- a/src/main/scala/lisa/mathematics/GroupTheory.scala
+++ b/src/main/scala/lisa/mathematics/GroupTheory.scala
@@ -578,6 +578,39 @@ object GroupTheory extends lisa.Main {
     thenHave(thesis) by Restate
   }
 
+  /**
+   * Theorem --- An element `x` of a group `(G, *)` is idempotent if and only if `x` is the neutral element.
+   */
+  val neutralElementIdempotence = Theorem(
+    (group(G, *), x ∈ G) |- (op(x, *, x) === x) <=> (x === identity(G, *))
+  ) {
+    assume(group(G, *))
+    assume(x ∈ G)
+
+    val neutralityEquality = have(op(identity(G, *), *, x) === x) by Tautology.from(identityNeutrality)
+
+    // Forward direction, using the equality x * x = x = e * x
+    // and concluding by right cancellation
+    have(op(x, *, x) === x |- x === identity(G, *)) subproof {
+      have(op(x, *, x) === x |- op(x, *, x) === x) by Hypothesis
+      have(op(x, *, x) === x |- op(x, *, x) === op(identity(G, *), *, x)) by Equalities(lastStep, neutralityEquality)
+      have((op(x, *, x) === x, identity(G, *) ∈ G) |- x === identity(G, *)) by Tautology.from(
+        lastStep,
+        rightCancellation of (x -> x, y -> x, z -> identity(G, *))
+      )
+      have(thesis) by Cut(identityInGroup, lastStep)
+    }
+    val forward = thenHave((op(x, *, x) === x) ==> (x === identity(G, *))) by Restate
+
+    have(x === identity(G, *) |- op(x, *, x) === x) by RightSubstEq(
+      List((x, identity(G, *))),
+      lambda(z, op(z, *, x) === x)
+    )(neutralityEquality)
+    val backward = thenHave((x === identity(G, *)) ==> (op(x, *, x) === x)) by Restate
+
+    have(thesis) by RightIff(forward, backward)
+  }
+
   //
   // 1.1 Subgroups
   //

From 6c8b441cfbc2562d4c3826a34776ce7a3acdaf76 Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Mon, 5 Jun 2023 19:13:59 +0200
Subject: [PATCH 54/68] Add homomorphisms

---
 .../scala/lisa/mathematics/GroupTheory.scala  | 32 ++++++++++++++++---
 1 file changed, 28 insertions(+), 4 deletions(-)

diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala
index 1e95f727..70d2f4d0 100644
--- a/src/main/scala/lisa/mathematics/GroupTheory.scala
+++ b/src/main/scala/lisa/mathematics/GroupTheory.scala
@@ -854,10 +854,34 @@ object GroupTheory extends lisa.Main {
     have(thesis) by Cut(lastStep, finalStep)
   }
 
-      // The main idea is to notice that for H = G, identity(G, *) must be in H per [[identityInGroup]]
-      // hence identity(G, *) ∉ H would be contradictory
-      have(subgroup(H, G, *) /\ identity(G, *) ∉ H |- ()) by Hypothesis
-      thenHave(∀(H, subgroup(H, G, *) ==> identity(G, *) ∉ H) |- subgroup(G, G, *) ==> identity(G, *) ∉ G) by InstantiateForall(G)
+  /**
+   * 2. Homomorphisms
+   */
+
+  // Extra group composition law
+  val ★ = variable
+
+  /**
+   * Definition --- A group homomorphism is a mapping `f: G -> H` from structures `G` and `H` equipped with binary operations `*` and `★` respectively,
+   * such that for all `x, y ∈ G`, we have* `f(x * y) = f(x) ★ f(y)`.
+   * 
+   * In the following, "homomorphism" always stands for "group homomorphism", i.e. `(G, *)` and `(H, ★)` are groups.
+   */
+  val homomorphism = DEF(f, G, *, H, ★) --> group(G, *) /\ group(H, ★) /\ functionFrom(f, G, H) /\ ∀(x, x ∈ G ==> ∀(y, y ∈ G ==> (app(f, op(x, *, y)) === op(app(f, x), ★, app(f, y)))))
+
+  /**
+   * Lemma --- Practical reformulation of the homomorphism definition.
+   */
+  val homomorphismApplication = Lemma(
+    (homomorphism(f, G, *, H, ★), x ∈ G, y ∈ G) |- app(f, op(x, *, y)) === op(app(f, x), ★, app(f, y))
+  ) {
+    assume(homomorphism(f, G, *, H, ★))
+    have(∀(x, x ∈ G ==> ∀(y, y ∈ G ==> (app(f, op(x, *, y)) === op(app(f, x), ★, app(f, y)))))) by Tautology.from(homomorphism.definition)
+    thenHave(x ∈ G ==> ∀(y, y ∈ G ==> (app(f, op(x, *, y)) === op(app(f, x), ★, app(f, y))))) by InstantiateForall(x)
+    thenHave((x ∈ G) |- ∀(y, y ∈ G ==> (app(f, op(x, *, y)) === op(app(f, x), ★, app(f, y))))) by Restate
+    thenHave((x ∈ G) |- y ∈ G ==> (app(f, op(x, *, y)) === op(app(f, x), ★, app(f, y)))) by InstantiateForall(y)
+    thenHave(thesis) by Restate
+  }
 
       have(∀(H, subgroup(H, G, *) ==> identity(G, *) ∉ H) |- identity(G, *) ∉ G) by Tautology.from(lastStep, isSubgroup)
     }

From a70cbd3d0600c4856523191588b0f0cc4b2b437f Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Mon, 5 Jun 2023 19:14:35 +0200
Subject: [PATCH 55/68] Prove homomorphism maps e_G to e_H

---
 .../scala/lisa/mathematics/GroupTheory.scala  | 57 ++++++++++++++++++-
 .../scala/lisa/mathematics/SetTheory.scala    | 34 +++++++++++
 2 files changed, 90 insertions(+), 1 deletion(-)

diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala
index 70d2f4d0..f6fa23ad 100644
--- a/src/main/scala/lisa/mathematics/GroupTheory.scala
+++ b/src/main/scala/lisa/mathematics/GroupTheory.scala
@@ -883,7 +883,62 @@ object GroupTheory extends lisa.Main {
     thenHave(thesis) by Restate
   }
 
-      have(∀(H, subgroup(H, G, *) ==> identity(G, *) ∉ H) |- identity(G, *) ∉ G) by Tautology.from(lastStep, isSubgroup)
+  /**
+   * Theorem --- If `f` is a group-homomorphism between `G` and `H`, then `f(e_G) = e_H`.
+   */
+  val homomorphismMapsIdentityToIdentity = Theorem(
+    homomorphism(f, G, *, H, ★) |- app(f, identity(G, *)) === identity(H, ★)
+  ) {
+    val e = identity(G, *)
+
+    val groupG = have(homomorphism(f, G, *, H, ★) |- group(G, *)) by Tautology.from(homomorphism.definition)
+    val groupH = have(homomorphism(f, G, *, H, ★) |- group(H, ★)) by Tautology.from(homomorphism.definition)
+    val functionF = have(homomorphism(f, G, *, H, ★) |- functionFrom(f, G, H)) by Tautology.from(homomorphism.definition)
+
+    val identityInG = have(homomorphism(f, G, *, H, ★) |- e ∈ G) by Cut(groupG, identityInGroup)
+
+    have((homomorphism(f, G, *, H, ★), e ∈ G) |- app(f, e) ∈ H) by Cut(
+      functionF,
+      functionAppInCodomain of (VariableLabel("t") -> e, x -> G, y -> H)
+    )
+    val appInH = have(homomorphism(f, G, *, H, ★) |- app(f, e) ∈ H) by Cut(identityInG, lastStep)
+
+    // 0. e * e = e (to apply substitution)
+    have(group(G, *) |- op(e, *, e) === e) by Cut(
+      identityInGroup,
+      neutralElementIdempotence of (x -> e)
+    )
+    val eq0 = have(homomorphism(f, G, *, H, ★) |- op(e, *, e) === e) by Cut(groupG, lastStep)
+    
+    // 1. f(e * e) = f(e)
+    have(app(f, e) === app(f, e)) by RightRefl
+    thenHave(op(e, *, e) === e |- app(f, op(e, *, e)) === app(f, e)) by RightSubstEq(
+      List((op(e, *, e), e)),
+      lambda(z, app(f, z) === app(f, e))
+    )
+    val eq1 = have(homomorphism(f, G, *, H, ★) |- app(f, op(e, *, e)) === app(f, e)) by Cut(eq0, lastStep)
+
+    // 2. f(e * e) = f(e) ★ f(e)
+    val eq2 = have(homomorphism(f, G, *, H, ★) |- app(f, op(e, *, e)) === op(app(f, e), ★, app(f, e))) by Cut(
+      identityInG,
+      homomorphismApplication of (x -> e, y -> e)
+    )
+
+    // 3. f(e) ★ f(e) = f(e)
+    val eq3 = have(homomorphism(f, G, *, H, ★) |- op(app(f, e), ★, app(f, e)) === app(f, e)) by Equalities(eq1, eq2)
+
+    // Conclude by idempotence
+    have((homomorphism(f, G, *, H, ★), app(f, e) ∈ H) |- (op(app(f, e), ★, app(f, e)) === app(f, e)) <=> (app(f, e) === identity(H, ★))) by Cut(
+      groupH,
+      neutralElementIdempotence of (x -> app(f, e), G -> H, * -> ★)
+    )
+    have(homomorphism(f, G, *, H, ★)|- (op(app(f, e), ★, app(f, e)) === app(f, e)) <=> (app(f, e) === identity(H, ★))) by Cut(
+      appInH,
+      lastStep
+    )
+    
+    have(thesis) by Tautology.from(lastStep, eq3)
+  }
     }
 
     have(group(G, *) |- isNeutral(identity(G, *), G, *)) by Definition(identity, identityUniqueness)(G, *)
diff --git a/src/main/scala/lisa/mathematics/SetTheory.scala b/src/main/scala/lisa/mathematics/SetTheory.scala
index 11257644..9af0cba5 100644
--- a/src/main/scala/lisa/mathematics/SetTheory.scala
+++ b/src/main/scala/lisa/mathematics/SetTheory.scala
@@ -2882,6 +2882,40 @@ object SetTheory extends lisa.Main {
     have(thesis) by Cut(surjfunc, funceq)
   }
 
+  /**
+   * Lemma --- If `f: x -> y` is a function and `z ∈ x`, then `f(z) ∈ y`.
+   */
+  val functionAppInCodomain = Lemma(
+    (functionFrom(f, x, y), in(t, x)) |- in(app(f, t), y)
+  ) {
+    have((functional(f), in(t, relationDomain(f))) |- in(pair(t, app(f, t)), f)) by Definition(app, functionApplicationUniqueness)(f, t)
+    have((functionFrom(f, x, y), in(t, relationDomain(f))) |- in(pair(t, app(f, t)), f)) by Cut(functionFromImpliesFunctional, lastStep)
+    thenHave((functionFrom(f, x, y), relationDomain(f) === x, in(t, x)) |- in(pair(t, app(f, t)), f)) by LeftSubstEq(
+      List((relationDomain(f), x)),
+      lambda(z, in(t, z))
+    )
+    val appDef = have((functionFrom(f, x, y), in(t, x)) |- in(pair(t, app(f, t)), f)) by Cut(functionFromImpliesDomainEq, lastStep)
+
+    have(∀(t, in(t, setOfFunctions(x, y)) <=> (in(t, powerSet(cartesianProduct(x, y))) /\ functionalOver(t, x)))) by Definition(setOfFunctions, setOfFunctionsUniqueness)(x, y)
+    thenHave(in(f, setOfFunctions(x, y)) <=> (in(f, powerSet(cartesianProduct(x, y))) /\ functionalOver(f, x))) by InstantiateForall(f)
+
+    have(functionFrom(f, x, y) |- ∀(t, in(t, f) ==> in(t, cartesianProduct(x, y)))) by Tautology.from(
+      lastStep,
+      functionFrom.definition,
+      powerSet.definition of (x -> f, y -> cartesianProduct(x, y)),
+      subset.definition of (x -> f, y -> cartesianProduct(x, y))
+    )
+    thenHave(functionFrom(f, x, y) |- in(pair(t, app(f, t)), f) ==> in(pair(t, app(f, t)), cartesianProduct(x, y))) by InstantiateForall(pair(t, app(f, t)))
+    thenHave((functionFrom(f, x, y), in(pair(t, app(f, t)), f)) |- in(pair(t, app(f, t)), cartesianProduct(x, y))) by Restate
+
+    have((functionFrom(f, x, y), in(pair(t, app(f, t)), f)) |- in(app(f, t), y)) by Tautology.from(
+      lastStep,
+      pairInCartesianProduct of (a -> t, b -> app(f, t))
+    )
+
+    have(thesis) by Cut(appDef, lastStep)
+  }
+
   /**
    * Theorem --- Cantor's Theorem
    *

From 51b47b5ed5ef8a11ea2d4036fbe0805502cc6410 Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Tue, 6 Jun 2023 16:07:03 +0200
Subject: [PATCH 56/68] Clean GroupTheory and add abelian group definition

---
 .../scala/lisa/mathematics/GroupTheory.scala  | 183 ++++++++++--------
 1 file changed, 102 insertions(+), 81 deletions(-)

diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala
index f6fa23ad..5d0bbf9f 100644
--- a/src/main/scala/lisa/mathematics/GroupTheory.scala
+++ b/src/main/scala/lisa/mathematics/GroupTheory.scala
@@ -179,10 +179,66 @@ object GroupTheory extends lisa.Main {
    */
   val group = DEF(G, *) --> binaryOperation(G, *) /\ associativityAxiom(G, *) /\ identityExistence(G, *) /\ inverseExistence(G, *)
 
+  /**
+   * Commutativity --- `*` is said to be commutative on `G` if `x * y = y * x` for all `x, y ∈ G`.
+   */
+  val commutativityAxiom = DEF(G, *) --> ∀(x, x ∈ G ==> ∀(y, y ∈ G ==> (op(x, *, y) === op(y, *, x))))
+
+  /**
+   * Abelian group --- A group is said to be abelian (or commutative) if every element commutes,
+   * i.e. it satisfies [[commutativityAxiom]].
+   */
+  val abelianGroup = DEF(G, *) --> group(G, *) /\ commutativityAxiom(G, *)
+
+  /**
+   * Alias for abelian group.
+   */
+  val commutativeGroup = abelianGroup
+
+  /**
+   * Lemma --- For elements `x, y, z` in a group `(G, *)`, we have `(xy)z = x(yz)`.
+   * 
+   * Practical reformulation of the [[associativityAxiom]].
+   */
+  val associativity = Lemma(
+    (group(G, *), x ∈ G, y ∈ G, z ∈ G) |- op(op(x, *, y), *, z) === op(x, *, op(y, *, z))
+  ) {
+    assume(group(G, *))
+
+    have(∀(x, x ∈ G ==> ∀(y, y ∈ G ==> ∀(z, z ∈ G ==> (op(op(x, *, y), *, z) === op(x, *, op(y, *, z))))))) by Tautology.from(
+      group.definition,
+      associativityAxiom.definition
+    )
+    thenHave(x ∈ G ==> ∀(y, y ∈ G ==> ∀(z, z ∈ G ==> (op(op(x, *, y), *, z) === op(x, *, op(y, *, z)))))) by InstantiateForall(x)
+    thenHave(x ∈ G |- ∀(y, y ∈ G ==> ∀(z, z ∈ G ==> (op(op(x, *, y), *, z) === op(x, *, op(y, *, z)))))) by Restate
+    thenHave(x ∈ G |- y ∈ G ==> ∀(z, z ∈ G ==> (op(op(x, *, y), *, z) === op(x, *, op(y, *, z))))) by InstantiateForall(y)
+    thenHave((x ∈ G, y ∈ G) |- ∀(z, z ∈ G ==> (op(op(x, *, y), *, z) === op(x, *, op(y, *, z))))) by Restate
+    thenHave((x ∈ G, y ∈ G) |- z ∈ G ==> (op(op(x, *, y), *, z) === op(x, *, op(y, *, z)))) by InstantiateForall(z)
+    thenHave((x ∈ G, y ∈ G, z ∈ G) |- (op(op(x, *, y), *, z) === op(x, *, op(y, *, z)))) by Restate
+  }
+
+  /**
+   * Lemma --- For elements `x, y` in an abelian group `(G, *)`, we have `xy = yx`.
+   * 
+   * Practical reformulation of [[commutativityAxiom]].
+   */
+  val commutativity = Lemma(
+    (abelianGroup(G, *), x ∈ G, y ∈ G) |- op(x, *, y) === op(y, *, x)
+  ) {
+    assume(abelianGroup(G, *))
+    
+    have(∀(x, x ∈ G ==> ∀(y, y ∈ G ==> (op(x, *, y) === op(y, *, x))))) by Tautology.from(
+      abelianGroup.definition,
+      commutativityAxiom.definition
+    )
+    thenHave(x ∈ G ==> ∀(y, y ∈ G ==> (op(x, *, y) === op(y, *, x)))) by InstantiateForall(x)
+    thenHave(x ∈ G |- ∀(y, y ∈ G ==> (op(x, *, y) === op(y, *, x)))) by Restate
+    thenHave(x ∈ G |- (y ∈ G ==> (op(x, *, y) === op(y, *, x)))) by InstantiateForall(y)
+    thenHave((x ∈ G, y ∈ G) |- ((op(x, *, y) === op(y, *, x)))) by Restate
+  }
+
   /**
    * Group operation is functional -- The group operation `*` is functional.
-   *
-   * Similarly to the lemma below, follows more or less directly from the definitions, but still useful.
    */
   val groupOperationIsFunctional = Lemma(
     group(G, *) |- functional(*)
@@ -195,9 +251,9 @@ object GroupTheory extends lisa.Main {
   }
 
   /**
-   * Group operation domain -- The domain of a group law is the cartesian product of the group with itself.
+   * Group operation domain -- The domain of a group law is the cartesian product of the group `G` with itself.
    *
-   * Follows directly from the definition of `binaryRelation`, but it is a useful lemma to have in some proofs.
+   * Follows directly from the definition of `binaryRelation`.
    */
   val groupOperationDomain = Lemma(
     group(G, *) |- relationDomain(*) === cartesianProduct(G, G)
@@ -232,28 +288,6 @@ object GroupTheory extends lisa.Main {
     have(thesis) by Cut(groupOperationDomain, lastStep)
   }
 
-  /**
-   * Lemma --- For elements `x, y, z` in a group `(G, *)`, we have `(xy)z = x(yz)`.
-   * 
-   * Simpler reformulation of the [[associativityAxiom]].
-   */
-  val associativity = Lemma(
-    (group(G, *), x ∈ G, y ∈ G, z ∈ G) |- op(op(x, *, y), *, z) === op(x, *, op(y, *, z))
-  ) {
-    assume(group(G, *))
-
-    have(∀(x, x ∈ G ==> ∀(y, y ∈ G ==> ∀(z, z ∈ G ==> (op(op(x, *, y), *, z) === op(x, *, op(y, *, z))))))) by Tautology.from(
-      group.definition,
-      associativityAxiom.definition
-    )
-    thenHave(x ∈ G ==> ∀(y, y ∈ G ==> ∀(z, z ∈ G ==> (op(op(x, *, y), *, z) === op(x, *, op(y, *, z)))))) by InstantiateForall(x)
-    thenHave(x ∈ G |- ∀(y, y ∈ G ==> ∀(z, z ∈ G ==> (op(op(x, *, y), *, z) === op(x, *, op(y, *, z)))))) by Restate
-    thenHave(x ∈ G |- y ∈ G ==> ∀(z, z ∈ G ==> (op(op(x, *, y), *, z) === op(x, *, op(y, *, z))))) by InstantiateForall(y)
-    thenHave((x ∈ G, y ∈ G) |- ∀(z, z ∈ G ==> (op(op(x, *, y), *, z) === op(x, *, op(y, *, z))))) by Restate
-    thenHave((x ∈ G, y ∈ G) |- z ∈ G ==> (op(op(x, *, y), *, z) === op(x, *, op(y, *, z)))) by InstantiateForall(z)
-    thenHave((x ∈ G, y ∈ G, z ∈ G) |- (op(op(x, *, y), *, z) === op(x, *, op(y, *, z)))) by Restate
-  }
-
   /**
    * Identity uniqueness --- In a group (G, *), an identity element is unique, i.e. if both `e * x = x * e = x` and
    * `f * x = x * f = x` for all `x`, then `e = f`.
@@ -283,9 +317,7 @@ object GroupTheory extends lisa.Main {
       val firstEq = thenHave(op(e, *, f) === f) by Tautology
 
       // 2. ef = e
-      have(e ∈ G |- ((op(f, *, e) === e) /\ (op(e, *, f) === e))) by InstFunSchema(Map(e -> f, f -> e))(neutrality)
-
-      have((op(f, *, e) === e) /\ (op(e, *, f) === e)) by Cut(membership of (e -> e), lastStep)
+      have((op(f, *, e) === e) /\ (op(e, *, f) === e)) by Cut(membership of (e -> e), neutrality of (e -> f, f -> e))
       val secondEq = thenHave(e === op(e, *, f)) by Tautology
 
       // 3. Conclude by transitivity
@@ -303,7 +335,7 @@ object GroupTheory extends lisa.Main {
   /**
    * Lemma --- The identity element is neutral by definition.
    */
-  val identityIsNeutral = Lemma(
+  private val identityIsNeutral = Lemma(
     group(G, *) |- isNeutral(identity(G, *), G, *)
   ) {
     have(thesis) by Definition(identity, identityUniqueness)(G, *)
@@ -312,7 +344,7 @@ object GroupTheory extends lisa.Main {
   /**
    * Lemma --- For any element `x` in a group `(G, *)`, we have `x * e = e * x = x`.
    * 
-   * Simpler reformulation of [[identityIsNeutral]].
+   * Practical reformulation of [[identityIsNeutral]].
    */
   val identityNeutrality = Lemma(
     (group(G, *), x ∈ G) |- (op(identity(G, *), *, x) === x) /\ (op(x, *, identity(G, *)) === x)
@@ -328,7 +360,7 @@ object GroupTheory extends lisa.Main {
   /**
    * Theorem --- The identity element belongs to the group.
    *
-   * This might seem like a silly theorem, but it is useful in [[groupNonEmpty]].
+   * This simple theorem has unexpected consequences, such as [[groupNonEmpty]].
    */
   val identityInGroup = Theorem(
     group(G, *) |- identity(G, *) ∈ G
@@ -418,7 +450,7 @@ object GroupTheory extends lisa.Main {
   /**
    * Lemma --- The inverse of `x` is an inverse of `x` (by definition).
    */
-  val inverseIsInverse = Lemma(
+  private val inverseIsInverse = Lemma(
     (group(G, *), x ∈ G) |- isInverse(inverse(x, G, *), x, G, *)
   ) {
     have(thesis) by Definition(inverse, inverseUniqueness)(x, G, *)
@@ -427,7 +459,7 @@ object GroupTheory extends lisa.Main {
   /**
    * Lemma --- The inverse element `y` of `x` is in `G`.
    */
-  val inverseIsInGroup = Lemma(
+  val inverseInGroup = Lemma(
     (group(G, *), x ∈ G) |- inverse(x, G, *) ∈ G
   ) {
     have(thesis) by Tautology.from(
@@ -440,58 +472,45 @@ object GroupTheory extends lisa.Main {
    * Theorem --- `y` is the inverse of `x` iff `x` is the inverse of `y`
    */
   val inverseSymmetry = Theorem(
-    group(G, *) |- ∀(x, x ∈ G ==> (isInverse(y, x, G, *) ==> isInverse(x, y, G, *)))
+    (group(G, *), x ∈ G) |- (isInverse(y, x, G, *) ==> isInverse(x, y, G, *))
   ) {
-    have((group(G, *), x ∈ G, isInverse(y, x, G, *)) |- isInverse(x, y, G, *)) subproof {
-      assume(group(G, *))
-      assume(x ∈ G)
-      assume(isInverse(y, x, G, *))
+    assume(group(G, *))
+    assume(x ∈ G)
 
-      have(y ∈ G /\ isNeutral(op(x, *, y), G, *) /\ isNeutral(op(y, *, x), G, *)) by Tautology.from(isInverse.definition)
-      thenHave(isNeutral(op(y, *, x), G, *) /\ isNeutral(op(x, *, y), G, *)) by Tautology
-      thenHave(x ∈ G /\ isNeutral(op(y, *, x), G, *) /\ isNeutral(op(x, *, y), G, *)) by Tautology
+    have(isInverse(y, x, G, *) |- y ∈ G /\ isNeutral(op(x, *, y), G, *) /\ isNeutral(op(y, *, x), G, *)) by Tautology.from(isInverse.definition)
+    thenHave(isInverse(y, x, G, *) |- isNeutral(op(y, *, x), G, *) /\ isNeutral(op(x, *, y), G, *)) by Tautology
+    thenHave(isInverse(y, x, G, *) |- x ∈ G /\ isNeutral(op(y, *, x), G, *) /\ isNeutral(op(x, *, y), G, *)) by Tautology
 
-      have(thesis) by Tautology.from(lastStep, isInverse.definition of (y -> x, x -> y))
-    }
-    thenHave(group(G, *) |- x ∈ G ==> (isInverse(y, x, G, *) ==> isInverse(x, y, G, *))) by Restate
-    thenHave(thesis) by RightForall
+    have(isInverse(y, x, G, *) |- isInverse(x, y, G, *)) by Tautology.from(lastStep, isInverse.definition of (y -> x, x -> y))
+    thenHave(thesis) by Restate
   }
 
   /**
    * Involution of the inverse -- For all `x`, `inverse(inverse(x)) = x`.
    */
   val inverseIsInvolutive = Theorem(
-    group(G, *) |- ∀(x, x ∈ G ==> (inverse(inverse(x, G, *), G, *) === x))
+    (group(G, *), x ∈ G) |- (inverse(inverse(x, G, *), G, *) === x)
   ) {
     assume(group(G, *))
 
-    val inverseCharacterization = have((group(G, *), x ∈ G) |- ((y === inverse(x, G, *)) <=> isInverse(y, x, G, *))) subproof {
-      assume(group(G, *))
-
+    val inverseCharacterization = have(x ∈ G |- ((y === inverse(x, G, *)) <=> isInverse(y, x, G, *))) subproof {
       have(x ∈ G |- ∀(y, (y === inverse(x, G, *)) <=> isInverse(y, x, G, *))) by Tautology.from(inverseUniqueness, inverse.definition)
       thenHave(thesis) by InstantiateForall(y)
     }
 
     // Use the symmetry of the inverse relation to prove that x is an inverse of inverse(x)
-    val satisfaction = have((group(G, *), x ∈ G) |- isInverse(x, inverse(x, G, *), G, *)) subproof {
-      assume(group(G, *))
-
-      have(∀(z, z ∈ G ==> (isInverse(inverse(x, G, *), z, G, *) ==> isInverse(z, inverse(x, G, *), G, *)))) by Tautology.from(inverseSymmetry of (y -> inverse(x, G, *)))
-      thenHave(x ∈ G ==> (isInverse(inverse(x, G, *), x, G, *) ==> isInverse(x, inverse(x, G, *), G, *))) by InstantiateForall(x)
-      val symmetryCut = thenHave((x ∈ G, isInverse(inverse(x, G, *), x, G, *)) |- isInverse(x, inverse(x, G, *), G, *)) by Restate
-
-      have(thesis) by Cut(inverseIsInverse, symmetryCut)
+    val satisfaction = have(x ∈ G |- isInverse(x, inverse(x, G, *), G, *)) subproof {
+      have((x ∈ G, isInverse(inverse(x, G, *), x, G, *)) |- isInverse(x, inverse(x, G, *), G, *)) by Restate.from(inverseSymmetry of (y -> inverse(x, G, *)))
+      have(thesis) by Cut(inverseIsInverse, lastStep)
     }
 
     // Conclude by uniqueness
     val u = inverse(inverse(x, G, *), G, *)
 
-    val characterization = have(in(inverse(x, G, *), G) |- ((x === u) <=> isInverse(x, inverse(x, G, *), G, *))) by Restate.from(inverseCharacterization of (x -> inverse(x, G, *), y -> x))
-    val eq = have((x ∈ G, in(inverse(x, G, *), G)) |- (x === u)) by Tautology.from(satisfaction, characterization)
+    val characterization = have(inverse(x, G, *) ∈ G |- ((x === u) <=> isInverse(x, inverse(x, G, *), G, *))) by Restate.from(inverseCharacterization of (x -> inverse(x, G, *), y -> x))
+    val eq = have((x ∈ G, inverse(x, G, *) ∈ G) |- (x === u)) by Tautology.from(satisfaction, characterization)
 
-    have(x ∈ G |- (x === u)) by Cut(inverseIsInGroup, eq)
-    thenHave(x ∈ G ==> (x === u)) by Restate
-    thenHave(thesis) by RightForall
+    have(thesis) by Cut(inverseInGroup, eq)
   }
 
   /**
@@ -506,7 +525,7 @@ object GroupTheory extends lisa.Main {
     val cancellation = have((group(G, *), x ∈ G, y ∈ G) |- op(i, *, op(x, *, y)) === y) subproof {
       // (ix)y = i(xy)
       val eq1 = have((group(G, *), x ∈ G, y ∈ G) |- op(op(i, *, x), *, y) === op(i, *, op(x, *, y))) by Cut(
-        inverseIsInGroup,
+        inverseInGroup,
         associativity of (x -> i, y -> x, z -> y)
       )
       
@@ -549,7 +568,7 @@ object GroupTheory extends lisa.Main {
     val cancellation = have((group(G, *), x ∈ G, y ∈ G) |- op(op(y, *, x), *, i) === y) subproof {
       // (xy)i = y(xi)
       val eq1 = have((group(G, *), x ∈ G, y ∈ G) |- op(op(y, *, x), *, i) === op(y, *, op(x, *, i))) by Cut(
-        inverseIsInGroup,
+        inverseInGroup,
         associativity of (x -> y, y -> x, z -> i)
       )
       
@@ -579,9 +598,9 @@ object GroupTheory extends lisa.Main {
   }
 
   /**
-   * Theorem --- An element `x` of a group `(G, *)` is idempotent if and only if `x` is the neutral element.
+   * Theorem --- An element `x` of a group `(G, *)` is idempotent if and only if `x` is the identity element.
    */
-  val neutralElementIdempotence = Theorem(
+  val identityIdempotence = Theorem(
     (group(G, *), x ∈ G) |- (op(x, *, x) === x) <=> (x === identity(G, *))
   ) {
     assume(group(G, *))
@@ -611,8 +630,11 @@ object GroupTheory extends lisa.Main {
     have(thesis) by RightIff(forward, backward)
   }
 
+  // TODO Direct product group
+  // TODO Permutation group
+
   //
-  // 1.1 Subgroups
+  // 2. Subgroups
   //
 
   /**
@@ -763,12 +785,12 @@ object GroupTheory extends lisa.Main {
 
     have(!(functional(*) /\ pair(x, y) ∈ relationDomain(*)) |- ()) by LeftNot(premises)
     thenHave(!functional(*) \/ !(pair(x, y) ∈ relationDomain(*)) |- ()) by Restate
-    thenHave(!functional(*) \/ !(pair(x, y) ∈ relationDomain(*)) |- (r === emptySet())) by Weakening
-    val neg = thenHave((!functional(*) \/ !(pair(x, y) ∈ relationDomain(*))) ==> (r === emptySet())) by Restate
+    thenHave(!functional(*) \/ !(pair(x, y) ∈ relationDomain(*)) |- (r === ∅)) by Weakening
+    val neg = thenHave((!functional(*) \/ !(pair(x, y) ∈ relationDomain(*))) ==> (r === ∅)) by Restate
 
     have(
       ((functional(*) /\ pair(x, y) ∈ relationDomain(*)) ==> pair(pair(x, y), r) ∈ *) /\
-        ((!functional(*) \/ !(pair(x, y) ∈ relationDomain(*))) ==> (r === emptySet()))
+        ((!functional(*) \/ !(pair(x, y) ∈ relationDomain(*))) ==> (r === ∅))
     ) by RightAnd(pos, neg)
 
     have(thesis) by Tautology.from(lastStep, characterization)
@@ -854,9 +876,13 @@ object GroupTheory extends lisa.Main {
     have(thesis) by Cut(lastStep, finalStep)
   }
 
-  /**
-   * 2. Homomorphisms
-   */
+  // TODO Subgroup inverse
+  // TODO Subgroup condition
+  // TODO Trivial subgroup
+
+  //
+  // 3. Homomorphisms
+  //
 
   // Extra group composition law
   val ★ = variable
@@ -906,7 +932,7 @@ object GroupTheory extends lisa.Main {
     // 0. e * e = e (to apply substitution)
     have(group(G, *) |- op(e, *, e) === e) by Cut(
       identityInGroup,
-      neutralElementIdempotence of (x -> e)
+      identityIdempotence of (x -> e)
     )
     val eq0 = have(homomorphism(f, G, *, H, ★) |- op(e, *, e) === e) by Cut(groupG, lastStep)
     
@@ -930,7 +956,7 @@ object GroupTheory extends lisa.Main {
     // Conclude by idempotence
     have((homomorphism(f, G, *, H, ★), app(f, e) ∈ H) |- (op(app(f, e), ★, app(f, e)) === app(f, e)) <=> (app(f, e) === identity(H, ★))) by Cut(
       groupH,
-      neutralElementIdempotence of (x -> app(f, e), G -> H, * -> ★)
+      identityIdempotence of (x -> app(f, e), G -> H, * -> ★)
     )
     have(homomorphism(f, G, *, H, ★)|- (op(app(f, e), ★, app(f, e)) === app(f, e)) <=> (app(f, e) === identity(H, ★))) by Cut(
       appInH,
@@ -939,9 +965,4 @@ object GroupTheory extends lisa.Main {
     
     have(thesis) by Tautology.from(lastStep, eq3)
   }
-    }
-
-    have(group(G, *) |- isNeutral(identity(G, *), G, *)) by Definition(identity, identityUniqueness)(G, *)
-    have(group(G, *) |- identity(G, *) ∈ G /\ ∀(x, (x ∈ G) ==> ((op(identity(G, *), *, x) === x) /\ (op(identity(G, *), *, e) === x)))) by Tautology.from(isNeutral.definition of (e -> identity(G, *)))
-  }*/
 }

From 5764ce734ea8e7ea164e10600899c3c1b7dd0bb4 Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Tue, 6 Jun 2023 17:42:17 +0200
Subject: [PATCH 57/68] Define kernel of homomorphism

---
 .../scala/lisa/mathematics/GroupTheory.scala  | 25 ++++++++++++++++++-
 1 file changed, 24 insertions(+), 1 deletion(-)

diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala
index 5d0bbf9f..43667de2 100644
--- a/src/main/scala/lisa/mathematics/GroupTheory.scala
+++ b/src/main/scala/lisa/mathematics/GroupTheory.scala
@@ -5,6 +5,7 @@ import lisa.automation.kernel.CommonTactics.Definition
 import lisa.automation.kernel.CommonTactics.Equalities
 import lisa.automation.kernel.CommonTactics.ExistenceAndUniqueness
 import lisa.automation.kernel.OLPropositionalSolver.Tautology
+import lisa.automation.settheory.SetTheoryTactics.UniqueComprehension
 import lisa.mathematics.FirstOrderLogic.equalityTransitivity
 import lisa.mathematics.FirstOrderLogic.existsOneImpliesExists
 import lisa.mathematics.FirstOrderLogic.substitutionInUniquenessQuantifier
@@ -26,7 +27,7 @@ object GroupTheory extends lisa.Main {
   // Group elements
   private val a, b = variable
   private val x, y, z = variable
-  private val u, v, w = variable
+  private val t, u, v, w = variable
 
   // Identity elements
   private val e, f = variable
@@ -965,4 +966,26 @@ object GroupTheory extends lisa.Main {
     
     have(thesis) by Tautology.from(lastStep, eq3)
   }
+
+  /**
+   * Kernel uniqueness --- The kernel of a homomorphism is well-defined.
+   */
+  val kernelUniqueness = Theorem(
+    homomorphism(f, G, *, H, ★) |- ∃!(z, ∀(t, (t ∈ z) <=> (t ∈ G /\ (app(f, t) === identity(H, ★)))))
+  ) {
+    // We apply the comprehension axiom here.
+    // It might seem odd that the homomorphism assumption is not needed for the set to be defined,
+    // but remember that [[app]] and [[identity]] default to the empty set when the assumptions are not met.
+    // We add the assumption of `f` being a homomorphism to discard any value when the assumptions do not hold.
+    have(∃!(z, ∀(t, (t ∈ z) <=> (t ∈ G /\ (app(f, t) === identity(H, ★)))))) by UniqueComprehension(
+      G,
+      lambda(Seq(t, G), app(f, t) === identity(H, ★))
+    )
+    thenHave(thesis) by Weakening
+  }
+
+  /**
+   * Kernel --- The kernel of a homomorphism `f: G -> H` is the set of elements `t ∈ G` such that `f(t) = e_H`.
+   */
+  val kernel = DEF(f, G, *, H, ★) --> TheConditional(z, ∀(t, (t ∈ z) <=> (t ∈ G /\ (app(f, t) === identity(H, ★)))))(kernelUniqueness)
 }

From 2c5f2e794f825d4a57cbfbd7604dcc5f86620929 Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Fri, 9 Jun 2023 15:27:38 +0200
Subject: [PATCH 58/68] Fix ExistenceAndUniqueness RHS

---
 .../lisa/automation/kernel/CommonTactics.scala      | 13 +++++++------
 1 file changed, 7 insertions(+), 6 deletions(-)

diff --git a/src/main/scala/lisa/automation/kernel/CommonTactics.scala b/src/main/scala/lisa/automation/kernel/CommonTactics.scala
index 609580d4..a08b5c58 100644
--- a/src/main/scala/lisa/automation/kernel/CommonTactics.scala
+++ b/src/main/scala/lisa/automation/kernel/CommonTactics.scala
@@ -60,6 +60,7 @@ object CommonTactics {
         proof.InvalidProofTactic("Could not infer correct right pivots.")
       } else {
         val gammaPrime = uniquenessSeq.left.filter(f => !isSame(f, phi) && !isSame(f, substPhi))
+        val deltaPrime = uniquenessSeq.right.filter(f => !isSame(f, (x === y)) && !isSame(f, (y === x)))
 
         TacticSubproof {
           // There's got to be a better way of importing have/thenHave/assume methods
@@ -75,13 +76,13 @@ object CommonTactics {
             lib.assume(f)
           }
 
-          val backward = lib.have(phi |- (substPhi ==> (x === y))) by Restate.from(uniqueness)
+          val backward = lib.have(phi |- (deltaPrime + (substPhi ==> (x === y)))) by Restate.from(uniqueness)
 
-          lib.have(phi |- ((x === y) <=> substPhi)) by RightIff(forward, backward)
-          lib.thenHave(phi |- ∀(y, (x === y) <=> substPhi)) by RightForall
-          lib.thenHave(phi |- ∃(x, ∀(y, (x === y) <=> substPhi))) by RightExists
-          lib.thenHave(∃(x, phi) |- ∃(x, ∀(y, (x === y) <=> substPhi))) by LeftExists
-          lib.thenHave(∃(x, phi) |- ∃!(x, phi)) by RightExistsOne
+          lib.have(phi |- (deltaPrime + ((x === y) <=> substPhi))) by RightIff(forward, backward)
+          lib.thenHave(phi |- (deltaPrime + ∀(y, (x === y) <=> substPhi))) by RightForall
+          lib.thenHave(phi |- (deltaPrime + ∃(x, ∀(y, (x === y) <=> substPhi)))) by RightExists
+          lib.thenHave(∃(x, phi) |- (deltaPrime + ∃(x, ∀(y, (x === y) <=> substPhi)))) by LeftExists
+          lib.thenHave(∃(x, phi) |- (deltaPrime + ∃!(x, phi))) by RightExistsOne
 
           lib.have(bot) by Cut(existence, lib.lastStep)
         }

From 79d45ab9565f889f151e05eda40dde086041b148 Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Sun, 11 Jun 2023 17:28:49 +0200
Subject: [PATCH 59/68] Simplify inverse symmetry proof

---
 .../scala/lisa/mathematics/GroupTheory.scala  | 47 ++++++++++---------
 1 file changed, 25 insertions(+), 22 deletions(-)

diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala
index 43667de2..54d90d5d 100644
--- a/src/main/scala/lisa/mathematics/GroupTheory.scala
+++ b/src/main/scala/lisa/mathematics/GroupTheory.scala
@@ -473,11 +473,17 @@ object GroupTheory extends lisa.Main {
    * Theorem --- `y` is the inverse of `x` iff `x` is the inverse of `y`
    */
   val inverseSymmetry = Theorem(
-    (group(G, *), x ∈ G) |- (isInverse(y, x, G, *) ==> isInverse(x, y, G, *))
+    (group(G, *), x ∈ G, y ∈ G) |- (y === inverse(x, G, *)) <=> (x === inverse(y, G, *))
   ) {
     assume(group(G, *))
-    assume(x ∈ G)
 
+    val inverseCharacterization = have(x ∈ G |- ((y === inverse(x, G, *)) <=> isInverse(y, x, G, *))) subproof {
+      have(x ∈ G |- ∀(y, (y === inverse(x, G, *)) <=> isInverse(y, x, G, *))) by Tautology.from(inverseUniqueness, inverse.definition)
+      thenHave(thesis) by InstantiateForall(y)
+    }
+
+    val forward = have(x ∈ G |- isInverse(y, x, G, *) ==> isInverse(x, y, G, *)) subproof {  
+    assume(x ∈ G)
     have(isInverse(y, x, G, *) |- y ∈ G /\ isNeutral(op(x, *, y), G, *) /\ isNeutral(op(y, *, x), G, *)) by Tautology.from(isInverse.definition)
     thenHave(isInverse(y, x, G, *) |- isNeutral(op(y, *, x), G, *) /\ isNeutral(op(x, *, y), G, *)) by Tautology
     thenHave(isInverse(y, x, G, *) |- x ∈ G /\ isNeutral(op(y, *, x), G, *) /\ isNeutral(op(x, *, y), G, *)) by Tautology
@@ -486,32 +492,29 @@ object GroupTheory extends lisa.Main {
     thenHave(thesis) by Restate
   }
 
+    val backward = forward of (x -> y, y -> x)
+
+    have((x ∈ G, y ∈ G) |- isInverse(y, x, G, *) <=> isInverse(x, y, G, *)) by RightIff(forward, backward)
+
+    have(thesis) by Tautology.from(
+      inverseCharacterization,
+      lastStep,
+      inverseCharacterization of (x -> y, y -> x)
+    )
+  }
+
   /**
    * Involution of the inverse -- For all `x`, `inverse(inverse(x)) = x`.
+   * 
+   * Direct corollary of [[inverseSymmetry]].
    */
   val inverseIsInvolutive = Theorem(
     (group(G, *), x ∈ G) |- (inverse(inverse(x, G, *), G, *) === x)
   ) {
-    assume(group(G, *))
-
-    val inverseCharacterization = have(x ∈ G |- ((y === inverse(x, G, *)) <=> isInverse(y, x, G, *))) subproof {
-      have(x ∈ G |- ∀(y, (y === inverse(x, G, *)) <=> isInverse(y, x, G, *))) by Tautology.from(inverseUniqueness, inverse.definition)
-      thenHave(thesis) by InstantiateForall(y)
-    }
-
-    // Use the symmetry of the inverse relation to prove that x is an inverse of inverse(x)
-    val satisfaction = have(x ∈ G |- isInverse(x, inverse(x, G, *), G, *)) subproof {
-      have((x ∈ G, isInverse(inverse(x, G, *), x, G, *)) |- isInverse(x, inverse(x, G, *), G, *)) by Restate.from(inverseSymmetry of (y -> inverse(x, G, *)))
-      have(thesis) by Cut(inverseIsInverse, lastStep)
-    }
-
-    // Conclude by uniqueness
-    val u = inverse(inverse(x, G, *), G, *)
-
-    val characterization = have(inverse(x, G, *) ∈ G |- ((x === u) <=> isInverse(x, inverse(x, G, *), G, *))) by Restate.from(inverseCharacterization of (x -> inverse(x, G, *), y -> x))
-    val eq = have((x ∈ G, inverse(x, G, *) ∈ G) |- (x === u)) by Tautology.from(satisfaction, characterization)
-
-    have(thesis) by Cut(inverseInGroup, eq)
+    have(thesis) by Tautology.from(
+      inverseSymmetry of (y -> inverse(x, G, *)),
+      inverseInGroup
+    )
   }
 
   /**

From 9582b7ba16815fcea233e670a2300ea76e9baabd Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Sun, 11 Jun 2023 18:35:46 +0200
Subject: [PATCH 60/68] Prove inverse cancellation

---
 .../scala/lisa/mathematics/GroupTheory.scala  | 32 +++++++++++++++++++
 1 file changed, 32 insertions(+)

diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala
index 54d90d5d..9becae9f 100644
--- a/src/main/scala/lisa/mathematics/GroupTheory.scala
+++ b/src/main/scala/lisa/mathematics/GroupTheory.scala
@@ -469,6 +469,38 @@ object GroupTheory extends lisa.Main {
     )
   }
 
+  /**
+   * Theorem --- For any element `x`, we have `x * inverse(x) = inverse(x) * x = e`.
+   */
+  val inverseCancellation = Theorem(
+    (group(G, *), x ∈ G) |- (op(x, *, inverse(x, G, *)) === identity(G, *)) /\ (op(inverse(x, G, *), *, x) === identity(G, *))
+  ) {
+    assume(group(G, *))
+
+    have(∀(y, (y === identity(G, *)) <=> isNeutral(y, G, *))) by Tautology.from(
+      identity.definition,
+      identityUniqueness
+    )
+    val idCharacterization = thenHave((y === identity(G, *)) <=> isNeutral(y, G, *)) by InstantiateForall(y)
+
+    assume(x ∈ G)
+    val inverseDef = have((inverse(x, G, *) ∈ G) /\ isNeutral(op(x, *, inverse(x, G, *)), G, *) /\ isNeutral(op(inverse(x, G, *), *, x), G, *)) by Tautology.from(
+      inverseIsInverse,
+      isInverse.definition of (y -> inverse(x, G, *))
+    )
+
+    val left = have(op(x, *, inverse(x, G, *)) === identity(G, *)) by Tautology.from(
+      inverseDef,
+      idCharacterization of (y -> op(x, *, inverse(x, G, *)))
+    )
+    val right = have(op(inverse(x, G, *), *, x) === identity(G, *)) by Tautology.from(
+      inverseDef,
+      idCharacterization of (y -> op(inverse(x, G, *), *, x))
+    )
+
+    have(thesis) by RightAnd(left, right)
+  }
+
   /**
    * Theorem --- `y` is the inverse of `x` iff `x` is the inverse of `y`
    */

From 3b705cc54dec0732ad7c03895247f45ea6b68b6e Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Sun, 11 Jun 2023 18:36:08 +0200
Subject: [PATCH 61/68] Prove more inverse theorems

---
 .../scala/lisa/mathematics/GroupTheory.scala  | 68 +++++++++++++++++++
 1 file changed, 68 insertions(+)

diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala
index 9becae9f..3d3de1fd 100644
--- a/src/main/scala/lisa/mathematics/GroupTheory.scala
+++ b/src/main/scala/lisa/mathematics/GroupTheory.scala
@@ -666,6 +666,74 @@ object GroupTheory extends lisa.Main {
     have(thesis) by RightIff(forward, backward)
   }
 
+  /**
+   * Theorem --- If `x * y = e` then `y = inverse(x)`.
+   * 
+   * This also implies that `x = inverse(y)` by [[inverseSymmetry]].
+   */
+  val inverseTest = Theorem(
+    (group(G, *), x ∈ G, y ∈ G) |- (op(x, *, y) === identity(G, *)) ==> (y === inverse(x, G, *))
+  ) {
+    assume(group(G, *))
+    assume(x ∈ G)
+    assume(y ∈ G)
+
+    val e = identity(G, *)
+
+    // 1. e = x * inverse(x)
+    val eq1 = have(op(x, *, inverse(x, G, *)) === e) by Tautology.from(
+      identityInGroup,
+      inverseCancellation
+    )
+
+    // 2. x * y = x * inverse(x)
+    have((op(x, *, y) === e) |- (op(x, *, y) === e) ) by Hypothesis
+    val eq2 = have((op(x, *, y) === e) |- op(x, *, y) === op(x, *, inverse(x, G, *))) by Equalities(eq1, lastStep)
+
+    // Conclude by left cancellation
+    have((op(x, *, y) === e, inverse(x, G, *) ∈ G) |- (y === inverse(x, G, *))) by Tautology.from(
+      lastStep,
+      leftCancellation of (z -> inverse(x, G, *))
+    )
+    have((op(x, *, y) === e) |- (y === inverse(x, G, *))) by Cut(inverseInGroup, lastStep)
+    thenHave(thesis) by Restate
+  }
+
+  /**
+   * Theorem --- The inverse of the identity element is itself.
+   */
+  val inverseOfIdentityIsIdentity = Theorem(
+    group(G, *) |- inverse(identity(G, *), G, *) === identity(G, *)
+  ) {
+    assume(group(G, *))
+
+    val e = identity(G, *)
+
+    have(x ∈ G |- ∀(y, (y === inverse(x, G, *)) <=> isInverse(y, x, G, *))) by Tautology.from(
+      inverseUniqueness,
+      inverse.definition
+    )
+    thenHave(x ∈ G |- (e === inverse(x, G, *)) <=> isInverse(e, x, G, *)) by InstantiateForall(e)
+    val characterization = have((e === inverse(e, G, *)) <=> isInverse(e, e, G, *)) by Cut(identityInGroup, lastStep of (x -> e))
+
+    // Prove that e is an inverse of e
+    val satisfaction = have(isInverse(e, e, G, *)) subproof {
+      val neutrality = have(op(e, *, e) === e) by Cut(identityInGroup, identityNeutrality of (x -> e))
+
+      have((op(e, *, e) === e) |- isNeutral(op(e, *, e), G, *)) by RightSubstEq(
+        List((op(e, *, e), e)),
+        lambda(z, isNeutral(z, G, *))
+      )(identityIsNeutral)
+
+      have(isNeutral(op(e, *, e), G, *)) by Cut(neutrality, lastStep)
+
+      have(e ∈ G /\ isNeutral(op(e, *, e), G, *) /\ isNeutral(op(e, *, e), G, *)) by RightAnd(identityInGroup, lastStep, lastStep)
+      have(thesis) by Tautology.from(lastStep, isInverse.definition of (x -> e, y -> e))
+    }
+
+    have(thesis) by Tautology.from(characterization, satisfaction)
+  }
+
   // TODO Direct product group
   // TODO Permutation group
 

From 37bbace8ef0d2655ce4dabb9a478ef9791670c22 Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Sun, 11 Jun 2023 18:40:51 +0200
Subject: [PATCH 62/68] Prove that kernel is a subgroup

---
 .../scala/lisa/mathematics/GroupTheory.scala  | 326 +++++++++++++++++-
 1 file changed, 310 insertions(+), 16 deletions(-)

diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala
index 3d3de1fd..3cb43eb4 100644
--- a/src/main/scala/lisa/mathematics/GroupTheory.scala
+++ b/src/main/scala/lisa/mathematics/GroupTheory.scala
@@ -45,8 +45,6 @@ object GroupTheory extends lisa.Main {
    *
    * This is useful in defining specific elements in groups, where their uniqueness (and existence) strongly rely
    * on the assumption of the group structure.
-   *
-   * TODO This should probably be merged with [[The]] with an additional `orElse` method, to be discussed
    */
   def TheConditional(u: VariableLabel, f: Formula)(just: runningSetTheory.Theorem, defaultValue: Term = ∅): The = {
     val seq = just.proposition
@@ -289,6 +287,46 @@ object GroupTheory extends lisa.Main {
     have(thesis) by Cut(groupOperationDomain, lastStep)
   }
 
+  /**
+   * Lemma --- If `x, y ∈ G`, then `x * y ∈ G`.
+   */
+  val groupIsClosedByProduct = Lemma(
+    (group(G, *), x ∈ G, y ∈ G) |- op(x, *, y) ∈ G
+  ) {
+    have(∀(t, (t ∈ relationRange(*)) <=> ∃(a, pair(a, t) ∈ *))) by Definition(relationRange, relationRangeUniqueness)(*)
+    val relationRangeDef = thenHave((op(x, *, y) ∈ relationRange(*)) <=> ∃(a, pair(a, op(x, *, y)) ∈ *)) by InstantiateForall(op(x, *, y))
+
+    val appDef = have(
+      (functional(*), pair(x, y) ∈ relationDomain(*)) |- pair(pair(x, y), op(x, *, y)) ∈ *
+    ) by Definition(app, functionApplicationUniqueness)(*, pair(x, y))
+
+    assume(group(G, *))
+    assume(x ∈ G)
+    assume(y ∈ G)
+
+    // Show that x * y is in relation range  
+    have(pair(pair(x, y), op(x, *, y)) ∈ *) by Tautology.from(
+      appDef,
+      groupOperationIsFunctional,
+      groupPairInOperationDomain
+    )
+    thenHave(∃(a, pair(a, op(x, *, y)) ∈ *)) by RightExists
+
+    val productInRelationRange = have(op(x, *, y) ∈ relationRange(*)) by Tautology.from(lastStep, relationRangeDef)
+
+    // Conclude by [[functionImpliesRangeSubsetOfCodomain]]
+    have(∀(t, t ∈ relationRange(*) ==> t ∈ G)) by Tautology.from(
+      group.definition,
+      binaryOperation.definition,
+      functionImpliesRangeSubsetOfCodomain of (f -> *, x -> cartesianProduct(G, G), y -> G),
+      subset.definition of (x -> relationRange(*), y -> G)
+    )
+    thenHave(op(x, *, y) ∈ relationRange(*) ==> op(x, *, y) ∈ G) by InstantiateForall(op(x, *, y))
+    thenHave(op(x, *, y) ∈ relationRange(*) |- op(x, *, y) ∈ G) by Restate
+    
+    have(thesis) by Cut(productInRelationRange, lastStep)
+  }
+
   /**
    * Identity uniqueness --- In a group (G, *), an identity element is unique, i.e. if both `e * x = x * e = x` and
    * `f * x = x * f = x` for all `x`, then `e = f`.
@@ -515,14 +553,14 @@ object GroupTheory extends lisa.Main {
     }
 
     val forward = have(x ∈ G |- isInverse(y, x, G, *) ==> isInverse(x, y, G, *)) subproof {  
-    assume(x ∈ G)
-    have(isInverse(y, x, G, *) |- y ∈ G /\ isNeutral(op(x, *, y), G, *) /\ isNeutral(op(y, *, x), G, *)) by Tautology.from(isInverse.definition)
-    thenHave(isInverse(y, x, G, *) |- isNeutral(op(y, *, x), G, *) /\ isNeutral(op(x, *, y), G, *)) by Tautology
-    thenHave(isInverse(y, x, G, *) |- x ∈ G /\ isNeutral(op(y, *, x), G, *) /\ isNeutral(op(x, *, y), G, *)) by Tautology
+      assume(x ∈ G)
+      have(isInverse(y, x, G, *) |- y ∈ G /\ isNeutral(op(x, *, y), G, *) /\ isNeutral(op(y, *, x), G, *)) by Tautology.from(isInverse.definition)
+      thenHave(isInverse(y, x, G, *) |- isNeutral(op(y, *, x), G, *) /\ isNeutral(op(x, *, y), G, *)) by Tautology
+      thenHave(isInverse(y, x, G, *) |- x ∈ G /\ isNeutral(op(y, *, x), G, *) /\ isNeutral(op(x, *, y), G, *)) by Tautology
 
-    have(isInverse(y, x, G, *) |- isInverse(x, y, G, *)) by Tautology.from(lastStep, isInverse.definition of (y -> x, x -> y))
-    thenHave(thesis) by Restate
-  }
+      have(isInverse(y, x, G, *) |- isInverse(x, y, G, *)) by Tautology.from(lastStep, isInverse.definition of (y -> x, x -> y))
+      thenHave(thesis) by Restate
+    }
 
     val backward = forward of (x -> y, y -> x)
 
@@ -982,6 +1020,27 @@ object GroupTheory extends lisa.Main {
 
   // TODO Subgroup inverse
   // TODO Subgroup condition
+
+  /**
+   * Theorem --- A subset `H ⊆ G` of a group `(G, *)` is a subgroup if and only if:
+   *   1. `H` is non-empty,
+   *   2. `H` is closed by products, and
+   *   3. `H` is closed by inversion.
+   * 
+   * It is often easier to prove the 3 conditions independently than using the definition directly.
+   * 
+   * Note that in the case where H is finite, conditions 1 and 2 are sufficient.
+   */
+  val subgroupCondition = Theorem(
+    (group(G, *), subset(H, G)) |- (subgroup(H, G, *) <=> (
+      (H !== ∅) /\
+      ∀(x, x ∈ H ==> ∀(y, y ∈ H ==> (op(x, *, y) ∈ H))) /\
+      ∀(x, x ∈ H ==> inverse(x, G, *) ∈ H)
+    ))
+  ) {
+    sorry
+  }
+
   // TODO Trivial subgroup
 
   //
@@ -1013,6 +1072,19 @@ object GroupTheory extends lisa.Main {
     thenHave(thesis) by Restate
   }
 
+  /**
+   * Lemma --- If `f` is a homomorphism, then `f(x) ∈ H` for all `x ∈ G`.
+   */
+  private val homomorphismAppInH = Lemma(
+    (homomorphism(f, G, *, H, ★), x ∈ G) |- app(f, x) ∈ H
+  ) {
+    have(homomorphism(f, G, *, H, ★) |- functionFrom(f, G, H)) by Tautology.from(homomorphism.definition)
+    have(thesis) by Cut(
+      lastStep,
+      functionAppInCodomain of (VariableLabel("t") -> x, VariableLabel("x") -> G, y -> H)
+    )
+  }
+
   /**
    * Theorem --- If `f` is a group-homomorphism between `G` and `H`, then `f(e_G) = e_H`.
    */
@@ -1023,15 +1095,9 @@ object GroupTheory extends lisa.Main {
 
     val groupG = have(homomorphism(f, G, *, H, ★) |- group(G, *)) by Tautology.from(homomorphism.definition)
     val groupH = have(homomorphism(f, G, *, H, ★) |- group(H, ★)) by Tautology.from(homomorphism.definition)
-    val functionF = have(homomorphism(f, G, *, H, ★) |- functionFrom(f, G, H)) by Tautology.from(homomorphism.definition)
 
     val identityInG = have(homomorphism(f, G, *, H, ★) |- e ∈ G) by Cut(groupG, identityInGroup)
-
-    have((homomorphism(f, G, *, H, ★), e ∈ G) |- app(f, e) ∈ H) by Cut(
-      functionF,
-      functionAppInCodomain of (VariableLabel("t") -> e, x -> G, y -> H)
-    )
-    val appInH = have(homomorphism(f, G, *, H, ★) |- app(f, e) ∈ H) by Cut(identityInG, lastStep)
+    val appInH = have(homomorphism(f, G, *, H, ★) |- app(f, e) ∈ H) by Cut(identityInG, homomorphismAppInH of (x -> e))
 
     // 0. e * e = e (to apply substitution)
     have(group(G, *) |- op(e, *, e) === e) by Cut(
@@ -1069,6 +1135,63 @@ object GroupTheory extends lisa.Main {
     
     have(thesis) by Tautology.from(lastStep, eq3)
   }
+  
+  /**
+   * Theorem --- If `f: G -> H` is a group homomorphism, then `f(inverse(x, G, *)) = inverse(f(x), H, ★)`.
+   */
+  val homomorphismMapsInverseToInverse = Theorem(
+    (homomorphism(f, G, *, H, ★), x ∈ G) |- app(f, inverse(x, G, *)) === inverse(app(f, x), H, ★)
+  ) {
+    assume(homomorphism(f, G, *, H, ★))
+    assume(x ∈ G)
+
+    val groupG = have(group(G, *)) by Tautology.from(homomorphism.definition)
+    val groupH = have(group(H, ★)) by Tautology.from(homomorphism.definition)
+
+    val eG = identity(G, *)
+    val eH = identity(H, ★)
+    val i = inverse(x, G, *)
+    val iInG = have(i ∈ G) by Cut(groupG, inverseInGroup)
+
+    // 1. f(x * inverse(x)) = f(x) f(inverse(x))
+    val eq1 = have(app(f, op(x, *, i)) === op(app(f, x), ★, app(f, i))) by Cut(
+      iInG,
+      homomorphismApplication of (y -> i)
+    )
+
+    // 2. f(x * inverse(x)) = f(e)
+    val cancellation = have(op(x, *, i) === eG) by Tautology.from(
+      groupG,
+      inverseCancellation
+    )
+
+    have(app(f, op(x, *, i)) === app(f, op(x, *, i))) by RightRefl
+    thenHave((op(x, *, i) === eG) |- (app(f, op(x, *, i)) === app(f, eG))) by RightSubstEq(
+      List((op(x, *, i), eG)),
+      lambda(z, app(f, op(x, *, i)) === app(f, z))
+    )
+
+    val eq2 = have(app(f, op(x, *, i)) === app(f, eG)) by Cut(cancellation, lastStep)
+
+    // 3. f(e) = e'
+    val eq3 = have(app(f, eG) === eH) by Tautology.from(homomorphismMapsIdentityToIdentity)
+
+    // 4. f(x)f(inverse(x)) = e'
+    val eq4 = have(op(app(f, x), ★, app(f, i)) === eH) by Equalities(eq1, eq2, eq3)
+
+    // Conclude
+    val conclusion = have((app(f, i) ∈ H) |- (app(f, i) === inverse(app(f, x), H, ★))) by Tautology.from(
+      groupH,
+      inverseTest of (G -> H, * -> ★, x -> app(f, x), y -> app(f, i)),
+      eq4,
+      homomorphismAppInH
+    )
+    have(app(f, i) ∈ H) by Cut(iInG, homomorphismAppInH of (x -> i))
+
+    have(thesis) by Cut(lastStep, conclusion)
+  }
+
+  // TODO Homomorphism composition once we have function composition
 
   /**
    * Kernel uniqueness --- The kernel of a homomorphism is well-defined.
@@ -1091,4 +1214,175 @@ object GroupTheory extends lisa.Main {
    * Kernel --- The kernel of a homomorphism `f: G -> H` is the set of elements `t ∈ G` such that `f(t) = e_H`.
    */
   val kernel = DEF(f, G, *, H, ★) --> TheConditional(z, ∀(t, (t ∈ z) <=> (t ∈ G /\ (app(f, t) === identity(H, ★)))))(kernelUniqueness)
+
+  // Shortcut alias
+  private val K = kernel(f, G, *, H, ★)
+
+  /**
+   * Lemma --- Reformulation of the kernel definition.
+   */
+  private val kernelDef = Lemma(
+    homomorphism(f, G, *, H, ★) |- (x ∈ K) <=> (x ∈ G /\ (app(f, x) === identity(H, ★)))
+  ) {
+    assume(homomorphism(f, G, *, H, ★))
+    have(∀(t, (t ∈ K) <=> (t ∈ G /\ (app(f, t) === identity(H, ★))))) by Definition(kernel, kernelUniqueness)(f, G, *, H, ★)
+    thenHave(thesis) by InstantiateForall(x)
+  }
+
+  /**
+   * Lemma --- The kernel is closed by products, i.e. if `x, y ∈ K`, then `x * y ∈ K`.
+   */
+  val kernelIsClosedByProducts = Lemma(
+    (homomorphism(f, G, *, H, ★), x ∈ K, y ∈ K) |- op(x, *, y) ∈ K
+  ) {
+    assume(homomorphism(f, G, *, H, ★))
+    assume(x ∈ K)
+    assume(y ∈ K)
+
+    val elemInG = have(x ∈ G) by Tautology.from(kernelDef)
+
+    val groupG = have(group(G, *)) by Tautology.from(homomorphism.definition)
+    val groupH = have(group(H, ★)) by Tautology.from(homomorphism.definition)
+
+    val e = identity(H, ★)
+    val eInH = have(e ∈ H) by Cut(groupH, identityInGroup of (G -> H, * -> ★))
+
+    // 1. f(x) ★ f(y) = f(x * y)
+    val eq1 = have(app(f, op(x, *, y)) === op(app(f, x), ★, app(f, y))) by Tautology.from(
+      homomorphismApplication,
+      elemInG,
+      elemInG of (x -> y)
+    )
+
+    // 2. f(x) ★ f(y) = e ★ e
+    val appValue = have(app(f, x) === e) by Tautology.from(kernelDef)
+    have(op(app(f, x), ★, app(f, y)) === op(app(f, x), ★, app(f, y))) by RightRefl
+    thenHave((app(f, x) === e, app(f, y) === e) |- op(app(f, x), ★, app(f, y)) === op(e, ★, e)) by RightSubstEq(
+      List((app(f, x), e), (app(f, y), e)),
+      lambda(Seq(a, b), op(app(f, x), ★, app(f, y)) === op(a, ★, b))
+    )
+
+    val eq2 = have(op(app(f, x), ★, app(f, y)) === op(e, ★, e)) by Tautology.from(
+      lastStep,
+      appValue,
+      appValue of (x -> y)
+    )
+
+    // 3. e ★ e = e
+    val eq3 = have(op(e, ★, e) === e) by Tautology.from(
+      identityNeutrality of (G -> H, * -> ★, x -> e),
+      groupH,
+      eInH
+    )
+
+    // 4. f(x * y) = e
+    val eq4 = have(app(f, op(x, *, y)) === e) by Equalities(eq1, eq2, eq3)
+
+    // Conclude that x * y ∈ K
+    have(op(x, *, y) ∈ G) by Tautology.from(
+      groupG,
+      elemInG,
+      elemInG of (x -> y),
+      groupIsClosedByProduct
+    )
+
+    have(op(x, *, y) ∈ G /\ (app(f, op(x, *, y)) === e)) by RightAnd(lastStep, eq4)
+    have(thesis) by Tautology.from(lastStep, kernelDef of (x -> op(x, *, y)))
+  }
+
+  /**
+   * Lemma --- The kernel is closed by inversion, i.e. if `x ∈ K` then `inverse(x, G, *) ∈ K`.
+   */
+  val kernelIsClosedByInversion = Lemma (
+    (homomorphism(f, G, *, H, ★), x ∈ K) |- inverse(x, G, *) ∈ K
+   ) {
+    assume(homomorphism(f, G, *, H, ★))
+    assume(x ∈ K)
+
+    val groupG = have(group(G, *)) by Tautology.from(homomorphism.definition)
+    val groupH = have(group(H, ★)) by Tautology.from(homomorphism.definition)
+    val elemInG = have(x ∈ G) by Tautology.from(kernelDef)
+
+    val e = identity(H, ★)
+    val appValue = have(app(f, x) === e) by Tautology.from(kernelDef)
+
+    // 1. f(inverse(x)) = inverse(f(x)) = inverse(e)
+    have(app(f, inverse(x, G, *)) === inverse(app(f, x), H, ★)) by Tautology.from(
+      homomorphismMapsInverseToInverse,
+      elemInG
+    )
+    thenHave((app(f, x) === e) |- (app(f, inverse(x, G, *)) === inverse(e, H, ★))) by RightSubstEq(
+      List((app(f, x), e)),
+      lambda(z, app(f, inverse(x, G, *)) === inverse(z, H, ★))
+    )
+    
+    val eq1 = have(app(f, inverse(x, G, *)) === inverse(e, H, ★)) by Cut(appValue, lastStep)
+
+    // 2. inverse(e) = e
+    val eq2 = have(inverse(e, H, ★) === e) by Cut(groupH, inverseOfIdentityIsIdentity of (G -> H, * -> ★))
+
+    // 3. Conclude
+    val eq3 = have(app(f, inverse(x, G, *)) === e) by Equalities(eq1, eq2)
+    have(inverse(x, G, *) ∈ G) by Tautology.from(
+      groupG,
+      elemInG,
+      inverseInGroup
+    )
+
+    have((inverse(x, G, *) ∈ G) /\ (app(f, inverse(x, G, *)) === e)) by RightAnd(lastStep, eq3)
+
+    have(thesis) by Tautology.from(lastStep, kernelDef of (x -> inverse(x, G, *)))
+  }
+
+  /**
+   * Theorem --- The kernel of a homomorphism `f: G -> H` is a subgroup of `G`.
+   */
+  val kernelIsSubgroup = Theorem(
+    homomorphism(f, G, *, H, ★) |- subgroup(kernel(f, G, *, H, ★), G, *)
+  ) {
+    assume(homomorphism(f, G, *, H, ★))
+    val groupG = have(group(G, *)) by Tautology.from(homomorphism.definition)
+
+    // We show that the kernel satisfies all requirements of [[subgroupCondition]]
+    have((x ∈ K) ==> (x ∈ G)) by Tautology.from(kernelDef)
+    thenHave(∀(x, x ∈ K ==> x ∈ G)) by RightForall
+    val kernelIsSubset = have(subset(K, G)) by Tautology.from(lastStep, subsetAxiom of (x -> K, y -> G))
+
+    // 1. kernel != ∅
+    have(identity(G, *) ∈ G) by Cut(groupG, identityInGroup)
+    have(identity(G, *) ∈ G /\ (app(f, identity(G, *)) === identity(H, ★))) by RightAnd(
+      lastStep,
+      homomorphismMapsIdentityToIdentity
+    )
+    have(identity(G, *) ∈ K) by Tautology.from(
+      lastStep,
+      kernelDef of (x -> identity(G, *))
+    )
+    val condition1 = have(K !== ∅) by Cut(lastStep, setWithElementNonEmpty of (y -> identity(G, *), x -> K))
+    
+    // 2. The kernel is closed by products
+    have(x ∈ K |- ((y ∈ K) ==> op(x, *, y) ∈ K)) by Restate.from(kernelIsClosedByProducts)
+    thenHave(x ∈ K |- ∀(y, (y ∈ K) ==> op(x, *, y) ∈ K)) by RightForall
+    thenHave(x ∈ K ==> ∀(y, (y ∈ K) ==> op(x, *, y) ∈ K)) by Restate
+    val condition2 = thenHave(∀(x, x ∈ K ==> ∀(y, (y ∈ K) ==> op(x, *, y) ∈ K))) by RightForall
+
+    // 3. The kernel is closed by inversion
+    have((x ∈ K) ==> (inverse(x, G, *) ∈ K)) by Restate.from(kernelIsClosedByInversion)
+    val condition3 = thenHave(∀(x, (x ∈ K) ==> (inverse(x, G, *) ∈ K))) by RightForall
+
+    // Conclude
+    have((K !== ∅) /\ ∀(x, x ∈ K ==> ∀(y, (y ∈ K) ==> op(x, *, y) ∈ K)) /\ ∀(x, (x ∈ K) ==> (inverse(x, G, *) ∈ K))) by RightAnd(
+      condition1, condition2, condition3
+    )
+
+    have(subgroup(K, G, *)) by Tautology.from(
+      lastStep,
+      subgroupCondition of (H -> K),
+      groupG,
+      kernelIsSubset
+    )
+  }
+
+  // TODO Kernel injectivity
+  // TODO Image is subgroup
 }

From a389de7c44d40c53b9221d2573b870283c1d30ba Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Sun, 11 Jun 2023 18:41:10 +0200
Subject: [PATCH 63/68] Define isomorphisms and automorphisms

---
 src/main/scala/lisa/mathematics/GroupTheory.scala | 12 ++++++++++++
 1 file changed, 12 insertions(+)

diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala
index 3cb43eb4..f239a9b8 100644
--- a/src/main/scala/lisa/mathematics/GroupTheory.scala
+++ b/src/main/scala/lisa/mathematics/GroupTheory.scala
@@ -1385,4 +1385,16 @@ object GroupTheory extends lisa.Main {
 
   // TODO Kernel injectivity
   // TODO Image is subgroup
+
+  /**
+   * Isomorphism --- An isomorphism `f: G -> H` is a bijective homomorphism.
+   * 
+   * In some sense, isomorphic groups are equivalent, up to relabelling their elements.
+   */
+  val isomorphism = DEF(f, G, *, H, ★) --> homomorphism(f, G, *, H, ★) /\ bijective(f, G, H)
+
+  /**
+   * Automorphism --- An automorphism is an isomorphism from a group to itself.
+   */
+  val automorphism = DEF(f, G, *) --> isomorphism(f, G, *, G, *)
 }

From d935f4ecf2859d9f1bf6d3e3f66732e29fba62c7 Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Sun, 11 Jun 2023 21:21:06 +0200
Subject: [PATCH 64/68] Prove subgroup inverse

---
 .../scala/lisa/mathematics/GroupTheory.scala  | 64 ++++++++++++++++++-
 1 file changed, 62 insertions(+), 2 deletions(-)

diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala
index f239a9b8..73b914d2 100644
--- a/src/main/scala/lisa/mathematics/GroupTheory.scala
+++ b/src/main/scala/lisa/mathematics/GroupTheory.scala
@@ -1018,8 +1018,68 @@ object GroupTheory extends lisa.Main {
     have(thesis) by Cut(lastStep, finalStep)
   }
 
-  // TODO Subgroup inverse
-  // TODO Subgroup condition
+  /**
+   * Theorem --- If `H` is a subgroup of `G`, then the inverse is the same as in the parent group. 
+   */
+  val subgroupInverse = Theorem(
+    (subgroup(H, G, *), x ∈ H) |- inverse(x, H, restrictedFunction(*, cartesianProduct(H, H))) === inverse(x, G, *)
+  ) {
+    assume(subgroup(H, G, *))
+    assume(x ∈ H)
+
+    have(∀(x, (x ∈ H) ==> (x ∈ G))) by Tautology.from(
+      subgroup.definition,
+      subset.definition of (x -> H, y -> G)
+    )
+    val subsetDef = thenHave((x ∈ H) ==> (x ∈ G)) by InstantiateForall(x)
+    val xInG = thenHave(x ∈ G) by Tautology
+
+    val groupG = have(group(G, *)) by Tautology.from(subgroup.definition)
+    val ★ = restrictedFunction(*, cartesianProduct(H, H))
+    val groupH = have(group(H, ★)) by Tautology.from(subgroup.definition)
+
+    val eG = identity(G, *)
+    val eH = identity(H, ★)
+
+    val inverseHInH = have(inverse(x, H, ★) ∈ H) by Cut(groupH, inverseInGroup of (G -> H, * -> ★))
+    val inverseHInG = have(inverse(x, H, ★) ∈ G) by Tautology.from(inverseHInH, subsetDef of (x -> inverse(x, H, ★)))
+
+    // 1. x * inverse(x, H, ★) = e_H
+    have((inverse(x, H, ★) ∈ H) |- (op(x, ★, inverse(x, H, ★)) === eH)) by Tautology.from(
+      groupH,
+      inverseCancellation of (G -> H, * -> ★)
+    )
+    have((inverse(x, H, ★) ∈ H) |- (op(x, *, inverse(x, H, ★)) === eH)) by Equalities(
+      lastStep,
+      subgroupOperation of (y -> inverse(x, H, ★))
+    )
+    
+    val eq1 = have(op(x, *, inverse(x, H, ★)) === eH) by Cut(inverseHInH, lastStep)
+
+    // 2. e_H = e_G
+    val eq2 = have(eH === eG) by Tautology.from(subgroupIdentity)
+
+    // 3. x * inverse(x, G, *) = e_G
+    val eq3 = have(op(x, *, inverse(x, G, *)) === eG) by Tautology.from(
+      groupG,
+      xInG,
+      inverseInGroup,
+      inverseCancellation
+    )
+
+    // 4. x * inverse(x, H, ★) === x * inverse(x, G, *)
+    have(op(x, *, inverse(x, H, ★)) === op(x, *, inverse(x, G, *))) by Equalities(eq1, eq2, eq3)
+
+    // Conclude by left cancellation
+    have(thesis) by Tautology.from(
+      lastStep,
+      groupG,
+      xInG,
+      inverseHInG,
+      inverseInGroup,
+      leftCancellation of (y -> inverse(x, H, ★), z -> inverse(x, G, *))
+    )
+  }
 
   /**
    * Theorem --- A subset `H ⊆ G` of a group `(G, *)` is a subgroup if and only if:

From c2b71ab175a8c2b86ce496c91ad2f995b6b536b2 Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Thu, 15 Jun 2023 20:43:19 +0200
Subject: [PATCH 65/68] Prove restricted functions application theorem

---
 .../lisa/mathematics/FirstOrderLogic.scala    |  39 ++++
 .../scala/lisa/mathematics/GroupTheory.scala  |  78 ++------
 .../scala/lisa/mathematics/SetTheory.scala    | 174 ++++++++++++++++++
 3 files changed, 225 insertions(+), 66 deletions(-)

diff --git a/src/main/scala/lisa/mathematics/FirstOrderLogic.scala b/src/main/scala/lisa/mathematics/FirstOrderLogic.scala
index 6159f98c..da2e8b0f 100644
--- a/src/main/scala/lisa/mathematics/FirstOrderLogic.scala
+++ b/src/main/scala/lisa/mathematics/FirstOrderLogic.scala
@@ -1,5 +1,6 @@
 package lisa.mathematics
 
+import lisa.automation.kernel.CommonTactics.Cases
 import lisa.automation.kernel.OLPropositionalSolver.Tautology
 import lisa.automation.kernel.SimplePropositionalSolver.*
 import lisa.automation.kernel.SimpleSimplifier.*
@@ -394,4 +395,42 @@ object FirstOrderLogic extends lisa.Main {
     thenHave(thesis) by Restate
   }
 
+  /**
+   * Theorem --- Unique existential quantifier distributes fully with a closed formula. 
+   */
+  val uniqueExistentialConjunctionWithClosedFormula = Theorem(
+    existsOne(x, P(x) /\ p()) <=> (existsOne(x, P(x)) /\ p())
+  ) {
+    val pos = have(p() |- existsOne(x, P(x) /\ p()) <=> (existsOne(x, P(x)) /\ p())) subproof {
+      have(p() |- (P(x) /\ p()) <=> P(x)) by Tautology
+      thenHave(p() |- forall(x, (P(x) /\ p()) <=> P(x))) by RightForall
+
+      have(p() |- existsOne(x, P(x) /\ p()) <=> (existsOne(x, P(x)))) by Cut(
+        lastStep,
+        uniqueExistentialEquivalenceDistribution of (
+          P -> lambda(x, P(x) /\ p()),
+          Q -> lambda(x, P(x))
+        )
+      )
+      thenHave(thesis) by Tautology
+    }
+
+    val neg = have(!p() |- existsOne(x, P(x) /\ p()) <=> (existsOne(x, P(x)) /\ p())) subproof {
+      have((!p(), (x === x) <=> (P(x) /\ p())) |- p()) by Tautology
+      thenHave((!p(), forall(y, (x === y) <=> (P(y) /\ p()))) |- p()) by LeftForall
+      thenHave((!p(), exists(x, forall(y, (x === y) <=> (P(y) /\ p())))) |- p()) by LeftExists
+      thenHave((!p(), existsOne(x, P(x) /\ p())) |- p()) by LeftExistsOne
+      thenHave((!p(), existsOne(x, P(x) /\ p())) |- ()) by LeftNot
+      thenHave((!p(), existsOne(x, P(x) /\ p())) |- existsOne(x, P(x)) /\ p()) by Weakening
+      val forward = thenHave(!p() |- existsOne(x, P(x) /\ p()) ==> (existsOne(x, P(x)) /\ p())) by Restate
+
+      have((!p(), existsOne(x, P(x)) /\ p()) |- ()) by Tautology
+      thenHave((!p(), existsOne(x, P(x)) /\ p()) |- existsOne(x, P(x) /\ p())) by Weakening
+      val backward = thenHave(!p() |- existsOne(x, P(x)) /\ p() ==> existsOne(x, P(x) /\ p())) by Restate
+      
+      have(thesis) by RightIff(forward, backward)
+    }
+
+    have(thesis) by Cases(pos, neg)
+  }
 }
diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala
index 73b914d2..0dcfd0e2 100644
--- a/src/main/scala/lisa/mathematics/GroupTheory.scala
+++ b/src/main/scala/lisa/mathematics/GroupTheory.scala
@@ -865,77 +865,23 @@ object GroupTheory extends lisa.Main {
   ) {
     val H2 = cartesianProduct(H, H)
     val ★ = restrictedFunction(*, H2)
-    val r = app(★, pair(x, y))
-
-    // We characterize op(x, *, y), and show that op(x, ★, y) satisfies all requirements
-    have(
-      ∀(
-        z,
-        (z === app(*, pair(x, y))) <=> (((functional(*) /\ in(pair(x, y), relationDomain(*))) ==> in(pair(pair(x, y), z), *)) /\
-          ((!functional(*) \/ !in(pair(x, y), relationDomain(*))) ==> (z === ∅)))
-      )
-    ) by Tautology.from(app.definition of (f -> *, x -> pair(x, y)), functionApplicationUniqueness)
-    val characterization = thenHave(
-      (r === app(*, pair(x, y))) <=> (((functional(*) /\ in(pair(x, y), relationDomain(*))) ==> in(pair(pair(x, y), r), *)) /\
-        ((!functional(*) \/ !in(pair(x, y), relationDomain(*))) ==> (r === ∅)))
-    ) by InstantiateForall(r)
-
-    // Prove that the premises of the first implication hold
-    val leftPremise = have(subgroup(H, G, *) |- functional(*)) subproof {
+
       assume(subgroup(H, G, *))
-      have(group(G, *)) by Tautology.from(subgroup.definition)
-      have(functional(*)) by Tautology.from(lastStep, groupOperationIsFunctional)
-    }
+    val groupG = have(group(G, *)) by Tautology.from(subgroup.definition)
+    val groupH = have(group(H, ★)) by Tautology.from(subgroup.definition)
 
-    val premises = have((subgroup(H, G, *), x ∈ H, y ∈ H) |- functional(*) /\ pair(x, y) ∈ relationDomain(*)) by RightAnd(
-      leftPremise,
-      subgroupPairInParentOperationDomain
+    have((x ∈ H, y ∈ H) |- pair(x, y) ∈ relationDomain(★)) by Tautology.from(
+      groupH, groupPairInOperationDomain of (G -> H, * -> ★)
     )
-
-    // We show that op(x, ★, y) satisfies the conclusion of the implication
-    val appDef = have(
-      (functional(★), pair(x, y) ∈ relationDomain(★)) |- pair(pair(x, y), r) ∈ ★
-    ) by Definition(app, functionApplicationUniqueness)(★, pair(x, y))
-
-    // Reduce the assumptions of the definition to our subgroup assumption
-    val reduction1 = have(subgroup(H, G, *) |- group(H, ★)) by Tautology.from(subgroup.definition)
-    val reduction2 = have(subgroup(H, G, *) |- functional(★)) by Tautology.from(lastStep, groupOperationIsFunctional of (G -> H, * -> ★))
-
-    val reduction3 = have((subgroup(H, G, *), pair(x, y) ∈ relationDomain(★)) |- pair(pair(x, y), r) ∈ ★) by Tautology.from(reduction2, appDef)
-
-    have((subgroup(H, G, *), x ∈ H, y ∈ H) |- pair(x, y) ∈ relationDomain(★)) by Cut(
-      reduction1,
-      groupPairInOperationDomain of (G -> H, * -> ★)
+    have((functional(*), x ∈ H, y ∈ H) |- op(x, ★, y) === op(x, *, y)) by Cut(
+      lastStep,
+      restrictedFunctionApplication of (f -> *, VariableLabel("d") -> H2, x -> pair(x, y))
     )
-    val reducedDef = have((subgroup(H, G, *), x ∈ H, y ∈ H) |- pair(pair(x, y), r) ∈ ★) by Cut(lastStep, reduction3)
-
-    have(∀(u, (u ∈ ★) <=> (u ∈ * /\ ∃(y, ∃(z, y ∈ H2 /\ (u === pair(y, z))))))) by Definition(restrictedFunction, restrictedFunctionUniqueness)(*, H2)
-    thenHave((u ∈ ★) <=> (u ∈ * /\ ∃(y, ∃(z, y ∈ H2 /\ (u === pair(y, z)))))) by InstantiateForall(u)
-    thenHave(u ∈ ★ ==> u ∈ *) by Tautology
-
-    val satisfaction = have((subgroup(H, G, *), x ∈ H, y ∈ H) |- pair(pair(x, y), r) ∈ *) by Tautology.from(
-      lastStep of (u -> pair(pair(x, y), r)),
-      reducedDef
+    have(thesis) by Tautology.from(
+      lastStep,
+      groupOperationIsFunctional,
+      groupG
     )
-
-    // Reconstruct the whole definition
-    assume(subgroup(H, G, *))
-    assume(x ∈ H)
-    assume(y ∈ H)
-
-    val pos = have((functional(*) /\ pair(x, y) ∈ relationDomain(*)) ==> pair(pair(x, y), r) ∈ *) by Tautology.from(premises, satisfaction)
-
-    have(!(functional(*) /\ pair(x, y) ∈ relationDomain(*)) |- ()) by LeftNot(premises)
-    thenHave(!functional(*) \/ !(pair(x, y) ∈ relationDomain(*)) |- ()) by Restate
-    thenHave(!functional(*) \/ !(pair(x, y) ∈ relationDomain(*)) |- (r === ∅)) by Weakening
-    val neg = thenHave((!functional(*) \/ !(pair(x, y) ∈ relationDomain(*))) ==> (r === ∅)) by Restate
-
-    have(
-      ((functional(*) /\ pair(x, y) ∈ relationDomain(*)) ==> pair(pair(x, y), r) ∈ *) /\
-        ((!functional(*) \/ !(pair(x, y) ∈ relationDomain(*))) ==> (r === ∅))
-    ) by RightAnd(pos, neg)
-
-    have(thesis) by Tautology.from(lastStep, characterization)
   }
 
   /**
diff --git a/src/main/scala/lisa/mathematics/SetTheory.scala b/src/main/scala/lisa/mathematics/SetTheory.scala
index 9af0cba5..23670618 100644
--- a/src/main/scala/lisa/mathematics/SetTheory.scala
+++ b/src/main/scala/lisa/mathematics/SetTheory.scala
@@ -5,6 +5,7 @@ import lisa.automation.kernel.OLPropositionalSolver.Tautology
 import lisa.automation.kernel.SimpleSimplifier.*
 import lisa.automation.settheory.SetTheoryTactics.*
 import lisa.mathematics.FirstOrderLogic.*
+import lisa.mathematics.FirstOrderLogic.existentialConjunctionWithClosedFormula
 
 /**
  * Set Theory Library
@@ -2399,6 +2400,117 @@ object SetTheory extends lisa.Main {
     have(thesis) by Tautology.from(lastStep, pairMembership)
   }
 
+  /**
+   * Lemma --- If `g` is the restriction of `f` to `x`, then `g` is a subset of `f`.
+   */
+  val restrictedFunctionIsSubset = Lemma(
+    subset(restrictedFunction(f, x), f)
+  ) {
+    have(∀(t, (t ∈ restrictedFunction(f, x)) <=> (t ∈ f /\ ∃(y, ∃(z, y ∈ x /\ (t === pair(y, z))))))) by Definition(restrictedFunction, restrictedFunctionUniqueness)(f, x)
+    thenHave((t ∈ restrictedFunction(f, x)) <=> (t ∈ f /\ ∃(y, ∃(z, y ∈ x /\ (t === pair(y, z)))))) by InstantiateForall(t)
+    thenHave(t ∈ restrictedFunction(f, x) ==> t ∈ f) by Tautology
+    thenHave(∀(t, t ∈ restrictedFunction(f, x) ==> t ∈ f)) by RightForall
+    
+    have(thesis) by Tautology.from(
+      lastStep,
+      subset.definition of (x -> restrictedFunction(f, x), y -> f)
+    )
+  }
+
+  /**
+   * Lemma --- If `f` is a function, so is any restriction of `f`.
+   */
+  val restrictedFunctionIsFunctional = Lemma(
+    functional(f) |- functional(restrictedFunction(f, x))
+  ) {
+    val g = restrictedFunction(f, x)
+    assume(functional(f))
+    
+    val functionalDef = have(relation(f) /\ ∀(x, ∃(y, in(pair(x, y), f)) ==> ∃!(y, in(pair(x, y), f)))) by Tautology.from(functional.definition)
+    thenHave(∀(x, ∃(y, in(pair(x, y), f)) ==> ∃!(y, in(pair(x, y), f)))) by Tautology
+    val fAppUniqueness = thenHave(∃(y, in(pair(t, y), f)) ==> ∃!(y, in(pair(t, y), f))) by InstantiateForall(t)
+    val fIsRelation = have(relation(f)) by Tautology.from(functionalDef)
+    
+    val gIsRelation = have(relation(g)) subproof {
+      // Since f is a relation and g is a subset of f, g is also a relation by subset transitivity
+      have(relationBetween(f, a, b) |- subset(f, cartesianProduct(a, b))) by Tautology.from(relationBetween.definition of (r -> f))
+      have(relationBetween(f, a, b) |- subset(g, f) /\ subset(f, cartesianProduct(a, b))) by RightAnd(restrictedFunctionIsSubset, lastStep)
+      have(relationBetween(f, a, b) |- subset(g, cartesianProduct(a, b))) by Cut(lastStep, subsetTransitivity of (a -> g, b -> f, c -> cartesianProduct(a, b)))
+
+      // Reconstruct the definition
+      have(relationBetween(f, a, b) |- relationBetween(g, a, b)) by Tautology.from(lastStep, relationBetween.definition of (r -> g))
+      thenHave(relationBetween(f, a, b) |- ∃(b, relationBetween(g, a, b))) by RightExists
+      thenHave(relationBetween(f, a, b) |- ∃(a, ∃(b, relationBetween(g, a, b)))) by RightExists
+      thenHave(∃(b, relationBetween(f, a, b)) |- ∃(a, ∃(b, relationBetween(g, a, b)))) by LeftExists
+      thenHave(∃(a, ∃(b, relationBetween(f, a, b))) |- ∃(a, ∃(b, relationBetween(g, a, b)))) by LeftExists
+
+      have(relation(f) |- relation(g)) by Tautology.from(
+        lastStep, relation.definition of (r -> f), relation.definition of (r -> g)
+      )
+      have(thesis) by Cut(fIsRelation, lastStep)
+    }
+
+    val gAppUniqueness = have(∀(t, ∃(y, pair(t, y) ∈ g) ==> ∃!(y, pair(t, y) ∈ g))) subproof { 
+      have((pair(t, a) ∈ restrictedFunction(f, x)) <=> (pair(t, a) ∈ f /\ in(t, x))) by Tautology.from(restrictedFunctionPairMembership)
+      val equiv = thenHave(∀(a, (pair(t, a) ∈ restrictedFunction(f, x)) <=> (pair(t, a) ∈ f /\ in(t, x)))) by RightForall
+
+      // Strategy:
+      // 1. ∃(a, (t, a) ∈ g)) <=> ∃(a, (t, a) ∈ f /\ t ∈ x)                [[restrictedFunctionPairMembership]]
+      // 2. ∃(a, (t, a) ∈ f /\ t ∈ x) <=> ∃(a, (t, a) ∈ f) /\ t ∈ x        [[existentialConjunctionWithClosedFormula]]
+      // 3. (∃(a, (t, a) ∈ f) /\ t ∈ x) ==> (∃!(a, (t, a) ∈ f) /\ t ∈ x)   [[fAppUniqueness]]
+      // 4. (∃!(a, (t, a) ∈ f) /\ t ∈ x) <=> ∃!(a, (t, a) ∈ f /\ t ∈ x)    [[uniqueExistentialConjunctionWithClosedFormula]]
+      // 5. ∃!(a, (t, a) ∈ f /\ t ∈ x) <=> ∃!(a, (t, a) ∈ g)               [[uniqueExistentialEquivalenceDistribution]]
+      val p = formulaVariable
+
+      val step1 = have(∃(a, pair(t, a) ∈ g) <=> ∃(a, pair(t, a) ∈ f /\ t ∈ x)) by Cut(
+        equiv,
+        existentialEquivalenceDistribution of (
+          P -> lambda(a, pair(t, a) ∈ g),
+          Q -> lambda(a, pair(t, a) ∈ f /\ t ∈ x)
+        )
+      )
+
+      val step2 = have((∃(a, pair(t, a) ∈ f) /\ t ∈ x) <=> (∃(a, pair(t, a) ∈ f) /\ t ∈ x)) by Tautology.from(
+        existentialConjunctionWithClosedFormula of (
+          P -> lambda(a, pair(t, a) ∈ f),
+          p -> lambda(Seq(), t ∈ x)
+        )
+      )
+
+      val step3 = have((∃(a, pair(t, a) ∈ f) /\ t ∈ x) ==> (∃!(a, pair(t, a) ∈ f) /\ t ∈ x)) by Tautology.from(fAppUniqueness)
+
+      val step4 = have((∃!(a, pair(t, a) ∈ f) /\ t ∈ x) <=> ∃!(a, pair(t, a) ∈ f /\ t ∈ x)) by Tautology.from(
+        uniqueExistentialConjunctionWithClosedFormula of (
+          P -> lambda(a, pair(t, a) ∈ f),
+          p -> lambda(Seq(), t ∈ x)
+        )
+      )
+
+      val step5 = have(∃!(a, pair(t, a) ∈ f /\ t ∈ x) <=> ∃!(a, pair(t, a) ∈ g)) by Tautology.from(
+        equiv,
+        uniqueExistentialEquivalenceDistribution of (
+          P -> lambda(a, pair(t, a) ∈ f /\ t ∈ x),
+          Q -> lambda(a, pair(t, a) ∈ g)
+        )
+      )
+
+      have(∃(y, pair(t, y) ∈ g) ==> ∃!(y, pair(t, y) ∈ g)) by Tautology.from(
+        step1,
+        step2,
+        step3,
+        step4,
+        step5
+      )
+      thenHave(∀(t, ∃(y, pair(t, y) ∈ g) ==> ∃!(y, pair(t, y) ∈ g))) by RightForall
+    }
+
+    have(relation(g) /\ ∀(t, ∃(y, in(pair(t, y), g)) ==> ∃!(y, in(pair(t, y), g)))) by RightAnd(gIsRelation, gAppUniqueness)
+    have(functional(g)) by Tautology.from(
+      lastStep,
+      functional.definition of (f -> g)
+    )
+  }
+
   /**
    * Restricted function domain -- For a function `f`, the domain of `f_x` is `x ∩ relationDomain(f)`.
    */
@@ -2473,6 +2585,68 @@ object SetTheory extends lisa.Main {
     have(thesis) by Tautology.from(domCharacterization, simplerCharacterization)
   }
 
+  /**
+   * Theorem --- A restricted function coincides with the original function on its domain.
+   * In other words, if `g = restrictedFunction(f, d)`, `x ∈ d`, then `g(x) = f(x)`.
+   */
+  val restrictedFunctionApplication = Theorem {
+    val g = restrictedFunction(f, d)
+    (functional(f), x ∈ relationDomain(g)) |- app(g, x) === app(f, x)
+  } {
+    val g = restrictedFunction(f, d)
+    val p = pair(x, app(g, x))
+
+    // Show that x ∈ relationDomain(g) ==> x ∈ relationDomain(f)
+    have(∀(x, x ∈ (d ∩ relationDomain(f)) <=> (x ∈ d /\ x ∈ relationDomain(f)))) by Definition(setIntersection, setIntersectionUniqueness)(d, relationDomain(f))
+    thenHave(x ∈ (d ∩ relationDomain(f)) <=> (x ∈ d /\ x ∈ relationDomain(f))) by InstantiateForall(x)
+    thenHave(x ∈ (d ∩ relationDomain(f)) ==> x ∈ relationDomain(f)) by Tautology
+    thenHave(x ∈ (d ∩ relationDomain(f)) |- x ∈ relationDomain(f)) by Restate
+    thenHave((x ∈ relationDomain(g), relationDomain(g) === (d ∩ relationDomain(f))) |- x ∈ relationDomain(f)) by LeftSubstEq(
+      List((relationDomain(g), (d ∩ relationDomain(f)))),
+      lambda(z, x ∈ z)
+    )
+    val domainInclusion = have(x ∈ relationDomain(g) |- x ∈ relationDomain(f)) by Cut(
+      restrictedFunctionDomain of (x -> d),
+      lastStep, 
+    )
+
+    // Characterize app(f, x)
+    val characterization = have(
+      (app(g, x) === app(f, x)) <=> (((functional(f) /\ (x ∈ relationDomain(f))) ==> (p ∈ f)) /\
+                                    ((!functional(f) \/ (x ∉ relationDomain(f))) ==> (app(g, x) === ∅)))
+    ) by InstantiateForall(app(g, x))(app.definition)
+
+    // Use the definition of restricted functions
+    have(∀(t, t ∈ g ==> t ∈ f)) by Tautology.from(
+      restrictedFunctionIsSubset of (x -> d),
+      subset.definition of (x -> g, y -> f)
+    )
+    thenHave(p ∈ g ==> p ∈ f) by InstantiateForall(p)
+    val gSubsetF = thenHave(p ∈ g |- p ∈ f) by Restate
+
+    // Show that (x, app(g, x)) ∈ f
+    have((functional(g), x ∈ relationDomain(g)) |- p ∈ g) by Definition(app, functionApplicationUniqueness)(g, x)
+    have((functional(f), x ∈ relationDomain(g)) |- p ∈ g) by Cut(restrictedFunctionIsFunctional of (x -> d), lastStep)
+    val membership = have((functional(f), x ∈ relationDomain(g)) |- p ∈ f) by Cut(lastStep, gSubsetF)
+
+    // Reconstruct the definition and conclude by the characterization
+    val premises = have((functional(f), x ∈ relationDomain(g)) |- (functional(f) /\ (x ∈ relationDomain(f)))) by Tautology.from(domainInclusion)
+    val pos = have((functional(f), x ∈ relationDomain(g)) |- ((functional(f) /\ (x ∈ relationDomain(f))) ==> (p ∈ f))) by Tautology.from(
+      premises,
+      membership
+    )
+
+    have((functional(f), x ∈ relationDomain(g), !(functional(f) /\ (x ∈ relationDomain(f)))) |- ()) by LeftNot(premises)
+    thenHave((functional(f), x ∈ relationDomain(g), !functional(f) \/ (x ∉ relationDomain(f))) |- ()) by Restate
+    thenHave((functional(f), x ∈ relationDomain(g), !functional(f) \/ (x ∉ relationDomain(f))) |- (app(g, x) === ∅)) by Weakening
+    val neg = thenHave((functional(f), x ∈ relationDomain(g)) |- (!functional(f) \/ (x ∉ relationDomain(f)) ==> (app(g, x) === ∅))) by Restate
+
+    have((functional(f), x ∈ relationDomain(g)) |- (((functional(f) /\ (x ∈ relationDomain(f))) ==> (p ∈ f)) /\
+                                                    ((!functional(f) \/ (x ∉ relationDomain(f))) ==> (app(g, x) === ∅)))) by RightAnd(pos, neg)
+
+    have(thesis) by Tautology.from(lastStep, characterization)
+  }
+
   /**
    * Restricted function cancellation --- Restricting a function to its relation domain does nothing.
    */

From f68d2df4df00718b38bf048108af3f25fcee0435 Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Thu, 15 Jun 2023 21:52:09 +0200
Subject: [PATCH 66/68] Rename subgroup operation

---
 .../scala/lisa/mathematics/GroupTheory.scala  | 162 +++++++++---------
 1 file changed, 79 insertions(+), 83 deletions(-)

diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala
index 0dcfd0e2..4a8c895d 100644
--- a/src/main/scala/lisa/mathematics/GroupTheory.scala
+++ b/src/main/scala/lisa/mathematics/GroupTheory.scala
@@ -779,6 +779,9 @@ object GroupTheory extends lisa.Main {
   // 2. Subgroups
   //
 
+  // By convention, this will always refer to the restricted operation.
+  private val ★ = restrictedFunction(*, cartesianProduct(H, H))
+
   /**
    * Subgroup --- `H` is a subgroup of `(G, *)` if `H` is a subset of `G`, such that the restriction of `*` to `H` is
    * a group law for `H`, i.e. `(H, *_H)` is a group.
@@ -861,11 +864,8 @@ object GroupTheory extends lisa.Main {
    * subgroup of `G`, then for elements `x, y ∈ H` we have `x ★ y = x * y`, where `★ = *_H`.
    */
   val subgroupOperation = Theorem(
-    (subgroup(H, G, *), x ∈ H, y ∈ H) |- (op(x, restrictedFunction(*, cartesianProduct(H, H)), y) === op(x, *, y))
+    (subgroup(H, G, *), x ∈ H, y ∈ H) |- (op(x, ★, y) === op(x, *, y))
   ) {
-    val H2 = cartesianProduct(H, H)
-    val ★ = restrictedFunction(*, H2)
-
       assume(subgroup(H, G, *))
     val groupG = have(group(G, *)) by Tautology.from(subgroup.definition)
     val groupH = have(group(H, ★)) by Tautology.from(subgroup.definition)
@@ -875,7 +875,7 @@ object GroupTheory extends lisa.Main {
     )
     have((functional(*), x ∈ H, y ∈ H) |- op(x, ★, y) === op(x, *, y)) by Cut(
       lastStep,
-      restrictedFunctionApplication of (f -> *, VariableLabel("d") -> H2, x -> pair(x, y))
+      restrictedFunctionApplication of (f -> *, VariableLabel("d") -> cartesianProduct(H, H), x -> pair(x, y))
     )
     have(thesis) by Tautology.from(
       lastStep,
@@ -888,10 +888,8 @@ object GroupTheory extends lisa.Main {
    * Lemma --- If `H` is a subgroup of `G`, then `e_H ∈ G`.
    */
   val subgroupIdentityInParent = Lemma(
-    subgroup(H, G, *) |- identity(H, restrictedFunction(*, cartesianProduct(H, H))) ∈ G 
+    subgroup(H, G, *) |- identity(H, ★) ∈ G 
   ) {
-    val ★ = restrictedFunction(*, cartesianProduct(H, H))
-    
     val identityInH = have(subgroup(H, G, *) |- identity(H, ★) ∈ H) by Tautology.from(
       subgroup.definition,
       identityInGroup of (G -> H, * -> ★)
@@ -912,9 +910,8 @@ object GroupTheory extends lisa.Main {
    * the parent group `(G, *)`.
    */
   val subgroupIdentity = Theorem(
-    subgroup(H, G, *) |- identity(H, restrictedFunction(*, cartesianProduct(H, H))) === identity(G, *)
+    subgroup(H, G, *) |- identity(H, ★) === identity(G, *)
   ) {
-    val ★ = restrictedFunction(*, cartesianProduct(H, H))
     val e_G = identity(G, *)
     val e_H = identity(H, ★)
 
@@ -968,7 +965,7 @@ object GroupTheory extends lisa.Main {
    * Theorem --- If `H` is a subgroup of `G`, then the inverse is the same as in the parent group. 
    */
   val subgroupInverse = Theorem(
-    (subgroup(H, G, *), x ∈ H) |- inverse(x, H, restrictedFunction(*, cartesianProduct(H, H))) === inverse(x, G, *)
+    (subgroup(H, G, *), x ∈ H) |- inverse(x, H, ★) === inverse(x, G, *)
   ) {
     assume(subgroup(H, G, *))
     assume(x ∈ H)
@@ -981,7 +978,6 @@ object GroupTheory extends lisa.Main {
     val xInG = thenHave(x ∈ G) by Tautology
 
     val groupG = have(group(G, *)) by Tautology.from(subgroup.definition)
-    val ★ = restrictedFunction(*, cartesianProduct(H, H))
     val groupH = have(group(H, ★)) by Tautology.from(subgroup.definition)
 
     val eG = identity(G, *)
@@ -1054,27 +1050,27 @@ object GroupTheory extends lisa.Main {
   //
 
   // Extra group composition law
-  val ★ = variable
+  val ** = variable
 
   /**
    * Definition --- A group homomorphism is a mapping `f: G -> H` from structures `G` and `H` equipped with binary operations `*` and `★` respectively,
-   * such that for all `x, y ∈ G`, we have* `f(x * y) = f(x) ★ f(y)`.
+   * such that for all `x, y ∈ G`, we have* `f(x * y) = f(x) ** f(y)`.
    * 
-   * In the following, "homomorphism" always stands for "group homomorphism", i.e. `(G, *)` and `(H, ★)` are groups.
+   * In the following, "homomorphism" always stands for "group homomorphism", i.e. `(G, *)` and `(H, **)` are groups.
    */
-  val homomorphism = DEF(f, G, *, H, ★) --> group(G, *) /\ group(H, ★) /\ functionFrom(f, G, H) /\ ∀(x, x ∈ G ==> ∀(y, y ∈ G ==> (app(f, op(x, *, y)) === op(app(f, x), ★, app(f, y)))))
+  val homomorphism = DEF(f, G, *, H, **) --> group(G, *) /\ group(H, **) /\ functionFrom(f, G, H) /\ ∀(x, x ∈ G ==> ∀(y, y ∈ G ==> (app(f, op(x, *, y)) === op(app(f, x), **, app(f, y)))))
 
   /**
    * Lemma --- Practical reformulation of the homomorphism definition.
    */
   val homomorphismApplication = Lemma(
-    (homomorphism(f, G, *, H, ★), x ∈ G, y ∈ G) |- app(f, op(x, *, y)) === op(app(f, x), ★, app(f, y))
+    (homomorphism(f, G, *, H, **), x ∈ G, y ∈ G) |- app(f, op(x, *, y)) === op(app(f, x), **, app(f, y))
   ) {
-    assume(homomorphism(f, G, *, H, ★))
-    have(∀(x, x ∈ G ==> ∀(y, y ∈ G ==> (app(f, op(x, *, y)) === op(app(f, x), ★, app(f, y)))))) by Tautology.from(homomorphism.definition)
-    thenHave(x ∈ G ==> ∀(y, y ∈ G ==> (app(f, op(x, *, y)) === op(app(f, x), ★, app(f, y))))) by InstantiateForall(x)
-    thenHave((x ∈ G) |- ∀(y, y ∈ G ==> (app(f, op(x, *, y)) === op(app(f, x), ★, app(f, y))))) by Restate
-    thenHave((x ∈ G) |- y ∈ G ==> (app(f, op(x, *, y)) === op(app(f, x), ★, app(f, y)))) by InstantiateForall(y)
+    assume(homomorphism(f, G, *, H, **))
+    have(∀(x, x ∈ G ==> ∀(y, y ∈ G ==> (app(f, op(x, *, y)) === op(app(f, x), **, app(f, y)))))) by Tautology.from(homomorphism.definition)
+    thenHave(x ∈ G ==> ∀(y, y ∈ G ==> (app(f, op(x, *, y)) === op(app(f, x), **, app(f, y))))) by InstantiateForall(x)
+    thenHave((x ∈ G) |- ∀(y, y ∈ G ==> (app(f, op(x, *, y)) === op(app(f, x), **, app(f, y))))) by Restate
+    thenHave((x ∈ G) |- y ∈ G ==> (app(f, op(x, *, y)) === op(app(f, x), **, app(f, y)))) by InstantiateForall(y)
     thenHave(thesis) by Restate
   }
 
@@ -1082,9 +1078,9 @@ object GroupTheory extends lisa.Main {
    * Lemma --- If `f` is a homomorphism, then `f(x) ∈ H` for all `x ∈ G`.
    */
   private val homomorphismAppInH = Lemma(
-    (homomorphism(f, G, *, H, ★), x ∈ G) |- app(f, x) ∈ H
+    (homomorphism(f, G, *, H, **), x ∈ G) |- app(f, x) ∈ H
   ) {
-    have(homomorphism(f, G, *, H, ★) |- functionFrom(f, G, H)) by Tautology.from(homomorphism.definition)
+    have(homomorphism(f, G, *, H, **) |- functionFrom(f, G, H)) by Tautology.from(homomorphism.definition)
     have(thesis) by Cut(
       lastStep,
       functionAppInCodomain of (VariableLabel("t") -> x, VariableLabel("x") -> G, y -> H)
@@ -1095,22 +1091,22 @@ object GroupTheory extends lisa.Main {
    * Theorem --- If `f` is a group-homomorphism between `G` and `H`, then `f(e_G) = e_H`.
    */
   val homomorphismMapsIdentityToIdentity = Theorem(
-    homomorphism(f, G, *, H, ★) |- app(f, identity(G, *)) === identity(H, ★)
+    homomorphism(f, G, *, H, **) |- app(f, identity(G, *)) === identity(H, **)
   ) {
     val e = identity(G, *)
 
-    val groupG = have(homomorphism(f, G, *, H, ★) |- group(G, *)) by Tautology.from(homomorphism.definition)
-    val groupH = have(homomorphism(f, G, *, H, ★) |- group(H, ★)) by Tautology.from(homomorphism.definition)
+    val groupG = have(homomorphism(f, G, *, H, **) |- group(G, *)) by Tautology.from(homomorphism.definition)
+    val groupH = have(homomorphism(f, G, *, H, **) |- group(H, **)) by Tautology.from(homomorphism.definition)
 
-    val identityInG = have(homomorphism(f, G, *, H, ★) |- e ∈ G) by Cut(groupG, identityInGroup)
-    val appInH = have(homomorphism(f, G, *, H, ★) |- app(f, e) ∈ H) by Cut(identityInG, homomorphismAppInH of (x -> e))
+    val identityInG = have(homomorphism(f, G, *, H, **) |- e ∈ G) by Cut(groupG, identityInGroup)
+    val appInH = have(homomorphism(f, G, *, H, **) |- app(f, e) ∈ H) by Cut(identityInG, homomorphismAppInH of (x -> e))
 
     // 0. e * e = e (to apply substitution)
     have(group(G, *) |- op(e, *, e) === e) by Cut(
       identityInGroup,
       identityIdempotence of (x -> e)
     )
-    val eq0 = have(homomorphism(f, G, *, H, ★) |- op(e, *, e) === e) by Cut(groupG, lastStep)
+    val eq0 = have(homomorphism(f, G, *, H, **) |- op(e, *, e) === e) by Cut(groupG, lastStep)
     
     // 1. f(e * e) = f(e)
     have(app(f, e) === app(f, e)) by RightRefl
@@ -1118,23 +1114,23 @@ object GroupTheory extends lisa.Main {
       List((op(e, *, e), e)),
       lambda(z, app(f, z) === app(f, e))
     )
-    val eq1 = have(homomorphism(f, G, *, H, ★) |- app(f, op(e, *, e)) === app(f, e)) by Cut(eq0, lastStep)
+    val eq1 = have(homomorphism(f, G, *, H, **) |- app(f, op(e, *, e)) === app(f, e)) by Cut(eq0, lastStep)
 
-    // 2. f(e * e) = f(e) ★ f(e)
-    val eq2 = have(homomorphism(f, G, *, H, ★) |- app(f, op(e, *, e)) === op(app(f, e), ★, app(f, e))) by Cut(
+    // 2. f(e * e) = f(e) ** f(e)
+    val eq2 = have(homomorphism(f, G, *, H, **) |- app(f, op(e, *, e)) === op(app(f, e), **, app(f, e))) by Cut(
       identityInG,
       homomorphismApplication of (x -> e, y -> e)
     )
 
-    // 3. f(e) ★ f(e) = f(e)
-    val eq3 = have(homomorphism(f, G, *, H, ★) |- op(app(f, e), ★, app(f, e)) === app(f, e)) by Equalities(eq1, eq2)
+    // 3. f(e) ** f(e) = f(e)
+    val eq3 = have(homomorphism(f, G, *, H, **) |- op(app(f, e), **, app(f, e)) === app(f, e)) by Equalities(eq1, eq2)
 
     // Conclude by idempotence
-    have((homomorphism(f, G, *, H, ★), app(f, e) ∈ H) |- (op(app(f, e), ★, app(f, e)) === app(f, e)) <=> (app(f, e) === identity(H, ★))) by Cut(
+    have((homomorphism(f, G, *, H, **), app(f, e) ∈ H) |- (op(app(f, e), **, app(f, e)) === app(f, e)) <=> (app(f, e) === identity(H, **))) by Cut(
       groupH,
-      identityIdempotence of (x -> app(f, e), G -> H, * -> ★)
+      identityIdempotence of (x -> app(f, e), G -> H, * -> **)
     )
-    have(homomorphism(f, G, *, H, ★)|- (op(app(f, e), ★, app(f, e)) === app(f, e)) <=> (app(f, e) === identity(H, ★))) by Cut(
+    have(homomorphism(f, G, *, H, **)|- (op(app(f, e), **, app(f, e)) === app(f, e)) <=> (app(f, e) === identity(H, **))) by Cut(
       appInH,
       lastStep
     )
@@ -1143,24 +1139,24 @@ object GroupTheory extends lisa.Main {
   }
   
   /**
-   * Theorem --- If `f: G -> H` is a group homomorphism, then `f(inverse(x, G, *)) = inverse(f(x), H, ★)`.
+   * Theorem --- If `f: G -> H` is a group homomorphism, then `f(inverse(x, G, *)) = inverse(f(x), H, **)`.
    */
   val homomorphismMapsInverseToInverse = Theorem(
-    (homomorphism(f, G, *, H, ★), x ∈ G) |- app(f, inverse(x, G, *)) === inverse(app(f, x), H, ★)
+    (homomorphism(f, G, *, H, **), x ∈ G) |- app(f, inverse(x, G, *)) === inverse(app(f, x), H, **)
   ) {
-    assume(homomorphism(f, G, *, H, ★))
+    assume(homomorphism(f, G, *, H, **))
     assume(x ∈ G)
 
     val groupG = have(group(G, *)) by Tautology.from(homomorphism.definition)
-    val groupH = have(group(H, ★)) by Tautology.from(homomorphism.definition)
+    val groupH = have(group(H, **)) by Tautology.from(homomorphism.definition)
 
     val eG = identity(G, *)
-    val eH = identity(H, ★)
+    val eH = identity(H, **)
     val i = inverse(x, G, *)
     val iInG = have(i ∈ G) by Cut(groupG, inverseInGroup)
 
     // 1. f(x * inverse(x)) = f(x) f(inverse(x))
-    val eq1 = have(app(f, op(x, *, i)) === op(app(f, x), ★, app(f, i))) by Cut(
+    val eq1 = have(app(f, op(x, *, i)) === op(app(f, x), **, app(f, i))) by Cut(
       iInG,
       homomorphismApplication of (y -> i)
     )
@@ -1183,12 +1179,12 @@ object GroupTheory extends lisa.Main {
     val eq3 = have(app(f, eG) === eH) by Tautology.from(homomorphismMapsIdentityToIdentity)
 
     // 4. f(x)f(inverse(x)) = e'
-    val eq4 = have(op(app(f, x), ★, app(f, i)) === eH) by Equalities(eq1, eq2, eq3)
+    val eq4 = have(op(app(f, x), **, app(f, i)) === eH) by Equalities(eq1, eq2, eq3)
 
     // Conclude
-    val conclusion = have((app(f, i) ∈ H) |- (app(f, i) === inverse(app(f, x), H, ★))) by Tautology.from(
+    val conclusion = have((app(f, i) ∈ H) |- (app(f, i) === inverse(app(f, x), H, **))) by Tautology.from(
       groupH,
-      inverseTest of (G -> H, * -> ★, x -> app(f, x), y -> app(f, i)),
+      inverseTest of (G -> H, * -> **, x -> app(f, x), y -> app(f, i)),
       eq4,
       homomorphismAppInH
     )
@@ -1203,15 +1199,15 @@ object GroupTheory extends lisa.Main {
    * Kernel uniqueness --- The kernel of a homomorphism is well-defined.
    */
   val kernelUniqueness = Theorem(
-    homomorphism(f, G, *, H, ★) |- ∃!(z, ∀(t, (t ∈ z) <=> (t ∈ G /\ (app(f, t) === identity(H, ★)))))
+    homomorphism(f, G, *, H, **) |- ∃!(z, ∀(t, (t ∈ z) <=> (t ∈ G /\ (app(f, t) === identity(H, **)))))
   ) {
     // We apply the comprehension axiom here.
     // It might seem odd that the homomorphism assumption is not needed for the set to be defined,
     // but remember that [[app]] and [[identity]] default to the empty set when the assumptions are not met.
     // We add the assumption of `f` being a homomorphism to discard any value when the assumptions do not hold.
-    have(∃!(z, ∀(t, (t ∈ z) <=> (t ∈ G /\ (app(f, t) === identity(H, ★)))))) by UniqueComprehension(
+    have(∃!(z, ∀(t, (t ∈ z) <=> (t ∈ G /\ (app(f, t) === identity(H, **)))))) by UniqueComprehension(
       G,
-      lambda(Seq(t, G), app(f, t) === identity(H, ★))
+      lambda(Seq(t, G), app(f, t) === identity(H, **))
     )
     thenHave(thesis) by Weakening
   }
@@ -1219,19 +1215,19 @@ object GroupTheory extends lisa.Main {
   /**
    * Kernel --- The kernel of a homomorphism `f: G -> H` is the set of elements `t ∈ G` such that `f(t) = e_H`.
    */
-  val kernel = DEF(f, G, *, H, ★) --> TheConditional(z, ∀(t, (t ∈ z) <=> (t ∈ G /\ (app(f, t) === identity(H, ★)))))(kernelUniqueness)
+  val kernel = DEF(f, G, *, H, **) --> TheConditional(z, ∀(t, (t ∈ z) <=> (t ∈ G /\ (app(f, t) === identity(H, **)))))(kernelUniqueness)
 
   // Shortcut alias
-  private val K = kernel(f, G, *, H, ★)
+  private val K = kernel(f, G, *, H, **)
 
   /**
    * Lemma --- Reformulation of the kernel definition.
    */
   private val kernelDef = Lemma(
-    homomorphism(f, G, *, H, ★) |- (x ∈ K) <=> (x ∈ G /\ (app(f, x) === identity(H, ★)))
+    homomorphism(f, G, *, H, **) |- (x ∈ K) <=> (x ∈ G /\ (app(f, x) === identity(H, **)))
   ) {
-    assume(homomorphism(f, G, *, H, ★))
-    have(∀(t, (t ∈ K) <=> (t ∈ G /\ (app(f, t) === identity(H, ★))))) by Definition(kernel, kernelUniqueness)(f, G, *, H, ★)
+    assume(homomorphism(f, G, *, H, **))
+    have(∀(t, (t ∈ K) <=> (t ∈ G /\ (app(f, t) === identity(H, **))))) by Definition(kernel, kernelUniqueness)(f, G, *, H, **)
     thenHave(thesis) by InstantiateForall(x)
   }
 
@@ -1239,44 +1235,44 @@ object GroupTheory extends lisa.Main {
    * Lemma --- The kernel is closed by products, i.e. if `x, y ∈ K`, then `x * y ∈ K`.
    */
   val kernelIsClosedByProducts = Lemma(
-    (homomorphism(f, G, *, H, ★), x ∈ K, y ∈ K) |- op(x, *, y) ∈ K
+    (homomorphism(f, G, *, H, **), x ∈ K, y ∈ K) |- op(x, *, y) ∈ K
   ) {
-    assume(homomorphism(f, G, *, H, ★))
+    assume(homomorphism(f, G, *, H, **))
     assume(x ∈ K)
     assume(y ∈ K)
 
     val elemInG = have(x ∈ G) by Tautology.from(kernelDef)
 
     val groupG = have(group(G, *)) by Tautology.from(homomorphism.definition)
-    val groupH = have(group(H, ★)) by Tautology.from(homomorphism.definition)
+    val groupH = have(group(H, **)) by Tautology.from(homomorphism.definition)
 
-    val e = identity(H, ★)
-    val eInH = have(e ∈ H) by Cut(groupH, identityInGroup of (G -> H, * -> ★))
+    val e = identity(H, **)
+    val eInH = have(e ∈ H) by Cut(groupH, identityInGroup of (G -> H, * -> **))
 
-    // 1. f(x) ★ f(y) = f(x * y)
-    val eq1 = have(app(f, op(x, *, y)) === op(app(f, x), ★, app(f, y))) by Tautology.from(
+    // 1. f(x) ** f(y) = f(x * y)
+    val eq1 = have(app(f, op(x, *, y)) === op(app(f, x), **, app(f, y))) by Tautology.from(
       homomorphismApplication,
       elemInG,
       elemInG of (x -> y)
     )
 
-    // 2. f(x) ★ f(y) = e ★ e
+    // 2. f(x) ** f(y) = e ** e
     val appValue = have(app(f, x) === e) by Tautology.from(kernelDef)
-    have(op(app(f, x), ★, app(f, y)) === op(app(f, x), ★, app(f, y))) by RightRefl
-    thenHave((app(f, x) === e, app(f, y) === e) |- op(app(f, x), ★, app(f, y)) === op(e, ★, e)) by RightSubstEq(
+    have(op(app(f, x), **, app(f, y)) === op(app(f, x), **, app(f, y))) by RightRefl
+    thenHave((app(f, x) === e, app(f, y) === e) |- op(app(f, x), **, app(f, y)) === op(e, **, e)) by RightSubstEq(
       List((app(f, x), e), (app(f, y), e)),
-      lambda(Seq(a, b), op(app(f, x), ★, app(f, y)) === op(a, ★, b))
+      lambda(Seq(a, b), op(app(f, x), **, app(f, y)) === op(a, **, b))
     )
 
-    val eq2 = have(op(app(f, x), ★, app(f, y)) === op(e, ★, e)) by Tautology.from(
+    val eq2 = have(op(app(f, x), **, app(f, y)) === op(e, **, e)) by Tautology.from(
       lastStep,
       appValue,
       appValue of (x -> y)
     )
 
-    // 3. e ★ e = e
-    val eq3 = have(op(e, ★, e) === e) by Tautology.from(
-      identityNeutrality of (G -> H, * -> ★, x -> e),
+    // 3. e ** e = e
+    val eq3 = have(op(e, **, e) === e) by Tautology.from(
+      identityNeutrality of (G -> H, * -> **, x -> e),
       groupH,
       eInH
     )
@@ -1300,32 +1296,32 @@ object GroupTheory extends lisa.Main {
    * Lemma --- The kernel is closed by inversion, i.e. if `x ∈ K` then `inverse(x, G, *) ∈ K`.
    */
   val kernelIsClosedByInversion = Lemma (
-    (homomorphism(f, G, *, H, ★), x ∈ K) |- inverse(x, G, *) ∈ K
+    (homomorphism(f, G, *, H, **), x ∈ K) |- inverse(x, G, *) ∈ K
    ) {
-    assume(homomorphism(f, G, *, H, ★))
+    assume(homomorphism(f, G, *, H, **))
     assume(x ∈ K)
 
     val groupG = have(group(G, *)) by Tautology.from(homomorphism.definition)
-    val groupH = have(group(H, ★)) by Tautology.from(homomorphism.definition)
+    val groupH = have(group(H, **)) by Tautology.from(homomorphism.definition)
     val elemInG = have(x ∈ G) by Tautology.from(kernelDef)
 
-    val e = identity(H, ★)
+    val e = identity(H, **)
     val appValue = have(app(f, x) === e) by Tautology.from(kernelDef)
 
     // 1. f(inverse(x)) = inverse(f(x)) = inverse(e)
-    have(app(f, inverse(x, G, *)) === inverse(app(f, x), H, ★)) by Tautology.from(
+    have(app(f, inverse(x, G, *)) === inverse(app(f, x), H, **)) by Tautology.from(
       homomorphismMapsInverseToInverse,
       elemInG
     )
-    thenHave((app(f, x) === e) |- (app(f, inverse(x, G, *)) === inverse(e, H, ★))) by RightSubstEq(
+    thenHave((app(f, x) === e) |- (app(f, inverse(x, G, *)) === inverse(e, H, **))) by RightSubstEq(
       List((app(f, x), e)),
-      lambda(z, app(f, inverse(x, G, *)) === inverse(z, H, ★))
+      lambda(z, app(f, inverse(x, G, *)) === inverse(z, H, **))
     )
     
-    val eq1 = have(app(f, inverse(x, G, *)) === inverse(e, H, ★)) by Cut(appValue, lastStep)
+    val eq1 = have(app(f, inverse(x, G, *)) === inverse(e, H, **)) by Cut(appValue, lastStep)
 
     // 2. inverse(e) = e
-    val eq2 = have(inverse(e, H, ★) === e) by Cut(groupH, inverseOfIdentityIsIdentity of (G -> H, * -> ★))
+    val eq2 = have(inverse(e, H, **) === e) by Cut(groupH, inverseOfIdentityIsIdentity of (G -> H, * -> **))
 
     // 3. Conclude
     val eq3 = have(app(f, inverse(x, G, *)) === e) by Equalities(eq1, eq2)
@@ -1344,9 +1340,9 @@ object GroupTheory extends lisa.Main {
    * Theorem --- The kernel of a homomorphism `f: G -> H` is a subgroup of `G`.
    */
   val kernelIsSubgroup = Theorem(
-    homomorphism(f, G, *, H, ★) |- subgroup(kernel(f, G, *, H, ★), G, *)
+    homomorphism(f, G, *, H, **) |- subgroup(kernel(f, G, *, H, **), G, *)
   ) {
-    assume(homomorphism(f, G, *, H, ★))
+    assume(homomorphism(f, G, *, H, **))
     val groupG = have(group(G, *)) by Tautology.from(homomorphism.definition)
 
     // We show that the kernel satisfies all requirements of [[subgroupCondition]]
@@ -1356,7 +1352,7 @@ object GroupTheory extends lisa.Main {
 
     // 1. kernel != ∅
     have(identity(G, *) ∈ G) by Cut(groupG, identityInGroup)
-    have(identity(G, *) ∈ G /\ (app(f, identity(G, *)) === identity(H, ★))) by RightAnd(
+    have(identity(G, *) ∈ G /\ (app(f, identity(G, *)) === identity(H, **))) by RightAnd(
       lastStep,
       homomorphismMapsIdentityToIdentity
     )
@@ -1397,7 +1393,7 @@ object GroupTheory extends lisa.Main {
    * 
    * In some sense, isomorphic groups are equivalent, up to relabelling their elements.
    */
-  val isomorphism = DEF(f, G, *, H, ★) --> homomorphism(f, G, *, H, ★) /\ bijective(f, G, H)
+  val isomorphism = DEF(f, G, *, H, **) --> homomorphism(f, G, *, H, **) /\ bijective(f, G, H)
 
   /**
    * Automorphism --- An automorphism is an isomorphism from a group to itself.

From 450985df6de4ee6e7126275ff392110f06ae8a38 Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Fri, 16 Jun 2023 14:37:00 +0200
Subject: [PATCH 67/68] Prove main subgroup test theorem

---
 .../scala/lisa/mathematics/GroupTheory.scala  | 542 +++++++++++++++++-
 .../scala/lisa/mathematics/SetTheory.scala    |  64 +++
 2 files changed, 588 insertions(+), 18 deletions(-)

diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala
index 4a8c895d..cdce4d64 100644
--- a/src/main/scala/lisa/mathematics/GroupTheory.scala
+++ b/src/main/scala/lisa/mathematics/GroupTheory.scala
@@ -25,7 +25,7 @@ object GroupTheory extends lisa.Main {
   private val * = variable
 
   // Group elements
-  private val a, b = variable
+  private val a, b, c, d = variable
   private val x, y, z = variable
   private val t, u, v, w = variable
 
@@ -866,7 +866,7 @@ object GroupTheory extends lisa.Main {
   val subgroupOperation = Theorem(
     (subgroup(H, G, *), x ∈ H, y ∈ H) |- (op(x, ★, y) === op(x, *, y))
   ) {
-      assume(subgroup(H, G, *))
+    assume(subgroup(H, G, *))
     val groupG = have(group(G, *)) by Tautology.from(subgroup.definition)
     val groupH = have(group(H, ★)) by Tautology.from(subgroup.definition)
 
@@ -875,7 +875,7 @@ object GroupTheory extends lisa.Main {
     )
     have((functional(*), x ∈ H, y ∈ H) |- op(x, ★, y) === op(x, *, y)) by Cut(
       lastStep,
-      restrictedFunctionApplication of (f -> *, VariableLabel("d") -> cartesianProduct(H, H), x -> pair(x, y))
+      restrictedFunctionApplication of (f -> *, d -> cartesianProduct(H, H), x -> pair(x, y))
     )
     have(thesis) by Tautology.from(
       lastStep,
@@ -1023,8 +1023,460 @@ object GroupTheory extends lisa.Main {
     )
   }
 
+  //
+  // 2.1 Main subgroup test
+  // 
+  // We define several useful lemmas to attack this easy, but long theorem to formalize
+  // 
+
+  private val nonEmpty = H !== ∅
+  private val closedByProducts = ∀(x, ∀(y, (x ∈ H /\ y ∈ H) ==> (op(x, *, y) ∈ H)))
+  private val closedByInverses = ∀(x, x ∈ H ==> (inverse(x, G, *) ∈ H))
+  private val subgroupConditions = nonEmpty /\ closedByProducts /\ closedByInverses
+
+  /**
+   * Lemma --- Reformulation of the subset definition.
+   */
+  private val subgroupConditionsSubset = Lemma(
+    (subset(H, G), x ∈ H) |- x ∈ G
+  ) {
+    assume(subset(H, G))
+    have(∀(x, x ∈ H ==> x ∈ G)) by Tautology.from(subset.definition of (x -> H, y -> G))
+    thenHave(x ∈ H ==> x ∈ G) by InstantiateForall(x)
+    thenHave(x ∈ H |- x ∈ G) by Restate
+  }
+
+  /**
+   * Lemma --- The subgroup conditions imply that `relationDomain(★) === cartesianProduct(H, H)`.
+   */
+  private val subgroupConditionsDomain = Lemma(
+    (group(G, *), subset(H, G), subgroupConditions) |- relationDomain(★) === cartesianProduct(H, H)
+  ) {
+    val H2 = cartesianProduct(H, H)
+    val G2 = cartesianProduct(G, G)
+
+    assume(group(G, *))
+    assume(subset(H, G))
+    assume(subgroupConditions)
+
+    have(relationDomain(★) === (H2 ∩ relationDomain(*))) by Tautology.from(restrictedFunctionDomain of (f -> *, x -> H2))
+    thenHave((relationDomain(*) === G2) |- relationDomain(★) === (H2 ∩ G2)) by RightSubstEq(
+      List((relationDomain(*), G2)),
+      lambda(z, relationDomain(★) === (H2 ∩ z))
+    )
+    val eq1 = have(relationDomain(★) === (H2 ∩ G2)) by Cut(groupOperationDomain, lastStep)
+
+    // Prove that (H2 ∩ G2) = H2
+    have(subset(H2, G2)) by Tautology.from(subsetsCartesianProduct of (a -> H, b -> G, c -> H, d -> G))
+    val eq2 = have((H2 ∩ G2) === H2) by Cut(
+      lastStep,
+      setIntersectionSubset of (x -> H2, y -> G2)
+    )
+
+    have(thesis) by Equalities(eq1, eq2)
+  }
+
+  /**
+   * Lemma --- The subgroup conditions imply that `(x, y)` is in the relation domain of `★`.
+   * 
+   * Analogous to [[groupPairInOperationDomain]].
+   */
+  private val subgroupConditionsPairInDomain = Lemma(
+    (group(G, *), subset(H, G), subgroupConditions, x ∈ H, y ∈ H) |- pair(x, y) ∈ relationDomain(★) 
+  ) {
+    assume(group(G, *))
+    assume(subset(H, G))
+    assume(subgroupConditions)
+    assume(x ∈ H)
+    assume(y ∈ H)
+
+    have(pair(x, y) ∈ cartesianProduct(H, H)) by Tautology.from(
+      pairInCartesianProduct of (a -> x, b -> y, x -> H, y -> H)
+    )
+    thenHave((relationDomain(★) === cartesianProduct(H, H)) |- pair(x, y) ∈ relationDomain(★)) by RightSubstEq(
+      List((relationDomain(★), cartesianProduct(H, H))),
+      lambda(z, pair(x, y) ∈ z)
+    )
+
+    have(thesis) by Cut(subgroupConditionsDomain, lastStep)
+  }
+
+  /**
+   * Lemma --- The subgroup conditions imply that `x ★ y = x * y`.
+   * 
+   * Analogous to [[subgroupOperation]].
+   */
+  private val subgroupConditionsOperation = Lemma(
+    (group(G, *), subset(H, G), subgroupConditions, x ∈ H, y ∈ H) |- op(x, ★, y) === op(x, *, y)
+  ) {
+    have(thesis) by Tautology.from(
+      subgroupConditionsPairInDomain,
+      groupOperationIsFunctional,
+      restrictedFunctionIsFunctional of (f -> *, x -> cartesianProduct(H, H)),
+      restrictedFunctionApplication of (f -> *, d -> cartesianProduct(H, H), x -> pair(x, y))
+    )
+  }
+
   /**
-   * Theorem --- A subset `H ⊆ G` of a group `(G, *)` is a subgroup if and only if:
+   * Lemma --- The subgroup conditions imply that `x ★ y ∈ H`.
+   */
+  private val subgroupConditionsProductClosure = Lemma(
+    (group(G, *), subset(H, G), subgroupConditions, x ∈ H, y ∈ H) |- op(x, ★, y) ∈ H
+  ) {
+    assume(group(G, *))
+    assume(subset(H, G))
+    assume(subgroupConditions)
+
+    have(closedByProducts) by Tautology
+    thenHave(∀(y, (x ∈ H /\ y ∈ H) ==> (op(x, *, y) ∈ H))) by InstantiateForall(x)
+    thenHave((x ∈ H /\ y ∈ H) ==> (op(x, *, y) ∈ H)) by InstantiateForall(y)
+    thenHave((x ∈ H, y ∈ H) |- (op(x, *, y) ∈ H)) by Restate
+    thenHave((x ∈ H, y ∈ H, op(x, ★, y) === op(x, *, y)) |- (op(x, ★, y) ∈ H)) by RightSubstEq(
+      List((op(x, ★, y), op(x, *, y))),
+      lambda(z, z ∈ H)
+    )
+
+    have(thesis) by Cut(subgroupConditionsOperation, lastStep)
+  }
+
+  /**
+   * Lemma --- The subgroup conditions imply that `★` is a binary relation on `H`.
+   */
+  private val subgroupConditionsBinaryRelation = Lemma(
+    (group(G, *), subset(H, G), subgroupConditions) |- binaryOperation(H, ★)
+  ) {
+    assume(group(G, *))
+    assume(subset(H, G))
+    assume(subgroupConditions)
+
+    val H2 = cartesianProduct(H, H)
+    val r = variable
+
+    have(∀(t, (t ∈ setOfFunctions(H2, H)) <=> (t ∈ powerSet(cartesianProduct(H2, H)) /\ functionalOver(t, H2)))) by Definition(setOfFunctions, setOfFunctionsUniqueness)(H2, H)
+    val setOfFunDef = thenHave((★ ∈ setOfFunctions(H2, H)) <=> (★ ∈ powerSet(cartesianProduct(H2, H)) /\ functionalOver(★, H2))) by InstantiateForall(★)
+
+    val fun = have(functional(★)) by Tautology.from(
+      groupOperationIsFunctional,
+      restrictedFunctionIsFunctional of (f -> *, x -> H2)
+    )
+    have(functional(★) /\ (relationDomain(★) === H2)) by RightAnd(fun, subgroupConditionsDomain)
+    val funOver = have(functionalOver(★, H2)) by Tautology.from(lastStep, functionalOver.definition of (f -> ★, x -> H2))
+
+    have(subset(★, cartesianProduct(relationDomain(★), relationRange(★)))) by Tautology.from(
+      fun,
+      functional.definition of (f -> ★),
+      relation.definition of (r -> ★),
+      relationImpliesRelationBetweenDomainAndRange of (r -> ★),
+      relationBetween.definition of (r -> ★, a -> relationDomain(★), b -> relationRange(★))
+    )
+    thenHave((relationDomain(★) === H2) |- subset(★, cartesianProduct(H2, relationRange(★)))) by RightSubstEq(
+      List((relationDomain(★), H2)),
+      lambda(z, subset(★, cartesianProduct(z, relationRange(★))))
+    )
+
+    val subsetDomRange = have(subset(★, cartesianProduct(H2, relationRange(★)))) by Cut(
+      subgroupConditionsDomain,
+      lastStep
+    )
+
+    // Prove that ★ is a subset of H2 x H
+    val left = have(subset(H2, H2)) by Tautology.from(subsetReflexivity of (x -> H2))
+    val right = have(subset(relationRange(★), H)) subproof {
+      // Use pullback to characterize t
+      val pullback = have(t ∈ relationRange(★) |- ∃(x, (x ∈ relationDomain(★)) /\ (app(★, x) === t))) by Tautology.from(
+        groupOperationIsFunctional,
+        restrictedFunctionIsFunctional of (f -> *, x -> H2),
+        inRangeImpliesPullbackExists of (f -> ★, z -> t)
+      )
+
+      have((x ∈ relationDomain(★)) <=> (x ∈ relationDomain(★))) by Restate
+      thenHave((relationDomain(★) === H2) |- (x ∈ relationDomain(★)) <=> (x ∈ H2)) by RightSubstEq(
+        List((relationDomain(★), H2)),
+        lambda(z, (x ∈ relationDomain(★)) <=> (x ∈ z))
+      )
+      val equiv1 = have((x ∈ relationDomain(★)) <=> (x ∈ H2)) by Cut(subgroupConditionsDomain, lastStep)
+      val equiv2 = have((x ∈ H2) <=> ∃(a, ∃(b, (x === pair(a, b)) /\ in(a, H) /\ in(b, H)))) by Tautology.from(
+        elemOfCartesianProduct of (t -> x, x -> H, y -> H)
+      )
+
+      // Use closure by products to show that app(★, x) ∈ H
+      have(closedByProducts) by Tautology
+      thenHave(∀(y, (a ∈ H /\ y ∈ H) ==> (op(a, *, y) ∈ H))) by InstantiateForall(a)
+      thenHave((a ∈ H /\ b ∈ H) ==> (op(a, *, b) ∈ H)) by InstantiateForall(b)
+      thenHave((a ∈ H, b ∈ H) |- (op(a, *, b) ∈ H)) by Restate
+      thenHave((a ∈ H, b ∈ H, op(a, ★, b) === op(a, *, b)) |- (op(a, ★, b) ∈ H)) by RightSubstEq(
+        List((op(a, ★, b), op(a, *, b))),
+        lambda(z, z ∈ H)
+      )
+
+      have((a ∈ H, b ∈ H) |- (op(a, ★, b) ∈ H)) by Cut(
+        subgroupConditionsOperation of (x -> a, y -> b),
+        lastStep
+      )
+      thenHave((x === pair(a, b), a ∈ H, b ∈ H) |- (app(★, x) ∈ H)) by RightSubstEq(
+        List((x, pair(a, b))),
+        lambda(z, app(★, z) ∈ H)
+      )
+      thenHave(((x === pair(a, b)) /\ a ∈ H /\ b ∈ H) |- (app(★, x) ∈ H)) by Restate
+      thenHave(∃(b, (x === pair(a, b)) /\ a ∈ H /\ b ∈ H) |- (app(★, x) ∈ H)) by LeftExists
+      thenHave(∃(a, ∃(b, (x === pair(a, b)) /\ a ∈ H /\ b ∈ H)) |- (app(★, x) ∈ H)) by LeftExists
+
+      have((x ∈ relationDomain(★)) |- (app(★, x) ∈ H)) by Tautology.from(lastStep, equiv1, equiv2)
+      thenHave((x ∈ relationDomain(★), app(★, x) === t) |- (t ∈ H)) by RightSubstEq(
+        List((app(★, x), t)),
+        lambda(z, z ∈ H)
+      )
+      thenHave((x ∈ relationDomain(★) /\ (app(★, x) === t)) |- (t ∈ H)) by Restate
+      thenHave(∃(x, x ∈ relationDomain(★) /\ (app(★, x) === t)) |- (t ∈ H)) by LeftExists
+
+      have(t ∈ relationRange(★) |- t ∈ H) by Cut(pullback, lastStep)
+      thenHave(t ∈ relationRange(★) ==> t ∈ H) by Restate
+      thenHave(∀(t, t ∈ relationRange(★) ==> t ∈ H)) by RightForall
+
+      have(thesis) by Tautology.from(lastStep, subset.definition of (x -> relationRange(★), y -> H))
+    }
+
+    have(subset(cartesianProduct(H2, relationRange(★)), cartesianProduct(H2, H))) by Tautology.from(
+      left,
+      right,
+      subsetsCartesianProduct of (a -> H2, b -> H2, c -> relationRange(★), d -> H)
+    )
+    have(subset(★, cartesianProduct(H2, H))) by Tautology.from(
+      lastStep,
+      subsetDomRange,
+      subsetTransitivity of (a -> ★, b -> cartesianProduct(H2, relationRange(★)), c -> cartesianProduct(H2, H))
+    )
+    have(★ ∈ powerSet(cartesianProduct(H2, H))) by Tautology.from(
+      lastStep,
+      powerSet.definition of (x -> ★, y -> cartesianProduct(H2, H))
+    )
+
+    have(★ ∈ powerSet(cartesianProduct(H2, H)) /\ functionalOver(★, H2)) by RightAnd(lastStep, funOver)
+
+    have(thesis) by Tautology.from(
+      lastStep,
+      setOfFunDef,
+      functionFrom.definition of (f -> ★, x -> H2, y -> H),
+      binaryOperation.definition of (G -> H, * -> ★)
+    )
+  }
+
+  /**
+   * Lemma --- The subgroup conditions imply associativity on `H`.
+   * 
+   * This directly follows from associativity on `G` and [[subgroupConditionsOperation]].
+   */
+  private val subgroupConditionsAssociativity = Lemma(
+    (group(G, *), subset(H, G), subgroupConditions) |- associativityAxiom(H, ★)
+  ) {
+    assume(group(G, *))
+    assume(subset(H, G))
+    assume(subgroupConditions)
+
+    have((x ∈ H, y ∈ H, z ∈ H) |- op(op(x, ★, y), ★, z) === op(x, ★, op(y, ★, z))) subproof {
+      assume(x ∈ H)
+      assume(y ∈ H)
+      assume(z ∈ H)
+
+      have(op(op(x, *, y), *, z) === op(x, *, op(y, *, z))) by Tautology.from(
+        associativity,
+        subgroupConditionsSubset,
+        subgroupConditionsSubset of (x -> y),
+        subgroupConditionsSubset of (x -> z)
+      )
+      thenHave((op(x, ★, y) === op(x, *, y), op(y, ★, z) === op(y, *, z)) |- (op(op(x, ★, y), *, z) === op(x, *, op(y, ★, z)))) by RightSubstEq(
+        List((op(x, ★, y), op(x, *, y)), (op(y, ★, z), op(y, *, z))),
+        lambda(Seq(a, b), op(a, *, z) === op(x, *, b))
+      )
+
+      have(op(op(x, ★, y), *, z) === op(x, *, op(y, ★, z))) by Tautology.from(
+        lastStep,
+        subgroupConditionsOperation,
+        subgroupConditionsOperation of (x -> y, y -> z)
+      )
+      thenHave((op(op(x, ★, y), ★, z) === op(op(x, ★, y), *, z), op(x, ★, op(y, ★, z)) === op(x, *, op(y, ★, z))) |- (op(op(x, ★, y), ★, z) === op(x, ★, op(y, ★, z)))) by RightSubstEq(
+        List((op(op(x, ★, y), ★, z), op(op(x, ★, y), *, z)), (op(x, ★, op(y, ★, z)), op(x, *, op(y, ★, z)))),
+        lambda(Seq(a, b), a === b)
+      )
+
+      have(op(op(x, ★, y), ★, z) === op(x, ★, op(y, ★, z))) by Tautology.from(
+        lastStep,
+        subgroupConditionsOperation of (x -> op(x, ★, y), y -> z),
+        subgroupConditionsOperation of (x -> x, y -> op(y, ★, z)),
+        subgroupConditionsProductClosure,
+        subgroupConditionsProductClosure of (x -> y, y -> z)
+      )
+    }
+
+    // Reconstruct the axiom in its closed form
+    thenHave((x ∈ H, y ∈ H) |- (z ∈ H) ==> (op(op(x, ★, y), ★, z) === op(x, ★, op(y, ★, z)))) by Restate
+    thenHave((x ∈ H, y ∈ H) |- ∀(z, (z ∈ H) ==> (op(op(x, ★, y), ★, z) === op(x, ★, op(y, ★, z))))) by RightForall
+    thenHave((x ∈ H) |- (y ∈ H) ==> ∀(z, (z ∈ H) ==> (op(op(x, ★, y), ★, z) === op(x, ★, op(y, ★, z))))) by Restate
+    thenHave((x ∈ H) |- ∀(y, (y ∈ H) ==> ∀(z, (z ∈ H) ==> (op(op(x, ★, y), ★, z) === op(x, ★, op(y, ★, z)))))) by RightForall
+    thenHave((x ∈ H) ==> ∀(y, (y ∈ H) ==> ∀(z, (z ∈ H) ==> (op(op(x, ★, y), ★, z) === op(x, ★, op(y, ★, z)))))) by Restate
+    thenHave(∀(x, (x ∈ H) ==> ∀(y, (y ∈ H) ==> ∀(z, (z ∈ H) ==> (op(op(x, ★, y), ★, z) === op(x, ★, op(y, ★, z))))))) by RightForall
+
+    have(thesis) by Tautology.from(lastStep, associativityAxiom.definition of (G -> H, * -> ★))
+  }
+
+  /**
+   * Lemma --- The subgroup conditions imply the existence of an identity element on `H`.
+   * 
+   * We show in particular that identity(G, *) is neutral on `H`.
+   */
+  private val subgroupConditionsIdentityExistence = Lemma(
+    (group(G, *), subset(H, G), subgroupConditions) |- identityExistence(H, ★)
+  ) {
+    assume(group(G, *))
+    assume(subset(H, G))
+    assume(subgroupConditions)
+
+    // We show that for an element x ∈ H:
+    // 1. inverse(x) ∈ H                   [[closedByInverses]]
+    // 2. x * inverse(x) ∈ H               [[closedByProducts]]
+    // 3. x * inverse(x) = identity(G, *)  [[inverseCancellation]]
+    // 4. identity(G, *) ∈ H               Substitution of 3. in 2.
+    // 5. isNeutral(identity(G, *), H, ★)  [[identityNeutrality]]
+    // 6. identityExistence(H, ★)          [[identityExistence]]
+    // We finally conclude by [[nonEmpty]].
+
+    // 1. inverse(x) ∈ H
+    have(closedByInverses) by Tautology
+    thenHave(((x ∈ H) ==> (inverse(x, G, *) ∈ H))) by InstantiateForall(x)
+    val step1 = thenHave((x ∈ H) |- (inverse(x, G, *) ∈ H)) by Restate
+
+    // 2. x * inverse(x) ∈ H
+    have(closedByProducts) by Tautology
+    thenHave(∀(y, (x ∈ H /\ y ∈ H) ==> (op(x, *, y) ∈ H))) by InstantiateForall(x)
+    thenHave((x ∈ H /\ inverse(x, G, *) ∈ H) ==> (op(x, *, inverse(x, G, *)) ∈ H)) by InstantiateForall(inverse(x, G, *))
+    thenHave((x ∈ H, inverse(x, G, *) ∈ H) |- (op(x, *, inverse(x, G, *)) ∈ H)) by Restate
+
+    val step2 = have((x ∈ H) |- (op(x, *, inverse(x, G, *)) ∈ H)) by Cut(step1, lastStep)
+
+    // 3. x * inverse(x) = identity(G, *)
+    val step3 = have((x ∈ H) |- op(x, *, inverse(x, G, *)) === identity(G, *)) by Tautology.from(
+      subgroupConditionsSubset,
+      inverseCancellation
+    )
+
+    // 4. identity(G, *) ∈ H
+    have((x ∈ H, op(x, *, inverse(x, G, *)) === identity(G, *)) |- (identity(G, *) ∈ H)) by RightSubstEq(
+      List((op(x, *, inverse(x, G, *)), identity(G, *))),
+      lambda(z, z ∈ H)
+    )(step2)
+    val step4 = have((x ∈ H) |- (identity(G, *) ∈ H)) by Cut(step3, lastStep)
+
+    // 5. isNeutral(identity(G, *), H, ★)
+    have((x ∈ H) |- (op(identity(G, *), *, x) === x) /\ (op(x, *, identity(G, *)) === x)) by Tautology.from(
+      subgroupConditionsSubset,
+      identityNeutrality
+    )
+    thenHave((x ∈ H, op(identity(G, *), ★, x) === op(identity(G, *), *, x), op(x, ★, identity(G, *)) === op(x, *, identity(G, *))) |- (op(identity(G, *), ★, x) === x) /\ (op(x, ★, identity(G, *)) === x)) by RightSubstEq(
+      List((op(identity(G, *), ★, x), op(identity(G, *), *, x)), (op(x, ★, identity(G, *)), op(x, *, identity(G, *)))),
+      lambda(Seq(a, b), (a === x) /\ (b === x))
+    )
+
+    have(x ∈ H |- (op(identity(G, *), ★, x) === x) /\ (op(x, ★, identity(G, *)) === x)) by Tautology.from(
+      lastStep,
+      step4,
+      subgroupConditionsOperation of (x -> identity(G, *), y -> x),
+      subgroupConditionsOperation of (x -> x, y -> identity(G, *))
+    )
+
+    thenHave((x ∈ H) ==> (op(identity(G, *), ★, x) === x) /\ (op(x, ★, identity(G, *)) === x)) by Restate
+    thenHave(∀(x, (x ∈ H) ==> (op(identity(G, *), ★, x) === x) /\ (op(x, ★, identity(G, *)) === x))) by RightForall
+    val step5 = have((x ∈ H) |- isNeutral(identity(G, *), H, ★)) by Tautology.from(
+      lastStep,
+      step4,
+      isNeutral.definition of (e -> identity(G, *), G -> H, * -> ★)
+    )
+
+    // 6. identityExistence(H, ★) 
+    thenHave((x ∈ H) |- ∃(e, isNeutral(e, H, ★))) by RightExists
+    val step6 = have((x ∈ H) |- identityExistence(H, ★)) by Tautology.from(lastStep, identityExistence.definition of (G -> H, * -> ★))
+
+    // Conclude by [[nonEmpty]]
+    thenHave(∃(x, x ∈ H) |- identityExistence(H, ★)) by LeftExists
+
+    have(thesis) by Tautology.from(lastStep, nonEmptySetHasElement of (x -> H))
+  }
+
+  /**
+   * Lemma --- The subgroup conditions imply that for all elements `x` in `H`, there exists an inverse in `H`.
+   */
+  private val subgroupConditionsInverseExistence = Lemma(
+    (group(G, *), subset(H, G), subgroupConditions) |- inverseExistence(H, ★)
+  ) {
+    assume(group(G, *))
+    assume(subset(H, G))
+    assume(subgroupConditions)
+
+    val i = inverse(x, G, *)
+    
+    have(closedByInverses) by Tautology
+    thenHave(x ∈ H ==> i ∈ H) by InstantiateForall(x)
+    val inverseInH = thenHave(x ∈ H |- i ∈ H) by Restate
+
+    // Show that a neutral element of G is also neutral in H
+    val neutralityInheritance = have((e ∈ H, isNeutral(e, G, *)) |- isNeutral(e, H, ★)) subproof {
+      assume(isNeutral(e, G, *))
+      have(∀(x, (x ∈ G) ==> ((op(e, *, x) === x) /\ (op(x, *, e) === x)))) by Tautology.from(isNeutral.definition)
+      thenHave((x ∈ G) ==> ((op(e, *, x) === x) /\ (op(x, *, e) === x))) by InstantiateForall(x)
+      thenHave(x ∈ G |- (op(e, *, x) === x) /\ (op(x, *, e) === x)) by Restate
+      
+      have(x ∈ H |- (op(e, *, x) === x) /\ (op(x, *, e) === x)) by Cut(subgroupConditionsSubset, lastStep)
+      thenHave((x ∈ H, op(e, ★, x) === op(e, *, x), op(x, ★, e) === op(x, *, e)) |- (op(e, ★, x) === x) /\ (op(x, ★, e) === x)) by RightSubstEq(
+        List((op(e, ★, x), op(e, *, x)), (op(x, ★, e), op(x, *, e))),
+        lambda(Seq(a, b), (a === x) /\ (b === x))
+      )
+
+      have((x ∈ H, e ∈ H) |- (op(e, ★, x) === x) /\ (op(x, ★, e) === x)) by Tautology.from(
+        lastStep,
+        subgroupConditionsOperation of (x -> e, y -> x),
+        subgroupConditionsOperation of (x -> x, y -> e)
+      )
+      thenHave(e ∈ H |- (x ∈ H) ==> (op(e, ★, x) === x) /\ (op(x, ★, e) === x)) by Restate
+      thenHave(e ∈ H |- ∀(x, (x ∈ H) ==> (op(e, ★, x) === x) /\ (op(x, ★, e) === x))) by RightForall
+
+      have(e ∈ H |- isNeutral(e, H, ★)) by Tautology.from(lastStep, isNeutral.definition of (G -> H, * -> ★))
+    }
+
+    // Show that i is neutral in H
+    have(x ∈ H |- isNeutral(op(x, *, i), G, *) /\ isNeutral(op(i, *, x), G, *)) by Tautology.from(
+      subgroupConditionsSubset,
+      inverseIsInverse,
+      isInverse.definition of (y -> inverse(x, G, *))
+    )
+    thenHave((x ∈ H, op(x, ★, i) === op(x, *, i), op(i, ★, x) === op(i, *, x)) |- isNeutral(op(x, ★, i), G, *) /\ isNeutral(op(i, ★, x), G, *)) by RightSubstEq(
+      List((op(x, ★, i), op(x, *, i)), (op(i, ★, x), op(i, *, x))),
+      lambda(Seq(a, b), isNeutral(a, G, *) /\ isNeutral(b, G, *))
+    )
+
+    have((x ∈ H, i ∈ H) |- isNeutral(op(x, ★, i), G, *) /\ isNeutral(op(i, ★, x), G, *)) by Tautology.from(
+      lastStep,
+      subgroupConditionsOperation of (x -> x, y -> i),
+      subgroupConditionsOperation of (x -> i, y -> x)
+    )
+
+    have((x ∈ H, i ∈ H) |- isNeutral(op(x, ★, i), H, ★) /\ isNeutral(op(i, ★, x), H, ★)) by Tautology.from(
+      lastStep,
+      neutralityInheritance of (e -> op(x, ★, i)),
+      neutralityInheritance of (e -> op(i, ★, x)),
+      subgroupConditionsProductClosure of (x -> x, y -> i),
+      subgroupConditionsProductClosure of (x -> i, y -> x)
+    )
+
+    have(x ∈ H |- (i ∈ H) /\ isNeutral(op(x, ★, i), H, ★) /\ isNeutral(op(i, ★, x), H, ★)) by Tautology.from(inverseInH, lastStep)
+    have(x ∈ H |- isInverse(i, x, H, ★)) by Tautology.from(lastStep, isInverse.definition of (y -> i, G -> H, * -> ★))
+    thenHave(x ∈ H |- ∃(y, isInverse(y, x, H, ★))) by RightExists
+    thenHave(x ∈ H ==> ∃(y, isInverse(y, x, H, ★))) by Restate
+    thenHave(∀(x, x ∈ H ==> ∃(y, isInverse(y, x, H, ★)))) by RightForall
+
+    have(thesis) by Tautology.from(lastStep, inverseExistence.definition of (G -> H, * -> ★))
+  }
+
+  /**
+   * Theorem (Main subgroup test) --- A subset `H ⊆ G` of a group `(G, *)` is a subgroup if and only if:
    *   1. `H` is non-empty,
    *   2. `H` is closed by products, and
    *   3. `H` is closed by inversion.
@@ -1033,14 +1485,69 @@ object GroupTheory extends lisa.Main {
    * 
    * Note that in the case where H is finite, conditions 1 and 2 are sufficient.
    */
-  val subgroupCondition = Theorem(
-    (group(G, *), subset(H, G)) |- (subgroup(H, G, *) <=> (
-      (H !== ∅) /\
-      ∀(x, x ∈ H ==> ∀(y, y ∈ H ==> (op(x, *, y) ∈ H))) /\
-      ∀(x, x ∈ H ==> inverse(x, G, *) ∈ H)
-    ))
+  val subgroupTest = Theorem(
+    (group(G, *), subset(H, G)) |- (subgroup(H, G, *) <=> subgroupConditions)
   ) {
-    sorry
+    assume(group(G, *))
+    assume(subset(H, G))
+
+    // The forward direction follow directly:
+    // 1. nonEmpty --> [[groupNonEmpty]]
+    // 2. closedByProducts --> [[subgroupOperation]] and [[groupIsClosedByProduct]]
+    // 3. closedByInverses --> [[subgroupInverse]] and [[inverseInGroup]]
+    have(subgroup(H, G, *) |- subgroupConditions) subproof {
+      assume(subgroup(H, G, *))
+      val groupH = have(group(H, ★)) by Tautology.from(subgroup.definition)
+
+      val condition1 = have(nonEmpty) by Cut(groupH, groupNonEmpty of (G -> H, * -> ★))
+      
+      have((x ∈ H, y ∈ H) |- op(x, ★, y) ∈ H) by Cut(groupH, groupIsClosedByProduct of (G -> H, * -> ★))
+      thenHave((x ∈ H, y ∈ H, op(x, ★, y) === op(x, *, y)) |- op(x, *, y) ∈ H) by RightSubstEq(
+        List((op(x, ★, y), op(x, *, y))),
+        lambda(z, z ∈ H)
+      )
+
+      have((x ∈ H, y ∈ H) |- op(x, *, y) ∈ H) by Cut(subgroupOperation, lastStep)
+      thenHave((x ∈ H /\ y ∈ H) ==> (op(x, *, y) ∈ H)) by Restate
+      thenHave(∀(y, (x ∈ H /\ y ∈ H) ==> (op(x, *, y) ∈ H))) by RightForall
+      val condition2 = thenHave(closedByProducts) by RightForall
+
+      have((x ∈ H) |- (inverse(x, H, ★) ∈ H)) by Cut(groupH, inverseInGroup of (G -> H, * -> ★))
+      thenHave((x ∈ H, inverse(x, H, ★) === inverse(x, G, *)) |- inverse(x, G, *) ∈ H) by RightSubstEq(
+        List((inverse(x, H, ★), inverse(x, G, *))),
+        lambda(z, z ∈ H)
+      )
+
+      have((x ∈ H) |- (inverse(x, G, *) ∈ H)) by Cut(subgroupInverse, lastStep)
+      thenHave((x ∈ H) ==> (inverse(x, G, *) ∈ H)) by Restate
+      val condition3 = thenHave(closedByInverses) by RightForall
+
+      have(subgroupConditions) by RightAnd(condition1, condition2, condition3)
+    }
+    val forward = thenHave(subgroup(H, G, *) ==> subgroupConditions) by Restate
+
+    // For the backward direction, we must prove that the conditions make (H, ★) satisfy the axioms of a group:
+    // 1. Closure by products (i.e. ★'s codomain is H): [[closedByProducts]]
+    // 2. Associativity: follows from G's associativity
+    // 3. Identity existence: follows from [[nonEmpty]], [[closedByProducts]] and [[closedByInverses]]
+    // 4. Inverse existence: [[closedByInverse]]
+    //
+    // This direction is quite painful to prove. Each step is presented in its own lemma for easier legibility.
+    have(subgroupConditions |- subgroup(H, G, *)) subproof {
+      assume(subgroupConditions)
+      have(binaryOperation(H, ★) /\ associativityAxiom(H, ★) /\ identityExistence(H, ★) /\ inverseExistence(H, ★)) by RightAnd(
+        subgroupConditionsBinaryRelation,
+        subgroupConditionsAssociativity,
+        subgroupConditionsIdentityExistence,
+        subgroupConditionsInverseExistence
+      )
+      have(group(H, ★)) by Tautology.from(lastStep, group.definition of (G -> H, * -> ★))
+      thenHave(group(G, *) /\ subset(H, G) /\ group(H, ★)) by Tautology
+      have(thesis) by Tautology.from(lastStep, subgroup.definition)
+    }
+    val backward = thenHave(subgroupConditions ==> subgroup(H, G, *)) by Restate
+
+    have(thesis) by RightIff(forward, backward)
   }
 
   // TODO Trivial subgroup
@@ -1345,7 +1852,7 @@ object GroupTheory extends lisa.Main {
     assume(homomorphism(f, G, *, H, **))
     val groupG = have(group(G, *)) by Tautology.from(homomorphism.definition)
 
-    // We show that the kernel satisfies all requirements of [[subgroupCondition]]
+    // We show that the kernel satisfies all requirements of [[subgroupTest]]
     have((x ∈ K) ==> (x ∈ G)) by Tautology.from(kernelDef)
     thenHave(∀(x, x ∈ K ==> x ∈ G)) by RightForall
     val kernelIsSubset = have(subset(K, G)) by Tautology.from(lastStep, subsetAxiom of (x -> K, y -> G))
@@ -1363,23 +1870,22 @@ object GroupTheory extends lisa.Main {
     val condition1 = have(K !== ∅) by Cut(lastStep, setWithElementNonEmpty of (y -> identity(G, *), x -> K))
     
     // 2. The kernel is closed by products
-    have(x ∈ K |- ((y ∈ K) ==> op(x, *, y) ∈ K)) by Restate.from(kernelIsClosedByProducts)
-    thenHave(x ∈ K |- ∀(y, (y ∈ K) ==> op(x, *, y) ∈ K)) by RightForall
-    thenHave(x ∈ K ==> ∀(y, (y ∈ K) ==> op(x, *, y) ∈ K)) by Restate
-    val condition2 = thenHave(∀(x, x ∈ K ==> ∀(y, (y ∈ K) ==> op(x, *, y) ∈ K))) by RightForall
+    have((x ∈ K /\ y ∈ K) ==> op(x, *, y) ∈ K) by Restate.from(kernelIsClosedByProducts)
+    thenHave(∀(y, (x ∈ K /\ y ∈ K) ==> op(x, *, y) ∈ K)) by RightForall
+    val condition2 = thenHave(∀(x, ∀(y, (x ∈ K /\ y ∈ K) ==> op(x, *, y) ∈ K))) by RightForall
 
     // 3. The kernel is closed by inversion
     have((x ∈ K) ==> (inverse(x, G, *) ∈ K)) by Restate.from(kernelIsClosedByInversion)
     val condition3 = thenHave(∀(x, (x ∈ K) ==> (inverse(x, G, *) ∈ K))) by RightForall
 
     // Conclude
-    have((K !== ∅) /\ ∀(x, x ∈ K ==> ∀(y, (y ∈ K) ==> op(x, *, y) ∈ K)) /\ ∀(x, (x ∈ K) ==> (inverse(x, G, *) ∈ K))) by RightAnd(
+    have((K !== ∅) /\ ∀(x, ∀(y, (x ∈ K /\ y ∈ K) ==> op(x, *, y) ∈ K)) /\ ∀(x, (x ∈ K) ==> (inverse(x, G, *) ∈ K))) by RightAnd(
       condition1, condition2, condition3
     )
 
     have(subgroup(K, G, *)) by Tautology.from(
       lastStep,
-      subgroupCondition of (H -> K),
+      subgroupTest of (H -> K),
       groupG,
       kernelIsSubset
     )
diff --git a/src/main/scala/lisa/mathematics/SetTheory.scala b/src/main/scala/lisa/mathematics/SetTheory.scala
index 23670618..295154bb 100644
--- a/src/main/scala/lisa/mathematics/SetTheory.scala
+++ b/src/main/scala/lisa/mathematics/SetTheory.scala
@@ -366,6 +366,20 @@ object SetTheory extends lisa.Main {
     thenHave(∀(y, !in(y, x)) |- (x === ∅)) by Tautology
   }
 
+  /**
+   * Theorem --- A non-empty set has an element.
+   * 
+   * Contra-positive of [[setWithNoElementsIsEmpty]].
+   */
+  val nonEmptySetHasElement = Theorem(
+    !(x === ∅) |- ∃(y, in(y, x))
+  ) {
+    have((!(x === ∅), ∀(y, !in(y, x))) |- (x === ∅)) by Weakening(setWithNoElementsIsEmpty)
+    thenHave((!(x === ∅), ∀(y, !in(y, x))) |- ()) by LeftNot
+    thenHave((!(x === ∅)) |- ! ∀(y, !in(y, x))) by RightNot
+    thenHave(thesis) by Restate
+  }
+
   /**
    * Theorem --- The empty set is a subset of every set.
    *
@@ -861,6 +875,25 @@ object SetTheory extends lisa.Main {
     infix def ∩(y: Term) = setIntersection(x, y)
   }
 
+  /**
+   * Theorem --- If `x ⊆ y` then `x ∩ y = x`.
+   */
+  val setIntersectionSubset = Theorem(
+    subset(x, y) |- (x ∩ y) === x
+  ) {
+    assume(subset(x, y))
+    have(∀(t, t ∈ x ==> t ∈ y)) by Tautology.from(subset.definition)
+    val subsetDef = thenHave(t ∈ x ==> t ∈ y) by InstantiateForall(t)
+
+    have(∀(t, (t ∈ (x ∩ y)) <=> (t ∈ x /\ t ∈ y))) by Definition(setIntersection, setIntersectionUniqueness)(x, y)
+    val intersectionDef = thenHave((t ∈ (x ∩ y)) <=> (t ∈ x /\ t ∈ y)) by InstantiateForall(t)
+
+    have((t ∈ (x ∩ y)) <=> (t ∈ x)) by Tautology.from(subsetDef, intersectionDef)
+    thenHave(∀(t, (t ∈ (x ∩ y)) <=> (t ∈ x))) by RightForall
+
+    have(thesis) by Tautology.from(lastStep, extensionalityAxiom of (x -> (x ∩ y), y -> x))
+  }
+
   val unaryIntersectionUniqueness = Theorem(
     ∃!(z, ∀(t, in(t, z) <=> (exists(b, in(b, x)) /\ ∀(b, in(b, x) ==> in(t, b)))))
   ) {
@@ -1672,6 +1705,37 @@ object SetTheory extends lisa.Main {
     have(thesis) by Tautology.from(lastStep, defUnfold)
   }
 
+  /**
+   * Theorem --- If `a ⊆ b` and `c ⊆ d`, then `(a * c) ⊆ (b * d)`.
+   */
+  val subsetsCartesianProduct = Theorem(
+    (subset(a, b), subset(c, d)) |- subset(cartesianProduct(a, c), cartesianProduct(b, d))
+  ) {
+    have(subset(a, b) |- ∀(x, x ∈ a ==> x ∈ b)) by Tautology.from(subset.definition of (x -> a, y -> b))
+    val subsetDef = thenHave(subset(a, b) |- x ∈ a ==> x ∈ b) by InstantiateForall(x)
+
+    assume(subset(a, b))
+    assume(subset(c, d))
+    have(((t === pair(x, y)) /\ x ∈ a /\ y ∈ c) |- ((t === pair(x, y)) /\ x ∈ b /\ y ∈ d)) by Tautology.from(
+      subsetDef,
+      subsetDef of (a -> c, b -> d, x -> y)
+    )
+    thenHave(((t === pair(x, y)) /\ x ∈ a /\ y ∈ c) |- ∃(y, ((t === pair(x, y)) /\ x ∈ b /\ y ∈ d))) by RightExists
+    thenHave(((t === pair(x, y)) /\ x ∈ a /\ y ∈ c) |- ∃(x, ∃(y, ((t === pair(x, y)) /\ x ∈ b /\ y ∈ d)))) by RightExists
+    thenHave(∃(y, ((t === pair(x, y)) /\ x ∈ a /\ y ∈ c)) |- ∃(x, ∃(y, ((t === pair(x, y)) /\ x ∈ b /\ y ∈ d)))) by LeftExists
+    thenHave(∃(x, ∃(y, ((t === pair(x, y)) /\ x ∈ a /\ y ∈ c))) |- ∃(x, ∃(y, ((t === pair(x, y)) /\ x ∈ b /\ y ∈ d)))) by LeftExists
+    thenHave(∃(x, ∃(y, ((t === pair(x, y)) /\ x ∈ a /\ y ∈ c))) ==> ∃(x, ∃(y, ((t === pair(x, y)) /\ x ∈ b /\ y ∈ d)))) by Restate
+
+    have(t ∈ cartesianProduct(a, c) ==> t ∈ cartesianProduct(b, d)) by Tautology.from(
+      lastStep,
+      elemOfCartesianProduct of (x -> a, y -> c),
+      elemOfCartesianProduct of (x -> b, y -> d)
+    )
+    thenHave(∀(t, t ∈ cartesianProduct(a, c) ==> t ∈ cartesianProduct(b, d))) by RightForall
+
+    have(thesis) by Tautology.from(lastStep, subset.definition of (x -> cartesianProduct(a, c), y -> cartesianProduct(b, d)))
+  }
+
   /**
    * Theorem --- the union of two Cartesian products is a subset of the product of unions.
    *

From ad952d6442feccc32ce090e5d9b04ce121be7700 Mon Sep 17 00:00:00 2001
From: dhalilov 
Date: Fri, 16 Jun 2023 14:40:06 +0200
Subject: [PATCH 68/68] scala{fix, fmt}

---
 .../automation/kernel/CommonTactics.scala     | 65 +++++++------
 .../lisa/mathematics/FirstOrderLogic.scala    |  4 +-
 .../scala/lisa/mathematics/GroupTheory.scala  | 93 ++++++++++---------
 .../scala/lisa/mathematics/SetTheory.scala    | 24 +++--
 4 files changed, 100 insertions(+), 86 deletions(-)

diff --git a/src/main/scala/lisa/automation/kernel/CommonTactics.scala b/src/main/scala/lisa/automation/kernel/CommonTactics.scala
index a08b5c58..ee401b17 100644
--- a/src/main/scala/lisa/automation/kernel/CommonTactics.scala
+++ b/src/main/scala/lisa/automation/kernel/CommonTactics.scala
@@ -8,10 +8,12 @@ import lisa.prooflib.BasicStepTactic.*
 import lisa.prooflib.ProofTacticLib.{_, given}
 import lisa.prooflib.SimpleDeducedSteps.*
 import lisa.prooflib.*
-import lisa.utils.KernelHelpers.{_, given}
 import lisa.utils.FOLPrinter
+import lisa.utils.KernelHelpers.{_, given}
+
 import scala.collection.immutable.Queue
-import scala.collection.mutable.{Map => MutableMap, Set => MutableSet}
+import scala.collection.mutable.{Map => MutableMap}
+import scala.collection.mutable.{Set => MutableSet}
 
 object CommonTactics {
 
@@ -270,18 +272,18 @@ object CommonTactics {
    * Γ, Γ', ... |- a(i) === a(j)
    * 
* Proves any equality induced transitively by the equalities of the premises. - * + * * Internally, we construct an undirected graph, where sides of an equality are represented by vertices, and * an edge between two terms `a` and `b` means that some premise proves `a === b` (or equivalently `b === a`). * We also keep the premise from which the equality stems as a label of the edge to construct the final antecedent of * the bottom sequent. - * + * * We can see that an equality `a === b` is provable from the premises if and only if `a` is reachable from `b`. * We thus run Breadth-First Search (BFS) on the graph starting from `a` to find the smallest solution (in terms of * sequent calculus steps), if any. */ object Equalities extends ProofTactic { - def apply(using lib: Library, proof: lib.Proof)(equalities: proof.Fact*)(bot: Sequent): proof.ProofTacticJudgement = { + def apply(using lib: Library, proof: lib.Proof)(equalities: proof.Fact*)(bot: Sequent): proof.ProofTacticJudgement = { // Construct the graph as an adjacency list for O(1) equality checks val graph = MutableMap[FOL.Term, List[FOL.Term]]() val premises = MutableMap[(FOL.Term, FOL.Term), proof.Fact]() @@ -305,7 +307,7 @@ object CommonTactics { premises += ((y, x) -> premise) } } - case _ => + case _ => if (error.isEmpty) { error = Some(proof.InvalidProofTactic("Right-hand side of premises should only contain equalities.")) } @@ -315,7 +317,7 @@ object CommonTactics { if (error.nonEmpty) { return error.get } - + if (bot.right.size != 1) { return proof.InvalidProofTactic(s"Right-hand side of bottom sequent expected exactly 1 formula, got ${bot.right.size}") } @@ -367,23 +369,25 @@ object CommonTactics { val initialStep = order(premises(y -> path.head), path.head, y) val u = variable - path.zip(path.tail).foldLeft(initialStep)((leftEq, vars) => { - val leftSeq = innerProof.getSequent(leftEq) - val (a, b) = vars - val premiseEq = premises(vars) - val premiseSeq = proof.getSequent(premiseEq) - - // Apply equality transitivity on (a === y) /\ (a === b) to get (b === y) - // TODO Watch for issue #161, as assumptions will break this step - lib.have((leftSeq.left + premiseSeq.right.head) |- b === y) by RightSubstEq( - List(getEqTerms(premiseEq)), - lambda(u, u === y) - )(leftEq) - val seq = (leftSeq.left ++ premiseSeq.left) |- b === y - val eq = lib.have(seq) by Cut(premiseEq, lib.lastStep) - - eq - }) + path + .zip(path.tail) + .foldLeft(initialStep)((leftEq, vars) => { + val leftSeq = innerProof.getSequent(leftEq) + val (a, b) = vars + val premiseEq = premises(vars) + val premiseSeq = proof.getSequent(premiseEq) + + // Apply equality transitivity on (a === y) /\ (a === b) to get (b === y) + // TODO Watch for issue #161, as assumptions will break this step + lib.have((leftSeq.left + premiseSeq.right.head) |- b === y) by RightSubstEq( + List(getEqTerms(premiseEq)), + lambda(u, u === y) + )(leftEq) + val seq = (leftSeq.left ++ premiseSeq.left) |- b === y + val eq = lib.have(seq) by Cut(premiseEq, lib.lastStep) + + eq + }) } } @@ -394,20 +398,19 @@ object CommonTactics { } } proof.InvalidProofTactic("Equality is not provable from the premises.") - + case _ => proof.InvalidProofTactic("Right-hand side of bottom sequent should be of the form x === y.") } } } - /** + /** *
-   * Γ, φ |- Δ, Σ   Γ, ¬φ |- Δ, Σ'  
+   * Γ, φ |- Δ, Σ   Γ, ¬φ |- Δ, Σ'
    * -----------------------------
    * Γ |- Δ
    * 
* - * * TODO: Extending the tactic to more general pivots */ object Cases extends ProofTactic { @@ -415,8 +418,10 @@ object CommonTactics { val seqPos = proof.getSequent(pos) val seqNeg = proof.getSequent(neg) - if (!(contains(seqPos.left, phi) && contains(seqNeg.left, !phi) && !contains(seqNeg.left, phi)) && - !(contains(seqPos.left, !phi) && contains(seqNeg.left, phi) && !contains(seqNeg.left, !phi))) { + if ( + !(contains(seqPos.left, phi) && contains(seqNeg.left, !phi) && !contains(seqNeg.left, phi)) && + !(contains(seqPos.left, !phi) && contains(seqNeg.left, phi) && !contains(seqNeg.left, !phi)) + ) { proof.InvalidProofTactic("The given sequent do not contain φ or ¬φ.") } else { val gamma = bot.left diff --git a/src/main/scala/lisa/mathematics/FirstOrderLogic.scala b/src/main/scala/lisa/mathematics/FirstOrderLogic.scala index da2e8b0f..1c695fc1 100644 --- a/src/main/scala/lisa/mathematics/FirstOrderLogic.scala +++ b/src/main/scala/lisa/mathematics/FirstOrderLogic.scala @@ -396,7 +396,7 @@ object FirstOrderLogic extends lisa.Main { } /** - * Theorem --- Unique existential quantifier distributes fully with a closed formula. + * Theorem --- Unique existential quantifier distributes fully with a closed formula. */ val uniqueExistentialConjunctionWithClosedFormula = Theorem( existsOne(x, P(x) /\ p()) <=> (existsOne(x, P(x)) /\ p()) @@ -427,7 +427,7 @@ object FirstOrderLogic extends lisa.Main { have((!p(), existsOne(x, P(x)) /\ p()) |- ()) by Tautology thenHave((!p(), existsOne(x, P(x)) /\ p()) |- existsOne(x, P(x) /\ p())) by Weakening val backward = thenHave(!p() |- existsOne(x, P(x)) /\ p() ==> existsOne(x, P(x) /\ p())) by Restate - + have(thesis) by RightIff(forward, backward) } diff --git a/src/main/scala/lisa/mathematics/GroupTheory.scala b/src/main/scala/lisa/mathematics/GroupTheory.scala index cdce4d64..09d9c071 100644 --- a/src/main/scala/lisa/mathematics/GroupTheory.scala +++ b/src/main/scala/lisa/mathematics/GroupTheory.scala @@ -6,11 +6,11 @@ import lisa.automation.kernel.CommonTactics.Equalities import lisa.automation.kernel.CommonTactics.ExistenceAndUniqueness import lisa.automation.kernel.OLPropositionalSolver.Tautology import lisa.automation.settheory.SetTheoryTactics.UniqueComprehension +import lisa.kernel.proof.SequentCalculus.Hypothesis import lisa.mathematics.FirstOrderLogic.equalityTransitivity import lisa.mathematics.FirstOrderLogic.existsOneImpliesExists import lisa.mathematics.FirstOrderLogic.substitutionInUniquenessQuantifier import lisa.mathematics.SetTheory.* -import lisa.kernel.proof.SequentCalculus.Hypothesis /** * Group theory, developed following Chapter 2 of S. Lang "Undergraduate Algebra". @@ -196,7 +196,7 @@ object GroupTheory extends lisa.Main { /** * Lemma --- For elements `x, y, z` in a group `(G, *)`, we have `(xy)z = x(yz)`. - * + * * Practical reformulation of the [[associativityAxiom]]. */ val associativity = Lemma( @@ -218,14 +218,14 @@ object GroupTheory extends lisa.Main { /** * Lemma --- For elements `x, y` in an abelian group `(G, *)`, we have `xy = yx`. - * + * * Practical reformulation of [[commutativityAxiom]]. */ val commutativity = Lemma( (abelianGroup(G, *), x ∈ G, y ∈ G) |- op(x, *, y) === op(y, *, x) ) { assume(abelianGroup(G, *)) - + have(∀(x, x ∈ G ==> ∀(y, y ∈ G ==> (op(x, *, y) === op(y, *, x))))) by Tautology.from( abelianGroup.definition, commutativityAxiom.definition @@ -304,7 +304,7 @@ object GroupTheory extends lisa.Main { assume(x ∈ G) assume(y ∈ G) - // Show that x * y is in relation range + // Show that x * y is in relation range have(pair(pair(x, y), op(x, *, y)) ∈ *) by Tautology.from( appDef, groupOperationIsFunctional, @@ -323,14 +323,14 @@ object GroupTheory extends lisa.Main { ) thenHave(op(x, *, y) ∈ relationRange(*) ==> op(x, *, y) ∈ G) by InstantiateForall(op(x, *, y)) thenHave(op(x, *, y) ∈ relationRange(*) |- op(x, *, y) ∈ G) by Restate - + have(thesis) by Cut(productInRelationRange, lastStep) } /** * Identity uniqueness --- In a group (G, *), an identity element is unique, i.e. if both `e * x = x * e = x` and * `f * x = x * f = x` for all `x`, then `e = f`. - * + * * This justifies calling `e` the identity element. */ val identityUniqueness = Theorem( @@ -382,7 +382,7 @@ object GroupTheory extends lisa.Main { /** * Lemma --- For any element `x` in a group `(G, *)`, we have `x * e = e * x = x`. - * + * * Practical reformulation of [[identityIsNeutral]]. */ val identityNeutrality = Lemma( @@ -552,7 +552,7 @@ object GroupTheory extends lisa.Main { thenHave(thesis) by InstantiateForall(y) } - val forward = have(x ∈ G |- isInverse(y, x, G, *) ==> isInverse(x, y, G, *)) subproof { + val forward = have(x ∈ G |- isInverse(y, x, G, *) ==> isInverse(x, y, G, *)) subproof { assume(x ∈ G) have(isInverse(y, x, G, *) |- y ∈ G /\ isNeutral(op(x, *, y), G, *) /\ isNeutral(op(y, *, x), G, *)) by Tautology.from(isInverse.definition) thenHave(isInverse(y, x, G, *) |- isNeutral(op(y, *, x), G, *) /\ isNeutral(op(x, *, y), G, *)) by Tautology @@ -575,7 +575,7 @@ object GroupTheory extends lisa.Main { /** * Involution of the inverse -- For all `x`, `inverse(inverse(x)) = x`. - * + * * Direct corollary of [[inverseSymmetry]]. */ val inverseIsInvolutive = Theorem( @@ -602,7 +602,7 @@ object GroupTheory extends lisa.Main { inverseInGroup, associativity of (x -> i, y -> x, z -> y) ) - + // (ix)y = y have((group(G, *), x ∈ G) |- ∀(y, (y ∈ G) ==> ((op(op(i, *, x), *, y) === y) /\ (op(y, *, op(i, *, x)) === y)))) by Tautology.from( inverseIsInverse, @@ -630,7 +630,7 @@ object GroupTheory extends lisa.Main { /** * Theorem --- In a group `(G, *)`, we have `yx = zx ==> y = z`. - * + * * Analoguous to [[leftCancellation]]. */ val rightCancellation = Theorem( @@ -645,7 +645,7 @@ object GroupTheory extends lisa.Main { inverseInGroup, associativity of (x -> y, y -> x, z -> i) ) - + // y(xi) = y have((group(G, *), x ∈ G) |- ∀(y, (y ∈ G) ==> ((op(op(x, *, i), *, y) === y) /\ (op(y, *, op(x, *, i)) === y)))) by Tautology.from( inverseIsInverse, @@ -706,7 +706,7 @@ object GroupTheory extends lisa.Main { /** * Theorem --- If `x * y = e` then `y = inverse(x)`. - * + * * This also implies that `x = inverse(y)` by [[inverseSymmetry]]. */ val inverseTest = Theorem( @@ -725,7 +725,7 @@ object GroupTheory extends lisa.Main { ) // 2. x * y = x * inverse(x) - have((op(x, *, y) === e) |- (op(x, *, y) === e) ) by Hypothesis + have((op(x, *, y) === e) |- (op(x, *, y) === e)) by Hypothesis val eq2 = have((op(x, *, y) === e) |- op(x, *, y) === op(x, *, inverse(x, G, *))) by Equalities(eq1, lastStep) // Conclude by left cancellation @@ -871,7 +871,8 @@ object GroupTheory extends lisa.Main { val groupH = have(group(H, ★)) by Tautology.from(subgroup.definition) have((x ∈ H, y ∈ H) |- pair(x, y) ∈ relationDomain(★)) by Tautology.from( - groupH, groupPairInOperationDomain of (G -> H, * -> ★) + groupH, + groupPairInOperationDomain of (G -> H, * -> ★) ) have((functional(*), x ∈ H, y ∈ H) |- op(x, ★, y) === op(x, *, y)) by Cut( lastStep, @@ -888,7 +889,7 @@ object GroupTheory extends lisa.Main { * Lemma --- If `H` is a subgroup of `G`, then `e_H ∈ G`. */ val subgroupIdentityInParent = Lemma( - subgroup(H, G, *) |- identity(H, ★) ∈ G + subgroup(H, G, *) |- identity(H, ★) ∈ G ) { val identityInH = have(subgroup(H, G, *) |- identity(H, ★) ∈ H) by Tautology.from( subgroup.definition, @@ -962,7 +963,7 @@ object GroupTheory extends lisa.Main { } /** - * Theorem --- If `H` is a subgroup of `G`, then the inverse is the same as in the parent group. + * Theorem --- If `H` is a subgroup of `G`, then the inverse is the same as in the parent group. */ val subgroupInverse = Theorem( (subgroup(H, G, *), x ∈ H) |- inverse(x, H, ★) === inverse(x, G, *) @@ -995,7 +996,7 @@ object GroupTheory extends lisa.Main { lastStep, subgroupOperation of (y -> inverse(x, H, ★)) ) - + val eq1 = have(op(x, *, inverse(x, H, ★)) === eH) by Cut(inverseHInH, lastStep) // 2. e_H = e_G @@ -1025,9 +1026,9 @@ object GroupTheory extends lisa.Main { // // 2.1 Main subgroup test - // + // // We define several useful lemmas to attack this easy, but long theorem to formalize - // + // private val nonEmpty = H !== ∅ private val closedByProducts = ∀(x, ∀(y, (x ∈ H /\ y ∈ H) ==> (op(x, *, y) ∈ H))) @@ -1078,11 +1079,11 @@ object GroupTheory extends lisa.Main { /** * Lemma --- The subgroup conditions imply that `(x, y)` is in the relation domain of `★`. - * + * * Analogous to [[groupPairInOperationDomain]]. */ private val subgroupConditionsPairInDomain = Lemma( - (group(G, *), subset(H, G), subgroupConditions, x ∈ H, y ∈ H) |- pair(x, y) ∈ relationDomain(★) + (group(G, *), subset(H, G), subgroupConditions, x ∈ H, y ∈ H) |- pair(x, y) ∈ relationDomain(★) ) { assume(group(G, *)) assume(subset(H, G)) @@ -1103,7 +1104,7 @@ object GroupTheory extends lisa.Main { /** * Lemma --- The subgroup conditions imply that `x ★ y = x * y`. - * + * * Analogous to [[subgroupOperation]]. */ private val subgroupConditionsOperation = Lemma( @@ -1263,7 +1264,7 @@ object GroupTheory extends lisa.Main { /** * Lemma --- The subgroup conditions imply associativity on `H`. - * + * * This directly follows from associativity on `G` and [[subgroupConditionsOperation]]. */ private val subgroupConditionsAssociativity = Lemma( @@ -1321,7 +1322,7 @@ object GroupTheory extends lisa.Main { /** * Lemma --- The subgroup conditions imply the existence of an identity element on `H`. - * + * * We show in particular that identity(G, *) is neutral on `H`. */ private val subgroupConditionsIdentityExistence = Lemma( @@ -1371,7 +1372,9 @@ object GroupTheory extends lisa.Main { subgroupConditionsSubset, identityNeutrality ) - thenHave((x ∈ H, op(identity(G, *), ★, x) === op(identity(G, *), *, x), op(x, ★, identity(G, *)) === op(x, *, identity(G, *))) |- (op(identity(G, *), ★, x) === x) /\ (op(x, ★, identity(G, *)) === x)) by RightSubstEq( + thenHave( + (x ∈ H, op(identity(G, *), ★, x) === op(identity(G, *), *, x), op(x, ★, identity(G, *)) === op(x, *, identity(G, *))) |- (op(identity(G, *), ★, x) === x) /\ (op(x, ★, identity(G, *)) === x) + ) by RightSubstEq( List((op(identity(G, *), ★, x), op(identity(G, *), *, x)), (op(x, ★, identity(G, *)), op(x, *, identity(G, *)))), lambda(Seq(a, b), (a === x) /\ (b === x)) ) @@ -1391,7 +1394,7 @@ object GroupTheory extends lisa.Main { isNeutral.definition of (e -> identity(G, *), G -> H, * -> ★) ) - // 6. identityExistence(H, ★) + // 6. identityExistence(H, ★) thenHave((x ∈ H) |- ∃(e, isNeutral(e, H, ★))) by RightExists val step6 = have((x ∈ H) |- identityExistence(H, ★)) by Tautology.from(lastStep, identityExistence.definition of (G -> H, * -> ★)) @@ -1412,7 +1415,7 @@ object GroupTheory extends lisa.Main { assume(subgroupConditions) val i = inverse(x, G, *) - + have(closedByInverses) by Tautology thenHave(x ∈ H ==> i ∈ H) by InstantiateForall(x) val inverseInH = thenHave(x ∈ H |- i ∈ H) by Restate @@ -1423,7 +1426,7 @@ object GroupTheory extends lisa.Main { have(∀(x, (x ∈ G) ==> ((op(e, *, x) === x) /\ (op(x, *, e) === x)))) by Tautology.from(isNeutral.definition) thenHave((x ∈ G) ==> ((op(e, *, x) === x) /\ (op(x, *, e) === x))) by InstantiateForall(x) thenHave(x ∈ G |- (op(e, *, x) === x) /\ (op(x, *, e) === x)) by Restate - + have(x ∈ H |- (op(e, *, x) === x) /\ (op(x, *, e) === x)) by Cut(subgroupConditionsSubset, lastStep) thenHave((x ∈ H, op(e, ★, x) === op(e, *, x), op(x, ★, e) === op(x, *, e)) |- (op(e, ★, x) === x) /\ (op(x, ★, e) === x)) by RightSubstEq( List((op(e, ★, x), op(e, *, x)), (op(x, ★, e), op(x, *, e))), @@ -1480,9 +1483,9 @@ object GroupTheory extends lisa.Main { * 1. `H` is non-empty, * 2. `H` is closed by products, and * 3. `H` is closed by inversion. - * + * * It is often easier to prove the 3 conditions independently than using the definition directly. - * + * * Note that in the case where H is finite, conditions 1 and 2 are sufficient. */ val subgroupTest = Theorem( @@ -1500,7 +1503,7 @@ object GroupTheory extends lisa.Main { val groupH = have(group(H, ★)) by Tautology.from(subgroup.definition) val condition1 = have(nonEmpty) by Cut(groupH, groupNonEmpty of (G -> H, * -> ★)) - + have((x ∈ H, y ∈ H) |- op(x, ★, y) ∈ H) by Cut(groupH, groupIsClosedByProduct of (G -> H, * -> ★)) thenHave((x ∈ H, y ∈ H, op(x, ★, y) === op(x, *, y)) |- op(x, *, y) ∈ H) by RightSubstEq( List((op(x, ★, y), op(x, *, y))), @@ -1562,7 +1565,7 @@ object GroupTheory extends lisa.Main { /** * Definition --- A group homomorphism is a mapping `f: G -> H` from structures `G` and `H` equipped with binary operations `*` and `★` respectively, * such that for all `x, y ∈ G`, we have* `f(x * y) = f(x) ** f(y)`. - * + * * In the following, "homomorphism" always stands for "group homomorphism", i.e. `(G, *)` and `(H, **)` are groups. */ val homomorphism = DEF(f, G, *, H, **) --> group(G, *) /\ group(H, **) /\ functionFrom(f, G, H) /\ ∀(x, x ∈ G ==> ∀(y, y ∈ G ==> (app(f, op(x, *, y)) === op(app(f, x), **, app(f, y))))) @@ -1614,7 +1617,7 @@ object GroupTheory extends lisa.Main { identityIdempotence of (x -> e) ) val eq0 = have(homomorphism(f, G, *, H, **) |- op(e, *, e) === e) by Cut(groupG, lastStep) - + // 1. f(e * e) = f(e) have(app(f, e) === app(f, e)) by RightRefl thenHave(op(e, *, e) === e |- app(f, op(e, *, e)) === app(f, e)) by RightSubstEq( @@ -1637,14 +1640,14 @@ object GroupTheory extends lisa.Main { groupH, identityIdempotence of (x -> app(f, e), G -> H, * -> **) ) - have(homomorphism(f, G, *, H, **)|- (op(app(f, e), **, app(f, e)) === app(f, e)) <=> (app(f, e) === identity(H, **))) by Cut( + have(homomorphism(f, G, *, H, **) |- (op(app(f, e), **, app(f, e)) === app(f, e)) <=> (app(f, e) === identity(H, **))) by Cut( appInH, lastStep ) - + have(thesis) by Tautology.from(lastStep, eq3) } - + /** * Theorem --- If `f: G -> H` is a group homomorphism, then `f(inverse(x, G, *)) = inverse(f(x), H, **)`. */ @@ -1802,9 +1805,9 @@ object GroupTheory extends lisa.Main { /** * Lemma --- The kernel is closed by inversion, i.e. if `x ∈ K` then `inverse(x, G, *) ∈ K`. */ - val kernelIsClosedByInversion = Lemma ( + val kernelIsClosedByInversion = Lemma( (homomorphism(f, G, *, H, **), x ∈ K) |- inverse(x, G, *) ∈ K - ) { + ) { assume(homomorphism(f, G, *, H, **)) assume(x ∈ K) @@ -1824,7 +1827,7 @@ object GroupTheory extends lisa.Main { List((app(f, x), e)), lambda(z, app(f, inverse(x, G, *)) === inverse(z, H, **)) ) - + val eq1 = have(app(f, inverse(x, G, *)) === inverse(e, H, **)) by Cut(appValue, lastStep) // 2. inverse(e) = e @@ -1868,7 +1871,7 @@ object GroupTheory extends lisa.Main { kernelDef of (x -> identity(G, *)) ) val condition1 = have(K !== ∅) by Cut(lastStep, setWithElementNonEmpty of (y -> identity(G, *), x -> K)) - + // 2. The kernel is closed by products have((x ∈ K /\ y ∈ K) ==> op(x, *, y) ∈ K) by Restate.from(kernelIsClosedByProducts) thenHave(∀(y, (x ∈ K /\ y ∈ K) ==> op(x, *, y) ∈ K)) by RightForall @@ -1880,7 +1883,9 @@ object GroupTheory extends lisa.Main { // Conclude have((K !== ∅) /\ ∀(x, ∀(y, (x ∈ K /\ y ∈ K) ==> op(x, *, y) ∈ K)) /\ ∀(x, (x ∈ K) ==> (inverse(x, G, *) ∈ K))) by RightAnd( - condition1, condition2, condition3 + condition1, + condition2, + condition3 ) have(subgroup(K, G, *)) by Tautology.from( @@ -1896,7 +1901,7 @@ object GroupTheory extends lisa.Main { /** * Isomorphism --- An isomorphism `f: G -> H` is a bijective homomorphism. - * + * * In some sense, isomorphic groups are equivalent, up to relabelling their elements. */ val isomorphism = DEF(f, G, *, H, **) --> homomorphism(f, G, *, H, **) /\ bijective(f, G, H) diff --git a/src/main/scala/lisa/mathematics/SetTheory.scala b/src/main/scala/lisa/mathematics/SetTheory.scala index 295154bb..770c2f6f 100644 --- a/src/main/scala/lisa/mathematics/SetTheory.scala +++ b/src/main/scala/lisa/mathematics/SetTheory.scala @@ -368,7 +368,7 @@ object SetTheory extends lisa.Main { /** * Theorem --- A non-empty set has an element. - * + * * Contra-positive of [[setWithNoElementsIsEmpty]]. */ val nonEmptySetHasElement = Theorem( @@ -2474,7 +2474,7 @@ object SetTheory extends lisa.Main { thenHave((t ∈ restrictedFunction(f, x)) <=> (t ∈ f /\ ∃(y, ∃(z, y ∈ x /\ (t === pair(y, z)))))) by InstantiateForall(t) thenHave(t ∈ restrictedFunction(f, x) ==> t ∈ f) by Tautology thenHave(∀(t, t ∈ restrictedFunction(f, x) ==> t ∈ f)) by RightForall - + have(thesis) by Tautology.from( lastStep, subset.definition of (x -> restrictedFunction(f, x), y -> f) @@ -2489,12 +2489,12 @@ object SetTheory extends lisa.Main { ) { val g = restrictedFunction(f, x) assume(functional(f)) - + val functionalDef = have(relation(f) /\ ∀(x, ∃(y, in(pair(x, y), f)) ==> ∃!(y, in(pair(x, y), f)))) by Tautology.from(functional.definition) thenHave(∀(x, ∃(y, in(pair(x, y), f)) ==> ∃!(y, in(pair(x, y), f)))) by Tautology val fAppUniqueness = thenHave(∃(y, in(pair(t, y), f)) ==> ∃!(y, in(pair(t, y), f))) by InstantiateForall(t) val fIsRelation = have(relation(f)) by Tautology.from(functionalDef) - + val gIsRelation = have(relation(g)) subproof { // Since f is a relation and g is a subset of f, g is also a relation by subset transitivity have(relationBetween(f, a, b) |- subset(f, cartesianProduct(a, b))) by Tautology.from(relationBetween.definition of (r -> f)) @@ -2509,12 +2509,14 @@ object SetTheory extends lisa.Main { thenHave(∃(a, ∃(b, relationBetween(f, a, b))) |- ∃(a, ∃(b, relationBetween(g, a, b)))) by LeftExists have(relation(f) |- relation(g)) by Tautology.from( - lastStep, relation.definition of (r -> f), relation.definition of (r -> g) + lastStep, + relation.definition of (r -> f), + relation.definition of (r -> g) ) have(thesis) by Cut(fIsRelation, lastStep) } - val gAppUniqueness = have(∀(t, ∃(y, pair(t, y) ∈ g) ==> ∃!(y, pair(t, y) ∈ g))) subproof { + val gAppUniqueness = have(∀(t, ∃(y, pair(t, y) ∈ g) ==> ∃!(y, pair(t, y) ∈ g))) subproof { have((pair(t, a) ∈ restrictedFunction(f, x)) <=> (pair(t, a) ∈ f /\ in(t, x))) by Tautology.from(restrictedFunctionPairMembership) val equiv = thenHave(∀(a, (pair(t, a) ∈ restrictedFunction(f, x)) <=> (pair(t, a) ∈ f /\ in(t, x)))) by RightForall @@ -2671,13 +2673,13 @@ object SetTheory extends lisa.Main { ) val domainInclusion = have(x ∈ relationDomain(g) |- x ∈ relationDomain(f)) by Cut( restrictedFunctionDomain of (x -> d), - lastStep, + lastStep ) // Characterize app(f, x) val characterization = have( (app(g, x) === app(f, x)) <=> (((functional(f) /\ (x ∈ relationDomain(f))) ==> (p ∈ f)) /\ - ((!functional(f) \/ (x ∉ relationDomain(f))) ==> (app(g, x) === ∅))) + ((!functional(f) \/ (x ∉ relationDomain(f))) ==> (app(g, x) === ∅))) ) by InstantiateForall(app(g, x))(app.definition) // Use the definition of restricted functions @@ -2705,8 +2707,10 @@ object SetTheory extends lisa.Main { thenHave((functional(f), x ∈ relationDomain(g), !functional(f) \/ (x ∉ relationDomain(f))) |- (app(g, x) === ∅)) by Weakening val neg = thenHave((functional(f), x ∈ relationDomain(g)) |- (!functional(f) \/ (x ∉ relationDomain(f)) ==> (app(g, x) === ∅))) by Restate - have((functional(f), x ∈ relationDomain(g)) |- (((functional(f) /\ (x ∈ relationDomain(f))) ==> (p ∈ f)) /\ - ((!functional(f) \/ (x ∉ relationDomain(f))) ==> (app(g, x) === ∅)))) by RightAnd(pos, neg) + have( + (functional(f), x ∈ relationDomain(g)) |- (((functional(f) /\ (x ∈ relationDomain(f))) ==> (p ∈ f)) /\ + ((!functional(f) \/ (x ∉ relationDomain(f))) ==> (app(g, x) === ∅))) + ) by RightAnd(pos, neg) have(thesis) by Tautology.from(lastStep, characterization) }