Zulip Chat Archive

Stream: Equational

Topic: 1485


Pace Nielsen (Oct 16 2024 at 16:42):

Since this equation has three variables, it is not approachable by translation-invariant substitution.

So, I tried a straightforward linear map xy=ax+by+cx\diamond y=ax+by+c. The resulting reduction system is ba1ab,a2=b2=0,bc=accba\mapsto 1-ab, a^2=b^2=0,bc=-ac-c. Using such a map with coefficients subject to these relations, all of the open implications are unfalsifiable.

Hopefully one of the other techniques will work.

Terence Tao (Oct 16 2024 at 19:41):

Thank you for reporting on the negative result. This is not as "fun" as reporting a positive result, but is a very useful part of the problem-solving process in mathematics (and one that we often don't show as much as we should in the public-facing side of mathematical research!).

EDIT: do you mean ba=1abba = 1-ab instead of ba1abba \to 1-ab?

Pace Nielsen (Oct 16 2024 at 19:45):

Terence Tao said:

EDIT: do you mean ba=1abba = 1-ab instead of ba1abba \to 1-ab?

Yes, that was a typo. Good catch!

Terence Tao (Oct 17 2024 at 01:40):

This is a weaker version of the central groupoid axiom 168 (and if right multiplication is surjective, it is equivalent to that axiom). So probably some of the ideas used to understand central groupoids (eg graph theoretic approaches) may be relevant here.

Pace Nielsen (Oct 17 2024 at 21:26):

Here are a few additional no-go results for other open implications. Both the translation-invariant approach (even using noncommutative groups) and the linear approach (using coefficients from a noncommutative ring) will not work for Equations 1728, 3352, and 3457. (The proposed implication holds under such maps.) If I find other examples where these two approaches don't work, I'll post them here.

[Edited later:] Add 1323 to the list.

Terence Tao (Oct 17 2024 at 23:44):

An initial attempt to port some of the most basic central groupoid theory (x = (y ◇ x) ◇ (x ◇ z)) into 1485 (x = (y ◇ x) ◇ (x ◇ (z ◇ y))). Suppose that G obeys 1485 (which is equivalent to its dual 2182, x = ((y ◇ z) ◇ x) ◇ (x ◇ y)). Define x ◇ G := {x ◇ y: y ∈ G} to be the right-multiples of x, and similarly G ◇ x := {y ◇ x: y ∈ G} to be the left-multiples.

Claim: a ∈ G ◇ b if and only if b ∈ a ◇ G.

Proof: If a ∈ G ◇ b then a = x ◇ b for some x, hence b = (x ◇ b) ◇ (b ◇ (x ◇ x)) by 1485, hence b ∈ a ◇ G. The converse implication is similar but uses 2182 instead of 1485.

Because of this, one can place a directed graph on G by writing an edge a → b if and only if a ∈ G ◇ b or equivalently b ∈ a ◇ G.

In the case of central groupoids, the axiom 168 is equivalent to the assertion that any two points a, b are connected in G by exactly one path of length two, namely a → a ◇ b → b. (Note that paths of length 2 necessarily take the form y ◇ x → x → x ◇ z.) In this more general situation in which we only assume the weaker axiom 1485 (or 2182), we still have this path a → a ◇ b → b, but it seems that now there are also more paths. I don't yet know whether one can express 1485 purely in terms of this graph.

Terence Tao (Oct 17 2024 at 23:54):

Perhaps one way to construct a 1485 magma G is to first construct a promising candidate directed graph, and then try to build a 1485-obeying magma operation that has that graph?

Terence Tao (Oct 18 2024 at 00:05):

OK, here is an attempt to encode 1485 into graph-theoretic language. We have a directed graph as before. A path of length two is said to be good if it is of the form a → a ◇ b → b. The following claims are obvious from the definitions for a 1485 magma:

Claim 1 Any two points are connected by exactly one good path of length two.

Claim 2 Any directed edge in the graph is the first component of at least one good path.

Claim 3 Any directed edge in the graph is the second component of at least one good path.

We also have

Claim 4: given a 5-cycle a → b → c → d → e → a, if the paths b → c → d and d → e → a are good, then the path c → d → e is also good.

Proof: by hypothesis, c = b ◇ d and e = d ◇ a. Also, since a → b, we have a = z ◇ b for some z. Then by 1485 d = (b ◇ d) ◇ (d ◇ (z ◇ b)) = c ◇ e, given the claim.

Conversely, a directed graph with a notion of good path that obeys Claims 1-4, obeys 1485, by defining x ◇ y to be the midpoint of the unique good path from x to y; this is well-defined by Claim 1. Claims 2, 3 then ensure that the directed graph relation a → b is equivalent to either b ∈ a ◇ G or a ∈ G ◇ b, and 1485 then follows by reversing the proof of Claim 4.

Will Sawin (Oct 18 2024 at 12:11):

So this suggests we should try to find a directed graph such that any two points are connected by at least one path of length two, and there are "fewer than expected" five-cycles, so that claim 4 is not as restrictive.

Terence Tao (Oct 18 2024 at 15:59):

Here are two examples of 1485 models that might be instructive. The first are the natural central groupoids. Here, G=S×SG = S \times S, and the directed graph relation is given by (a,b)(b,c)(a,b) \to (b,c) for all b,cSb,c \in S. It is easy to see that any two points are connected by precisely one path of length two, and we designate all such paths to be good, then the other axioms are easy to verify. The magma operation is then (a,b)(c,d)=(b,c)(a,b) \diamond (c,d) = (b,c).

The other example (listed as the minimal size nontrivial magma obeying 1485 in equation explorer) is the magma on {0,1}\{0,1\} in which 00=10 \diamond 0 = 1 and 01=10=11=00 \diamond 1 = 1 \diamond 0 = 1 \diamond 1 = 0. This corresponds to the directed graph with edges 000 \to 0, 010 \to 1, and 101 \to 0. The good paths are 0100 \to 1 \to 0, 0010 \to 0 \to 1, 1011 \to 0 \to 1, and 1001 \to 0 \to 0; 0000 \to 0 \to 0 is the only path of length 2 that is not good. Up to cyclic permutation, the 5-cycles are 000000 0 \to 0 \to 0 \to 0 \to 0 \to 0, 010000 0 \to 1 \to 0 \to 0 \to 0 \to 0, and 010100 0 \to 1 \to 0 \to 1 \to 0 \to 0; they all obey the 5-cycle axiom.

Bhavik Mehta (Oct 18 2024 at 16:07):

Just to clarify for my own sake, aren't the implications from 1485 solved by the argument here: https://leanprover.zulipchat.com/#narrow/channel/458659-Equational/topic/Equation.202126/near/476396542? They are not, I mixed up equation 2162 and equation 2126

Bhavik Mehta (Oct 18 2024 at 18:34):

The maps txt:(xG)yx(Gy)t \mapsto x \diamond t : (x \diamond G) \diamond y \rightarrow x \diamond (G \diamond y) and ttyt \mapsto t \diamond y are inverse to each other. (Proof: observe (x((xt)y))y = (xt)y, and dually for the opposite direction). This is the closest analogy I could find to the fact in central groupoids that the obvious maps between xGx \diamond G and GyG \diamond y are inverse, but it would be nice if something stronger were true

Terence Tao (Oct 19 2024 at 02:31):

Another way of saying this is that if LxL_x is the left multiplication map and RyR_y is the right multiplication map then LxRyL_x R_y and RyLxR_y L_x are idempotents: LxRyLxRy=LxRyL_x R_y L_x R_y = L_x R_y and RyLxRyLx=RyLxR_y L_x R_y L_x = R_y L_x. In contrast, for central groupoids one has the stronger laws RyLxRy=RyR_y L_x R_y = R_y and LxRyLx=LxL_x R_y L_x = L_x.

Terence Tao (Oct 19 2024 at 14:36):

In the spirit of reporting negative results: I experimented with whether one could start with one complete model G of 1485 (e.g., a central groupoid), and enlarge it by adding additional elements to create larger model G'. I started with trying to add a single element 0 that was not already in G, but quickly found that the law 0 = ((y ◇ 0) ◇ (0 ◇ (z ◇ y))would almost certainly be inconsistent with the existing law in G, since most likely y ◇ 0 and 0 ◇ (z ◇ y) would already lie in G. In retrospect this was going to be difficult given the fact that for finite central groupoids at least, the cardinality is known to be a perfect square. I was then motivated by the identity (n+1)^2 = n^2 + n + n + 1 to try instead to adjoin a new element 0 to G together with two additional sets 0 ◇ G and G ◇ 0, with 0 idempotent (0 ◇ 0 = 0). Then every element x of the original magma G acquires a factorization as x = (0 ◇ x) ◇ (x ◇ 0), and we might als be able to impose absorptive laws such as 0 ◇ (0 ◇ G) = 0. I'm still struggling somewhat to see if one can still preserve the law 1485 in this way, there are a lot of cases to check. I know it works for natural central groupoids, but that's a significantly stronger structure than 1485.

Terence Tao (Oct 19 2024 at 21:29):

It may be useful to consider translation invariant models of 1485. Here, G is a group, and the directed graph is a Cayley graph generated by some subset S in G, thus the edges take the form x → xs for s in S and x in G. The good paths then take the form x → xs → xst for x in G and (s,t) in some set E of pairs in S × S; the magma law is given by x ⋄ xst = xs for x in G and (s,t) in E. The axioms of a 1485 magma in this model are then:

  1. The two coordinate projections from E to S are surjective.
  2. The product map (s,t) ↦ st is a bijection from E to G.
  3. If s₁ s₂ s₃ s₄ s₅ = 1 for some s₁, s₂, s₃, s₄, s₅ in S, and (s₁, s₂), (s₃, s₄) lie in E, then (s₂, s₃) also lies in E.

Central groupoids in this model correspond to the case where E is all of S × S.

At present I can't think of any non-trivial data G, S, E that obey axioms 1-3, but presumably if we can construct such a model, it will be useful for ruling out implications from 1485.

Terence Tao (Oct 20 2024 at 03:47):

Pace Nielsen said:

Here are a few additional no-go results for other open implications. Both the translation-invariant approach (even using noncommutative groups) and the linear approach (using coefficients from a noncommutative ring) will not work for Equations 1728, 3352, and 3457.

Did you mean 3475 instead of 3457? 3457 has no open outgoing implications.

Pace Nielsen (Oct 20 2024 at 03:50):

I believe I did mean 3475, but I'll double-check that on Monday at work (I kept my notes). [Yes it was a typo. That's all moot now that the implication is resolved.]

Pace Nielsen (Oct 21 2024 at 15:44):

Can someone with access to Vampire check if the non-implications stay non-implications (at least conjecturally) when we assume commutativity? That would allow us to work with non-directed graphs, which might simplify some considerations.

Daniel Weber (Oct 21 2024 at 15:54):

Unfortunately no, together with commutativity it implies 152 and 1483

Terence Tao (Oct 21 2024 at 17:05):

Given that the natural central groupoids (a,b)(c,d)=(b,c)(a,b) \diamond (c,d) = (b,c) are very far from commutative, it was a long shot that commutativity would be the right axiom here, but still good to know; I've marked this as one further "immunity" for this equation over in the outstanding equations thread.

Terence Tao (Oct 21 2024 at 17:26):

Natural central groupoids also qualify as a linear construction xy=Ax+Byx \diamond y= Ax + By (with A(a,b)=(b,0)A(a,b) = (b,0) and B(c,d)=(0,c)B(c,d) = (0,c), but linear constructions have already been ruled out previously. I have a vague hope that some local modification of a natural central groupoid (perhaps using some variant of a greedy method, but with the base being a natural central groupoid rather than a blank multiplication table - or perhaps a mostly blank table with a natural central groupoid in one corner already filled in ) could be helpful, but this is as far as I got with the idea so far.

Kevin M (Oct 21 2024 at 21:04):

Has this been checked up to finite magma of size 8?
The smallest non-trivial magma is of size 2, while the smallest non-trivial central groupoid is of size 4 and has some symmetry. So I was hoping there was some way to get an interesting size 8 magma that is vaguely a "semidirect product" using the symmetry of the size 4 central groupoid. I haven't had luck and was wondering if size 8 has already been fully explored for non-direct-product results, in which case I can give up chasing this phantom.

Negative result: with some slow python code I searched for a size 6 magma satisfying eqn 1485 and the search completed without finding anything. (As a code check, I found 24 of size 4. It sounds like magma of size 4 have been fully searched. Is there a way I can check against the current search results to check if 24 is the correct count?)

Terence Tao (Oct 21 2024 at 21:26):

I think it is a very good idea to try to locate some novel examples of finite 1485 magmas that aren't just products of central groupoids and the order 2 example.

Kevin M (Oct 22 2024 at 11:23):

Something else that might lead to a novel example of eqn 1485:
For finite central groupoids, there exists an x such that x = (x ◇ x). This is not necessary for eqn 1485, and while the numerical evidence is obviously sparse, the only known examples of eqn 1485 which are not central groupoids do not have such an element.

So does eqn 1485 and the existence of an idempotent x imply the central groupoid law? If true, this could help searches by knowing what to avoid. If false, a counter-example would give a novel example of eqn 1485.

I am not familiar with Vampire. Is that something which can be checked with Vampire?

thormikkel (Oct 22 2024 at 11:46):

Has any work been done on predicting what order the finite magmas following this rule might have? The finite central groupoids have very restricted orders. It's interesting that there seemed to be no magmas of order 6 following this rule, but there are magmas of order 2 and 4.

I also noticed that there are no two distinct elements x,y where (x ◇ x) = (y ◇ y), from the examples that have been given. This might be obvious from some traversal of the implication graph, but if this law holds then it narrows down the search a little bit.

There is also a strange coincidence with 4! = 24 being the number of legal permutations of the diagonal of the multiplication table under the law (x ◇ x) != (y ◇ y) for distinct x,y, while it is also the number of magmas of order 4 obeying 1485. Although from eyeballing the magmas, there are clearly some "forbidden" permutations of the diagonal.

Ian Duleba (Oct 22 2024 at 13:10):

I have conducted testing for magmas up to size 7x7 using MiniZinc+chuffed. I will share the results to facilitate further exploration by others in this area:

  • 2x2: <1s to find an example - 2 magmas found
  • 3x3: <1s - none
  • 4x4: <1s to find an example - 24 magmas found
  • 5x5: 4s - none
  • 6x6: 2m - none
  • 7x7: 30m - none
  • 8x8: ~30s to find an example - +1800 so far (work in progress). No new counterexamples found.

You can access the magmas here.

Ian Duleba (Oct 22 2024 at 13:11):

If there is interest, I can prepare a PR to include this data for broader use.

Michael Bucko (Oct 22 2024 at 14:05):

Ian Duleba [schrieb]

You can access the magmas [here](https://github.com/iduleba/Equational-Theories/tree/main/eq1485).

Could you please share the code? I'm still learning how to integrate MiniZinc into my workflow.

Kevin M (Oct 22 2024 at 14:06):

I updated my code to C, and made it more greedy. Can confirm no 7x7, and 8x8 is still running.
However it made it far enough in 8x8 to say:

  • there is no 8x8 with an idempotent element
  • picking some results at random, here is one that does not appear to be a product of size 2 or size 4
    [[1, 0, 0, 1, 2, 2, 3, 3],
    [0, 0, 0, 0, 3, 3, 3, 3],
    [4, 4, 5, 4, 5, 5, 4, 5],
    [7, 4, 5, 7, 6, 6, 4, 5],
    [1, 0, 0, 1, 6, 6, 3, 3],
    [7, 4, 5, 7, 2, 2, 4, 5],
    [4, 4, 5, 4, 5, 5, 4, 5],
    [0, 0, 0, 0, 3, 3, 3, 3]]

Michael Bucko (Oct 22 2024 at 14:10):

Kevin M schrieb:

I updated my code to C, and made it more greedy. Can confirm no 7x7, and 8x8 is still running.

Kevin, could you also please make the code available? I'd love to have a look.

Terence Tao (Oct 22 2024 at 15:40):

thormikkel said:

I also noticed that there are no two distinct elements x,y where (x ◇ x) = (y ◇ y), from the examples that have been given. This might be obvious from some traversal of the implication graph, but if this law holds then it narrows down the search a little bit.

This at least we can explain: 1485 implies 151, x = (x ◇ x) ◇ (x ◇ x), so the map x → x ◇ x is an involution (and in particular is a bijection, even in infinite models).

Pace Nielsen (Oct 22 2024 at 15:55):

Terence Tao said:

thormikkel said:

I also noticed that there are no two distinct elements x,y where (x ◇ x) = (y ◇ y), from the examples that have been given. This might be obvious from some traversal of the implication graph, but if this law holds then it narrows down the search a little bit.

This at least we can explain: 1485 implies 151, x = (x ◇ x) ◇ (x ◇ x), so the map x → x ◇ x is an involution (and in particular is a bijection, even in infinite models).

The implication 1485 -> 151 is open, is it not?

Terence Tao (Oct 22 2024 at 16:00):

Oops, you're right. (Getting quite sloppy the last day or so... should be more careful.) So this phenomenon is not yet explained, on the other hand refuting it would give us a new anti-implication.

I suppose one could try adding the axioms 0 ≠ 1, 0 ◇ 0 = 1 ◇ 1 to 1485 and see if anything interesting comes out of it.

Kevin M (Oct 22 2024 at 16:06):

Okay, here's the code if you are curious. Its nothing fancy, just a "hotspot / greedy" recursive brute force.
https://github.com/bafflingbits/brute1485
The MiniZinc approach should be much more flexible.

I haven't had a chance to categorize the 8x8 results yet. There are a lot more than I was expecting: 35520 with (0 ◇ 0) = 1. Everything after that should just be even more permutations of those results, so I stopped the program.

thormikkel (Oct 22 2024 at 16:08):

Kevin M said:

Okay, here's the code if you are curious. Its nothing fancy, just a "hotspot / greedy" recursive brute force.
https://github.com/bafflingbits/brute1485
The MiniZinc approach should be much more flexible.

I haven't had a chance to categorize the 8x8 results yet. There are a lot more than I was expecting: 35520 with (0 ◇ 0) = 1. Everything after that should just be even more permutations of those results, so I stopped the program.

I'll conjecture that the total number will be 8! = 40320 just for the fun of it :big_smile: adhering to the diagonal permutation law

Terence Tao (Oct 22 2024 at 16:21):

Kevin M said:

[[1, 0, 0, 1, 2, 2, 3, 3],
[0, 0, 0, 0, 3, 3, 3, 3],
[4, 4, 5, 4, 5, 5, 4, 5],
[7, 4, 5, 7, 6, 6, 4, 5],
[1, 0, 0, 1, 6, 6, 3, 3],
[7, 4, 5, 7, 2, 2, 4, 5],
[4, 4, 5, 4, 5, 5, 4, 5],
[0, 0, 0, 0, 3, 3, 3, 3]]

One noticeable thing about this table is that there are several repeated rows: elements x,xx, x' such that xy=xyx \diamond y = x' \diamond y for all yy. (Also several repeated columns, and several "almost repeated" rows and columns.) Certainly natural central groupoids (a,b)(c,d)=(b,c)(a,b) \diamond (c,d) = (b,c) have this property, since (a,b)(c,d)=(a,b)(c,d)(a,b) \diamond (c,d) = (a',b) \diamond (c,d). But I don't know if all central groupoids, let alone weak central groupoids, have some law that explains this phenomenon.

thormikkel (Oct 22 2024 at 16:29):

I observed the same phenomenon for order 4 magmas. There seems to be some combinatorial structure here, and a semi-strong law of permitted configurations of the diagonal elements in the multiplication table. The same order of diagonal elements appears multiple times, but with columns and rows either permuted or transposed. Originally I thought that the diagonal elements had a unique configuration for n! total magmas, but this is not the case. The total number of magmas still seem to be closely following this law though.

Michael Bucko (Oct 22 2024 at 16:38):

Kevin M schrieb:

Okay, here's the code if you are curious. Its nothing fancy, just a "hotspot / greedy" recursive brute force.
https://github.com/bafflingbits/brute1485

Thank you for sharing. Why did you pick the brute force vs MiniZinc then?
Also, if brute force worked here, then perhaps it could work with all the other remaining equations?

thormikkel (Oct 22 2024 at 17:17):

Kevin M said:

Okay, here's the code if you are curious. Its nothing fancy, just a "hotspot / greedy" recursive brute force.
https://github.com/bafflingbits/brute1485
The MiniZinc approach should be much more flexible.

I haven't had a chance to categorize the 8x8 results yet. There are a lot more than I was expecting: 35520 with (0 ◇ 0) = 1. Everything after that should just be even more permutations of those results, so I stopped the program.

None of these magmas had (x ◇ x) = (y ◇ y) for distinct x,y by a check in python.

Kevin M (Oct 22 2024 at 17:27):

Michael Bucko said:

Thank you for sharing. Why did you pick the brute force vs MiniZinc then?

I wasn't aware of MiniZinc til Ian Duleba posted the results above. :)
If we learn more constraints to reduce the search space, likely the more general tools will incorporate that information easier.

Kevin M (Oct 22 2024 at 17:32):

thormikkel said:

I observed the same phenomenon for order 4 magmas. There seems to be some combinatorial structure here, and a semi-strong law of permitted configurations of the diagonal elements in the multiplication table. The same order of diagonal elements appears multiple times, but with columns and rows either permuted or transposed. Originally I thought that the diagonal elements had a unique configuration for n! total magmas, but this is not the case. The total number of magmas still seem to be closely following this law though.

There are only two unique magmas in the n=4 case. One is the example central groupoid, and the other is just a direct product of (N=2) x (N=2). Both of those have an automorphism group of size 2, so they have 4!/2 different permutations of the "product table".

Michael Bucko (Oct 22 2024 at 17:35):

Kevin M schrieb:

Michael Bucko said:

Thank you for sharing. Why did you pick the brute force vs MiniZinc then?

I wasn't aware of MiniZinc til Ian Duleba posted the results above. :)
If we learn more constraints to reduce the search space, likely the more general tools will incorporate that information easier.

Reminds me of grid search. Perhaps a neural net could learn the constraints when trained on enough equations. I think generating lots of high quality math insight datasets could eventually get us to such neural nets.
Btw. I learnt today about PySat (from another post). Check it out too.

Kevin M (Oct 22 2024 at 18:28):

Okay I grouped the equivalent magmas in the 8x8 search.
There are 10 unique magma (up to relabelling of elements), 2 of which are direct products, so 8 new examples showed up in 8x8.
https://github.com/bafflingbits/brute1485/blob/main/n8_unique.txt

Terence Tao (Oct 22 2024 at 18:43):

Definitely seeing a lot of repeated rows and columns here. For instance in magma #7 the eight rows break off into four pairs of identical rows, though the same isn't quite true for the eight columns. I looked at all the equations known or conjectured to be implied by either 1485 or the stronger central groupoid axiom 168 and don't yet have any plausible explanation for this phenomenon yet, but it feels significant to me.

Bhavik Mehta (Oct 22 2024 at 19:21):

Are any of these isomorphic to a dual of another? For central groupoids there are 6 of order 9, two are self dual and the remaining 4 pair up as iso to a dual

Jose Brox (Oct 22 2024 at 22:07):

Kevin M ha dicho:

For finite central groupoids, there exists an x such that x = (x ◇ x). This is not necessary for eqn 1485, and while the numerical evidence is obviously sparse, the only known examples of eqn 1485 which are not central groupoids do not have such an element.

So does eqn 1485 and the existence of an idempotent x imply the central groupoid law? If true, this could help searches by knowing what to avoid. If false, a counter-example would give a novel example of eqn 1485.

I've run both Prover9 (for trying to get the implication) and Mace4 (for trying to get a countermodel) for more than 10 minutes on weak central groupoid + existence of an element a such that aa=a against central groupoid. Both searches have been fruitless. In addition I have also tried Prover9 with the Knuth-Bendix algorithm for 5 minutes.

Pace Nielsen (Oct 22 2024 at 22:11):

This is a long-shot idea. 1485 is equivalent to its dual. If you add that dual rule, does it help the search for examples?

Jose Brox (Oct 22 2024 at 22:25):

Pace Nielsen ha dicho:

This is a long-shot idea. 1485 is equivalent to its dual. If you add that dual rule, does it help the search for examples?

It seems to help going somewhat faster, at least with Mace4. I'm looking now for some model of weak central grupoid with two elements such that aa=bb. It is looking at size 16 now, with no examples found (if I'm not mistaken, the search is exhaustive).

Terence Tao (Oct 22 2024 at 23:57):

Following Knuth's paper, let A be the adjacency matrix of the directed graph associated to a weak central groupoid, thus A has entry 1 at (x,y) (or, in graph notation x → y), if and only if y = x ◇ z for some z, or equivalently x = w ◇ y for some w. Knuth observed that the central groupoid condition is equivalent to having A^2 = J where J is the all-ones matrix, which is a rank one matrix. That is to say, any two points x, y are connected by exactly one path of length two, which then defines the product: x → x ◇ y → y.

Experimentally, it appears that for at least some of the order 8 magmas above, A^2 is now a rank two matrix, which after permuting rows and columns is all 2s in the upper left 4 x 4 block, and all 1s elswehere. Thus, any two x, y are connected by exactly one path of length two which is x → x ◇ y → y, except when x, y are both in {0,1,2,3}, in which case there is also an alternate path x → z → ynot coming from the magma operation.

Not sure what to make of this yet. EDIT: for the order 2 magma, there is a similar phenomenon: the adjacency matrix is [[1,1],[1,0]], and the square is [[2,1],[1,1]].

Terence Tao (Oct 23 2024 at 00:24):

This vaguely suggests to me that while many of the order 8 magmas are not direct products of an order 2 magma and an order 4 magma, they may still be "semi-direct products" in some sense. Having a semi-direct product construction in hand could well generate a lot of further magmas that could potentially be used to generate anti-implications.

Terence Tao (Oct 23 2024 at 00:40):

In that vein, I conjecture that each of the order 8 magmas has a homomorphism to the order 2 magma, with each element of the latter being mapped to by four elements of the former. One can detect what the homomorphism ought to be by computing A^2: the elements x that get to have two paths to themselves (x → x ◇ x → x and some other path) would go to one element of the magma, and the elements that only have one path to themselves (which is necessarily x → x ◇ x → x) would go to the other.

This would partially describe the multiplication table (it divides the 8 x 8 table into four 4 x 4 blocks; the top left block only has entries in the second half of the magma, and the other three blocks have entries in the first half, as per the multiplication table [[1, 0], [0, 0]] of the order 2 magma), but then one still needs to understand the structure in each of the blocks.

Terence Tao (Oct 23 2024 at 00:55):

For instance, starting with magma #4

[[1, 0, 0, 1, 2, 2, 3, 3],
 [0, 0, 0, 0, 3, 3, 3, 3],
 [4, 4, 4, 4, 5, 5, 5, 5],
 [6, 4, 4, 7, 7, 6, 5, 5],
 [1, 0, 0, 1, 2, 2, 3, 3],
 [6, 4, 4, 7, 7, 6, 5, 5],
 [0, 0, 0, 0, 5, 5, 5, 5],
 [4, 4, 4, 4, 3, 3, 3, 3]]

I can rearrange and relabel the rows and columns (guided by the computation of A^2) to get

[[4, 4, 5, 5, 0, 0, 1, 1],
 [6, 7, 7, 6, 2, 2, 3, 3],
 [4, 4, 5, 5, 0, 0, 1, 1],
 [6, 7, 7, 6, 2, 2, 3, 3],
 [0, 0, 1, 1, 0, 0, 1, 1],
 [2, 2, 3, 3, 2, 2, 3, 3],
 [0, 0, 3, 3, 0, 0, 3, 3],
 [2, 2, 1, 1, 2, 2, 1, 1]]

where the block structure becomes apparent. In this case the three blocks taking values in {0,1,2,3} are in fact identical. I don't know if this is true in general (this is beginning to reach the limits of my python coding ability).

Ian Duleba (Oct 23 2024 at 03:10):

Apparently there are several ways one can relabel the elements in #4 so that the top left block only has entries in the second half of the magma and the other three blocks have entries in the first half, however, I couldn't find a permutation of the labels such that the three blocks taking values in {0,1,2,3} are in fact identical for #4 (in your example, the last row is slightly different).

Ian Duleba (Oct 23 2024 at 03:12):

This python script generates all permutations of labels for a magma and outputs magmas with the mentioned properties.

Terence Tao (Oct 23 2024 at 03:13):

Ah, well spotted. So the blocks tend to be almost identical, but not quite - I guess they are something like a modified version of a direct product of an order 2 and order 4 weak central groupoid, where the modification only affects a relatively small portion of the multiplication table.

Terence Tao (Oct 23 2024 at 03:16):

I was going to propose looking for 16 x 16 magmas with this block structure but without the absolutely identical blocks, this is probably too many variables for an ATP to handle. There do seem to be other symmetries lying around though. (I wish I understood why some rows and some columns are identical, in particular.)

Ian Duleba (Oct 23 2024 at 03:40):

#0 has the property you had in mind:

[[1, 0, 3, 2, 5, 4, 7, 6],
 [0, 0, 0, 0, 0, 0, 0, 0],
 [3, 0, 3, 0, 0, 3, 0, 3],
 [2, 0, 0, 2, 5, 7, 7, 5],
 [5, 0, 0, 5, 5, 0, 0, 5],
 [4, 0, 3, 7, 0, 4, 7, 3],
 [7, 0, 0, 7, 0, 7, 7, 0],
 [6, 0, 3, 5, 5, 3, 0, 6]]

becomes:

[[5, 7, 7, 5, 1, 3, 3, 1],
 [7, 4, 7, 4, 3, 0, 3, 0],
 [7, 7, 7, 7, 3, 3, 3, 3],
 [5, 4, 7, 6, 1, 0, 3, 2],
 [1, 3, 3, 1, 1, 3, 3, 1],
 [3, 0, 3, 0, 3, 0, 3, 0],
 [3, 3, 3, 3, 3, 3, 3, 3],
 [1, 0, 3, 2, 1, 0, 3, 2]]

Terence Tao (Oct 23 2024 at 03:55):

Thanks! Though this one may be "too" good - it looks like it's the direct product of [[1,0],[0,0]] and [[1,3,3,1],[3,0,3,0],[3,3,3,3],[1,0,3,2]] (think of 0,1,2,3,4,5,6,7 as being relabeled as (0,0), (1,0), (2,0), (3,0), (0,1), (1,1), (2,1), (3,1). I think the latter order 4 magma is itself a direct product of two copies of [[1,0],[0,0]] though I can't quite do the relabeling in my head. I was hoping for something where three blocks were identical but the fourth wasn't just a shift of the other three (note in this example that the top left block is the same block as the rest, but with all entries shifted by 4).

Clearly there is some sort of pattern here in general, but it's sort of a weird pattern that only repeats some of the time but not others... reminds me a little bit of aperiodic tilings of the plane (e.g., Penrose tilings) in some ways.

Ian Duleba (Oct 23 2024 at 04:51):

oh, I see. Unfortunately, with the "shift check" in place, none of the 10 magmas can be rearranged to create three identical blocks following that structure.

Terence Tao (Oct 23 2024 at 05:16):

OK, good to know, thanks.

Terence Tao (Oct 23 2024 at 05:51):

OK, I think I can now interpret this pattern graph-theoretically. Consider a directed graph G (which may have loops) whose vertices come in two classes, which I will arbitrarily call "low" and "high" (these correspond to the labels 0,1,2,3 and 4,5,6,7 in the above examples after ordering). We assume the following axioms.

  1. Every low vertex has outgoing and incoming edges to both low and high vertices.
  2. Every high vertex has outgoing and incoming edges to low vertices, but none to high vertices. Thus, the pattern high → high is forbidden.
  3. One can get from any vertex x to any vertex y by a path x → x ◇ y → y of length two that passes through an intermediate vertex x ◇ y that is low unless x and y are both low, in which case x ◇ y is high. (It is possible for some of these vertices to coincide.)
  4. The only other paths of length two that can exist are those paths x → z → y in which all three vertices are low. That is to say, the only non-good paths permitted are of the form low → low → low.

In terms of matrices, the adjacency matrix has the block form [[A,B], [C,0]] where BC = AB = CA = CB = J, where J is the all ones block. The square of the matrix is then [[J+A^2, J], [J, J]]. In the examples we have it seems we also have A^2=J, but for the purposes of generating a weak central groupoid this does not seem necessary. One can compare this axiom system to Knuth's axiom for the directed graph associated to a central groupoid, namely that one can get from any vertex to any vertex by exactly one path of length two (in matrix terms, this is A^2=J).

Claim: any such graph generates a weak central groupoid (1485).

Proof: it suffices to verify Claims 1-4 from my previous comment, where the good paths are the ones identified as x → x ◇ y → y. Claims 1-3 follow from axioms 1-3. For claim 4, we suppose for contradiction that we have a 5-cycle a → b → c → d → e → a where the paths b → c → d and d → e → a are good but the path c → d → e is bad. By axiom 5, c, d, e are all low, and this forces a and b to be high. But a high vertex cannot connect to a high vertex by axiom 2.

The next step for me is to try to interpret what some of the open implications from 1485, such as 166 or 1483, look like under this graph model.

Terence Tao (Oct 23 2024 at 06:09):

Alas, all such models obey both 166 and 1483, and hence all other open implications. It seems just allowing non-good low → low → low paths is not enough of a departure from the central groupoid axiom. But perhaps some more complicated modification of this idea can still work.

Proof of 166 x = (y ◇ x) ◇ (x ◇ x). The relevant graph geometry here is y ◇ x → x → x ◇ x. The only way this path is not good is if all three vertices are low; but if x is low then x ◇ x is high, so 166 holds.

Proof of 1483 x = (y ◇ x) ◇ (x ◇ (y ◇ z)). Now the geometry is y → y ◇ x → x → x ◇ (y ◇ z) → y ◇ z and y → y ◇ z. The only way that the path y ◇ x → x → x ◇ (y ◇ z) is not good is if all three vertices are low, which force y and y ◇ z to be high; but then they cannot be adjacent.

thormikkel (Oct 23 2024 at 08:18):

Could the identical rows simply be "pigeonholed" in this way due to the small order of the magma combined with the 1485 law? Since the "row" X o Y appears in the (dual) law, maybe there is a combinatorial argument as to why small magmas must take this form, but that this structure doesn't necessarily hold generally for larger orders.

Another semi-interesting observation: The identical rows only come in multiples of 2. Maybe this is obvious, but why should there not be an order 8 1485-magma with 3 pairs of identical rows for instance? Is this related to there being no 1485 magma of order 6?

Jose Brox (Oct 23 2024 at 09:57):

Reporting negative results: I have conducted the searches suggested by @Kevin M and @thormikkel related to squares in weak central groupoids, with the dual speed-up suggested by @Pace Nielsen . I have run Prover9/Mace4 for 12 hours:
1) Up to size 15, there is no weak central groupoid with an element a such that a^2=a which is not a central groupoid. Neither on an extensive search in size 16.
2) Nevertheless, no proof has been found that an element a with a^2=a necessarily forces the weak to become strong.
3) Up to size 17 there is no weak central groupoid with two elements such that a^2=b^2 (also tested extensively in size 18).

Michael Bucko (Oct 23 2024 at 10:38):

Jose Brox schrieb:

Reporting negative results: I have conducted the searches suggested by Kevin M and thormikkel related to squares in weak central groupoids, with the dual speed-up suggested by Pace Nielsen . I have run Prover9/Mace4 for 12 hours:
1) Up to size 15, there is no weak central groupoid with an element a such that a^2=a which is not a central groupoid. Neither on an extensive search in size 16.
2) Nevertheless, no proof has been found that an element a with a^2=a necessarily forces the weak to become strong.
3) Up to size 17 there is no weak central groupoid with two elements such that a^2=b^2 (also tested extensively in size 18).

Jose, can we perhaps document how we're running those provers? Some screenshots, datasets, arguments -- that'd be very helpful.

Jose Brox (Oct 23 2024 at 10:51):

Michael Bucko ha dicho:

.

Sure! Here you have screenshots, for the output report I'd need to kill the processes (I'm keeping them running a little more, just in case).
Prover9 1485 same squares no model up to size 17.JPG
Prover9 1485 idempotent central groupoid.JPG

Michael Bucko (Oct 23 2024 at 10:57):

Jose Brox schrieb:

Michael Bucko ha dicho:

.

Sure! Here you have screenshots, for the output report I'd need to kill the processes (I'm keeping them running a little more, just in case).
Prover9 1485 same squares no model up to size 17.JPG
Prover9 1485 idempotent central groupoid.JPG

This is very helpful. Thank you!

Daniel Weber (Oct 23 2024 at 11:10):

xxx \diamond x is an involution on {yz,zy=y}\{y \mid \exists z, z \diamond y = y\}

Daniel Weber (Oct 23 2024 at 11:31):

Is there a weak central groupoid of order 9 which isn't strong?

Jose Brox (Oct 23 2024 at 11:34):

Daniel Weber ha dicho:

Is there a weak central groupoid of order 9 which isn't strong?

No, there isn't, according to Mace4.

thormikkel (Oct 23 2024 at 11:35):

Daniel Weber said:

Is there a weak central groupoid of order 9 which isn't strong?

Good question, more generally one could conjecture that every weak central groupoid of order n^2 is also strong

Daniel Weber (Oct 23 2024 at 11:36):

thormikkel said:

Daniel Weber said:

Is there a weak central groupoid of order 9 which isn't strong?

Good question, more generally one could conjecture that every weak central groupoid of order n^2 is also strong

There is one of order 4, though

thormikkel (Oct 23 2024 at 11:38):

Daniel Weber said:

thormikkel said:

Daniel Weber said:

Is there a weak central groupoid of order 9 which isn't strong?

Good question, more generally one could conjecture that every weak central groupoid of order n^2 is also strong

There is one of order 4, though

Sorry, I meant n^2 where n is not even. I'm thinking there might be weak central groupoids of order 4*3 = 12 because it contains an even square

Jose Brox (Oct 23 2024 at 11:39):

There are no weak central groupoids of order 12, according to Mace4.

Daniel Weber (Oct 23 2024 at 11:42):

What sizes have weak central groupoids? I know of 2,4,8,16.

thormikkel (Oct 23 2024 at 11:52):

Are any of the order 16 groupoids known to be "sporadic" and not a product of lower orders? Would be interesting to see how many constant rows there are for these objects too. Might be unfeasable computationally..

Jose Brox (Oct 23 2024 at 12:02):

Daniel Weber ha dicho:

What sizes have weak central groupoids? I know of 2,4,8,16.

Also size 9! But all of them are strong, according to Mace4 (I provide an example). No other sizes up to 16.

Mace4 1485 model size 9.txt

thormikkel (Oct 23 2024 at 12:06):

Maybe the stronger conjecture here is that the orders of sporadic weak central groupoids are on the form 2^n? That would be beautiful

Jose Brox (Oct 23 2024 at 12:25):

thormikkel ha dicho:

Maybe the stronger conjecture here is that the orders of sporadic weak central groupoids are on the form 2^n? That would be beautiful

Does your definition of "sporadic" also include the condition of not being strong? Because there are no central groupoids of order 3, so all central groupoids of order 9 are sporadic in the sense of not being a product.

thormikkel (Oct 23 2024 at 12:30):

Jose Brox said:

thormikkel ha dicho:

Maybe the stronger conjecture here is that the orders of sporadic weak central groupoids are on the form 2^n? That would be beautiful

Does your definition of "sporadic" also include the condition of not being strong? Because there are no central groupoids of order 3, so all central groupoids of order 9 are sporadic in the sense of not being a product.

Yes that's right, strictly weak groupoids is probably a better nomenclature

thormikkel (Oct 23 2024 at 12:38):

Just to summarize the previous discussion:

We can define "strictly weak central groupoids" as magmas that obey #1485 but not the central groupoid law #168. We have 3 conjectures:

  • Finite strictly weak central groupoids have no idempotent element. This is in contrast to "strong" central groupoids that have n \sqrt{n} idempotents.

  • Finite strictly weak central groupoids have no distinct a,b so that a2=b2a^2 = b^2

  • Finite strictly weak sporadic central groupoids have order 2n2^n

"Sporadic" means that it is not a product of lower order magmas obeying 1485. In the finite realm it's probably smart to study the 6 unique sporadic variants of order 8 found here. They have in common that 4 or 8 pairs of rows in the operation table are identical.

I have no idea how you would construct an infinite strictly weak central groupoid.

Jose Brox (Oct 23 2024 at 12:42):

thormikkel

Ok, so no relation to the groupoids being products of smaller groupoids, I understood you wanted that in your definition.

Jose Brox (Oct 23 2024 at 12:44):

Otherwise, we have strictly weak central groupoids of size 18 (product of sizes 2 and 9), etc.

Daniel Weber (Oct 23 2024 at 12:49):

Some conclusions of 1485, from the partial greedy ruleset:

  • (x2)2x=x2=x(x2)2(x^2)^2 \diamond x = x^2 = x \diamond (x^2)^2
  • If xx=xx \diamond x = x and yx=yy \diamond x = y or xy=yx \diamond y = y then y=xy = x.
  • ((x2)2)2=x2((x^2)^2)^2 = x^2
  • If yz=zyy \diamond z = z \diamond y then ((yz)2)2=yz((y \diamond z)^2)^2 = y \diamond z (this subsumes the above)
  • If xy=yx \diamond y = y and yz=zy \diamond z = z then xz=yx \diamond z = y
  • If xy=xx \diamond y = x and zx=zz \diamond x = z then z2=zz^2 = z
  • If (x(yx))z=x(x \diamond (y \diamond x)) \diamond z = x for any y,zy, z then (x2)2=x(x^2)^2 = x

Kevin M (Oct 23 2024 at 14:06):

Terence Tao said:

This vaguely suggests to me that while many of the order 8 magmas are not direct products of an order 2 magma and an order 4 magma, they may still be "semi-direct products" in some sense. Having a semi-direct product construction in hand could well generate a lot of further magmas that could potentially be used to generate anti-implications.

Yes, hope for these phantom "semi-direct products" is what made me want to look at order 8 in the first place.
Unfortunately I have not been able to tease out the details.

Unfortunately the next size with these "products" would be 16, which would be very messy due to the multiple ways this could be factored. Than leaves 18 as the next size to give "clean" information on this phenomena. @Jose Brox 's calculations leave some hope that 18 might just barely be within the ability to calculate.

Terence Tao (Oct 23 2024 at 14:25):

Daniel Weber said:

Some conclusions of 1485, from the partial greedy ruleset:

This looks like a useful list. Is it possible to do a similar list for the other four outstanding equations on the other thread for 1516 and 854 on their respective threads?

Pace Nielsen (Oct 23 2024 at 16:32):

The "high/low" graph observation reminded me of something I tried a few days ago, but failed to make work. Maybe someone else will have more luck with it.

To control paths between vertices, use a tripartite graph. Thus, the vertices are partitioned into three sets A,B,CA,B,C (which I took to be countably infinite), with no edges between points in the same set. Each pair (x,y)(ABC)×(ABC)(x,y)\in (A\cup B\cup C)\times (A\cup B\cup C) must have a good path. If the points x,yx,y don't belong to the same set, then choose the connecting point to be in the third set. Otherwise, cyclically choose a connector (i.e., if x,yAx,y\in A choose the connector from BB, etc...).

The rules needed for the graph to do what we want turn into some sort of combinatorial claims about partitions of the pairs of vertices.

Terence Tao (Oct 23 2024 at 16:56):

Hmm, this construction would require there to exist a weak central groupoid on three elements (collapse each of A, B, C to a single point), which we know not to be the case.

But we do have weak central groupoids on four elements, so possibly some version of the construction works there. Ideally the base graph should be such that the directed graph is genuinely directed (x → y is not equivalent to y → x), this was the main issue with the two-element directed graph (in which 1 → 0, 0 → 1, and 0 → 0, or if you prefer high → low, low → high, and low → low) as the undirected nature of the graph blocked such constructions from being counterexamples to 161 or 1483.

Michael Bucko (Oct 23 2024 at 18:25):

Do we have a tptp for this one?

Jose Brox (Oct 23 2024 at 19:32):

thormikkel ha dicho:

Are any of the order 16 groupoids known to be "sporadic" and not a product of lower orders? Would be interesting to see how many constant rows there are for these objects too. Might be unfeasable computationally..

I've been running Mace4 on this, and after 5 CPU hours it has found an example of weak central groupoid of size 16 which is not strong (Idk if it is a product).

The finite magma explorer says it doesn't refute any unknown implication.

Mace4 1485 model size 16 not strong (array format).txt

thormikkel (Oct 23 2024 at 19:39):

Wow, nice work! This table also has the property that exactly half of them (8) come in identical pairs. There are certainly an intriguing number of powers of 2 associated to these objects!

Terence Tao (Oct 23 2024 at 19:43):

Don't have time for a full analysis, but the square A^2 of the adjacency matrix looks like it has an interesting structure:

  4   1   1   4   4   4   1   1   2   2   2   2   2   2   2   2
  1   1   1   1   1   1   1   1   1   1   1   1   1   1   1   1
  1   1   1   1   1   1   1   1   1   1   1   1   1   1   1   1
  4   1   1   4   4   4   1   1   2   2   2   2   2   2   2   2
  4   1   1   4   4   4   1   1   2   2   2   2   2   2   2   2
  4   1   1   4   4   4   1   1   2   2   2   2   2   2   2   2
  1   1   1   1   1   1   1   1   1   1   1   1   1   1   1   1
  1   1   1   1   1   1   1   1   1   1   1   1   1   1   1   1
  2   1   1   2   2   2   1   1   2   1   1   2   1   2   2   1
  2   1   1   2   2   2   1   1   1   2   2   1   2   1   1   2
  2   1   1   2   2   2   1   1   1   2   2   1   2   1   1   2
  2   1   1   2   2   2   1   1   2   1   1   2   1   2   2   1
  2   1   1   2   2   2   1   1   1   2   2   1   2   1   1   2
  2   1   1   2   2   2   1   1   2   1   1   2   1   2   2   1
  2   1   1   2   2   2   1   1   2   1   1   2   1   2   2   1
  2   1   1   2   2   2   1   1   1   2   2   1   2   1   1   2

Terence Tao (Oct 23 2024 at 19:44):

Perhaps this suggests a relabeling that would make the multiplication table more obviously structured.

Terence Tao (Oct 23 2024 at 19:47):

In fact it looks like a block version of the tensor product of the order two A^2 matrix [[2,1],[1,1]] with itself, which is [[4,2,2,1],[2,2,1,1],[2,1,2,1],[1,1,1,1]], I think.

Terence Tao (Oct 23 2024 at 19:49):

Have to run off to teach now, but my guess is that this order 16 magma can be grouped into four groups of four vertices (low-low, low-high, high-low, high-high), with the multiplication law behaving differently in each block (e.g. I would imagine the product of two low-low entries to be high-high). Based on the A^2 structure it seems like it is 0,3,4,5 which would be low-low, for instance.

Terence Tao (Oct 23 2024 at 19:50):

It could also just collapse to the direct product of an order 8 weak central groupoid and an order 2 weak central groupoid.

Terence Tao (Oct 23 2024 at 21:33):

OK, here is a permuted multiplication table:

 12  13  12  13   8   8  10  10   4   5   4   5   0   0   2   2
 12  13  12  13   8   8  10  10   4   5   4   5   0   0   2   2
 14  15  14  15   9   9  11  11   6   7   6   7   1   1   3   3
 14  15  14  15   9   9  11  11   6   7   6   7   1   1   3   3
  8  10   8  10   8   8  10  10   0   2   0   2   0   0   2   2
  9  11   9  11   9   9  11  11   1   3   1   3   1   1   3   3
  8  10   8  10   8   8  10  10   0   2   0   2   0   0   2   2
  9  11   9  11   9   9  11  11   1   3   1   3   1   1   3   3
  4   5   4   5   0   0   2   2   4   5   4   5   0   0   2   2
  4   5   4   5   0   0   2   2   4   5   4   5   0   0   2   2
  6   7   6   7   1   1   3   3   6   7   6   7   1   1   3   3
  6   7   6   7   1   1   3   3   6   7   6   7   1   1   3   3
  0   2   0   2   0   0   2   2   0   2   0   2   0   0   2   2
  1   3   1   3   1   1   3   3   1   3   1   3   1   1   3   3
  0   2   0   2   0   0   2   2   0   2   0   2   0   0   2   2
  1   3   1   3   1   1   3   3   1   3   1   3   1   1   3   3

or [[12, 13, 12, 13, 8, 8, 10, 10, 4, 5, 4, 5, 0, 0, 2, 2], [12, 13, 12, 13, 8, 8, 10, 10, 4, 5, 4, 5, 0, 0, 2, 2], [14, 15, 14, 15, 9, 9, 11, 11, 6, 7, 6, 7, 1, 1, 3, 3], [14, 15, 14, 15, 9, 9, 11, 11, 6, 7, 6, 7, 1, 1, 3, 3], [8, 10, 8, 10, 8, 8, 10, 10, 0, 2, 0, 2, 0, 0, 2, 2], [9, 11, 9, 11, 9, 9, 11, 11, 1, 3, 1, 3, 1, 1, 3, 3], [8, 10, 8, 10, 8, 8, 10, 10, 0, 2, 0, 2, 0, 0, 2, 2], [9, 11, 9, 11, 9, 9, 11, 11, 1, 3, 1, 3, 1, 1, 3, 3], [4, 5, 4, 5, 0, 0, 2, 2, 4, 5, 4, 5, 0, 0, 2, 2], [4, 5, 4, 5, 0, 0, 2, 2, 4, 5, 4, 5, 0, 0, 2, 2], [6, 7, 6, 7, 1, 1, 3, 3, 6, 7, 6, 7, 1, 1, 3, 3], [6, 7, 6, 7, 1, 1, 3, 3, 6, 7, 6, 7, 1, 1, 3, 3], [0, 2, 0, 2, 0, 0, 2, 2, 0, 2, 0, 2, 0, 0, 2, 2], [1, 3, 1, 3, 1, 1, 3, 3, 1, 3, 1, 3, 1, 1, 3, 3], [0, 2, 0, 2, 0, 0, 2, 2, 0, 2, 0, 2, 0, 0, 2, 2], [1, 3, 1, 3, 1, 1, 3, 3, 1, 3, 1, 3, 1, 1, 3, 3]] if you prefer.

Definitely looks like it is an extension of the product of two copies of the order 2 weak central groupoid. Unfortunately this base still has a directed graph that is symmetric (being the product of two directed graphs that already had this property), so it doesn't rule out anything new. [EDIT: actually I'm not 100% sure about this last statement, let me double check.]

If only we had a genuinely different base model than the order 2 model and the central groupoids to start from...

Terence Tao (Oct 23 2024 at 21:52):

OK, this particular example can be viewed as an extension of an order 2 weak central groupoid in two different ways: in the above numbering, either take 0-7 to be "low" and 8-15 to be "high", or 0-3, 8-11 to be "low" and 4-7, 12-15 to be "high". Either way the previous arguments rule out this type of construction generating new anti-implications.

Somehow one needs to modify this construction to allow a small number of carefully selected "high -> high" edges into the directed graph, but any naive attempt to do so ruins the axiom system unfortunately.

One very minor point by the way about the axiom system Claims 1-4 from my previous comment: Claims 2-3 are essentially redundant and can be dropped, because if they fail one can simply prune the graph by removing any edge that isn't the first (or isn't the second) component of a good path. So the only axioms needed are Claim 1 (any two points connected by exactly one good path) and Claim 4 (the strange axiom about 5-cycles).

Terence Tao (Oct 23 2024 at 21:56):

At this point the following ambitious conjecture does not seem to be ruled out: every weak central groupoid is either a central groupoid, or else has a homomorphism to the order 2 weak central groupoid (i.e., one can partition the elements into "low" and "high" elements, such that low ◇ low = high, and all other products are low). This ambitious conjecture would imply (in the positive) all the open implications! I don't think we have a counterexample for it currently. One weak piece of evidence against the conjecture: as the order 16 example (or just the product of two order 2 examples) shows, the homomorphism is not unique, so if the conjecture is true, it is unlikely that there is a "canonical" way to find the homomorphism.

Michael Bucko (Oct 23 2024 at 22:08):

Terence Tao schrieb:

At this point the following ambitious conjecture does not seem to be ruled out: every weak central groupoid is either a central groupoid, or else has a homomorphism to the order 2 weak central groupoid (i.e., one can partition the elements into "low" and "high" elements, such that low ◇ low = high, and all other products are low). This ambitious conjecture would imply (in the positive) all the open implications! I don't think we have a counterexample for it currently. One weak piece of evidence against the conjecture: as the order 16 example (or just the product of two order 2 examples) shows, the homomorphism is not unique, so if the conjecture is true, it is unlikely that there is a "canonical" way to find the homomorphism.

Maybe I don't fully understand, but wouldn't that imply that the complexity of weak central groupoids is, in some sense, inherently limited, i.e. that there must exist kind of a partitioning / centralization mechanism for creating such structures?

Terence Tao (Oct 23 2024 at 23:14):

Yes, this is an extremely strong conjecture, and I expect it to in fact be false, but we will need to find some way to refute it in order to achieve our expected goal of showing that all the open implications are false. Or, perhaps by some miracle, the conjecture is true, and then we have found a very rare collection of positive implications that were out of reach of ATPs.

Kevin M (Oct 24 2024 at 01:12):

Terence Tao said:

One weak piece of evidence against the conjecture: as the order 16 example (or just the product of two order 2 examples) shows, the homomorphism is not unique, so if the conjecture is true, it is unlikely that there is a "canonical" way to find the homomorphism.

If that size 16 is some kind of product with the order 2 "weak central" magma, then at a minimum it must factor into 3 parts (2)x(2)x(4). So unless I'm misunderstanding, are the two different ways to "factor out" an order 2 magma just due to there being two copies in there?

Terence Tao (Oct 24 2024 at 01:18):

The factoring may not be a direct product (or even a "semidirect product", whatever that means here): an order 16 magma might map onto an order 2 magma without there being an order 8 "quotient". (This is in contrast to the situation of groups, in which every surjective group homomorphism comes with a normal kernel that creates a short exact sequence. The difference is that a group contains an idempotent element - the identity - that can be used to locate the kernel, but the order 2 magma we have here has no idempotents.)

Kevin M (Oct 24 2024 at 01:19):

Regarding the strong conjecture, I'm more worried that the reason for a collection of implications that cannot be resolved is that they are true for all finite examples.

In that mind, I'm not sure how to do this, but is it possible to create an infinite "weak central" magma by taking this "hi/low" product with the (N=2) weak central magma and an infinite strong central magma? Since the infinite strong central magmas do not have idempotent elements, maybe this would lead to an interesting new example.

Terence Tao (Oct 24 2024 at 01:23):

Well, we can certainly take the direct product of any two 1485 magmas G, H and create another 1485 magma G × H (with the obvious law (g,h) ◇ (g',h') := (g ◇ g', h ◇ h')). So this already creates an infinite 1485 magma without idempotents. It doesn't get around the strong conjecture though, since if either of the two factors has a homomorphism onto the order 2 magma, then so does the dierct product.

Kevin M (Oct 24 2024 at 01:45):

Terence Tao said:

OK, here is a permuted multiplication table: ... [size 16 magma]

This has such nice blocks of 4x4 that I figured it could be simplified a bit more.

By swapping around elements I eventually figured out that this is equal to a direct product:

direct product (N=2 weak) x (N=2 weak) x (N=4 'strong')
=
 1 0   x   1 0   x   0 0 1 1
 0 0       0 0       2 2 3 3
                     0 0 1 1
                     2 2 3 3
=

 3 2  1 0    x   0 0 1 1
 2 2  0 0        2 2 3 3
                 0 0 1 1
 1 0  1 0        2 2 3 3
 0 0  0 0

=

[[ 12, 12, 13, 13,  8,  8,  9,  9,  4,  4,  5,  5,  0,  0,  1,  1],
 [ 14, 14, 15, 15, 10, 10, 11, 11,  6,  6,  7,  7,  2,  2,  3,  3],
 [ 12, 12, 13, 13,  8,  8,  9,  9,  4,  4,  5,  5,  0,  0,  1,  1],
 [ 14, 14, 15, 15, 10, 10, 11, 11,  6,  6,  7,  7,  2,  2,  3,  3],
 [  8,  8,  9,  9,  8,  8,  9,  9,  0,  0,  1,  1,  0,  0,  1,  1],
 [ 10, 10, 11, 11, 10, 10, 11, 11,  2,  2,  3,  3,  2,  2,  3,  3],
 [  8,  8,  9,  9,  8,  8,  9,  9,  0,  0,  1,  1,  0,  0,  1,  1],
 [ 10, 10, 11, 11, 10, 10, 11, 11,  2,  2,  3,  3,  2,  2,  3,  3],
 [  4,  4,  5,  5,  0,  0,  1,  1,  4,  4,  5,  5,  0,  0,  1,  1],
 [  6,  6,  7,  7,  2,  2,  3,  3,  6,  6,  7,  7,  2,  2,  3,  3],
 [  4,  4,  5,  5,  0,  0,  1,  1,  4,  4,  5,  5,  0,  0,  1,  1],
 [  6,  6,  7,  7,  2,  2,  3,  3,  6,  6,  7,  7,  2,  2,  3,  3],
 [  0,  0,  1,  1,  0,  0,  1,  1,  0,  0,  1,  1,  0,  0,  1,  1],
 [  2,  2,  3,  3,  2,  2,  3,  3,  2,  2,  3,  3,  2,  2,  3,  3],
 [  0,  0,  1,  1,  0,  0,  1,  1,  0,  0,  1,  1,  0,  0,  1,  1],
 [  2,  2,  3,  3,  2,  2,  3,  3,  2,  2,  3,  3,  2,  2,  3,  3]]

Then relabel (0, ...  ,15) -> (0, 2, 1, 3, 4, 5, 6, 7, 8, 10, 9, 11, 12, 13, 14, 15)

results in the size 16 table you found the nice layout in.

Pace Nielsen (Oct 24 2024 at 03:39):

Just to record it here: When a linear map xy=ax+by+cx\diamond y = ax+by+c satisfies 1485, we get the very nice (noncommutative) ring relations ab+ba=1ab+ba=1, a2=b2=0a^2=b^2=0, and ac+bc+c=0ac+bc+c=0. These are exactly the same relations needed to satisfy the central groupoid law.

Pace Nielsen (Oct 24 2024 at 04:55):

The translation-invariant ansatz was very helpful in defeating many implications. So I figured I'd record my thoughts about how it interacts with 1485. Take as carrier set a (possibly noncommutative) group GG. Given any x,y,zGx,y,z\in G, we can write y=px,z=qxy=px,z=qx for some (unique) p,qGp,q\in G.

Now, take ab=f(ba1)aa\diamond b = f(ba^{-1})a for some function f ⁣:GGf\colon G\to G. If I did the computations correctly [somebody should double-check me], this gives us the functional equation

f(f(f(pq1)q)p1f(p1)1)=p1f(p1)1.f(f(f(pq^{-1})q)p^{-1}f(p^{-1})^{-1}) = p^{-1}f(p^{-1})^{-1}.

Making the substitutions r=pq1r=pq^{-1} and s=p1s=p^{-1}, we have

f(f(f(r)r1s1)sf(s)1)=sf(s)1.f(f(f(r)r^{-1}s^{-1}) s f(s)^{-1})=s f(s)^{-1}.

To simplify things a little more, define a new function g ⁣:GGg\colon G\to G by the rule g(x):=f(x)xg(x):=f(x)x. Then the functional equation becomes

g(g(g(r)s1)g(r)g(s)1))=g(r)1g(g(r)s1)1.g(g(g(r)s^{-1})g(r)g(s)^{-1}))=g(r)^{-1}g(g(r)s^{-1})^{-1}.

Maybe this leads to gg being trivial; just putting it out there.

thormikkel (Oct 24 2024 at 07:09):

Terence Tao said:

Yes, this is an extremely strong conjecture, and I expect it to in fact be false, but we will need to find some way to refute it in order to achieve our expected goal of showing that all the open implications are false. Or, perhaps by some miracle, the conjecture is true, and then we have found a very rare collection of positive implications that were out of reach of ATPs.

Has this conjecture been tested against all unique order 8 weak central groupoids? It does look like your approach for magma #4 should hold generally for all of them.

It should be noted that @Jose Brox 's search for order 16 weak central groupoids was conducted using the assumptions a^2 != a, a^2 != b^2

EDIT: I stand corrected! See below.

Do we even know if 1485 + idempotency implies 168 for finite magmas? I think @Kevin M asked for this earlier. Daniel Weber's rule set has one condition that would imply the existence of an idempotent element but it might be too strong to give us something novel.

Looking at the infinite case, Knuth's paper has an elementary method of constructing free strong central groupoids. Naively following Knuth's procedure, we could try to produce a free weak central groupoid with any number of generators by the set of all words formed by the generators and the binary operation, replacing any subword on the form (b⋄a)⋄(a⋄(c⋄b)) with a (eq 1845) and ((b⋄c)⋄a)⋄(a⋄b) with a (the dual equation). Additionally we can replace (((a)^2)^2)^2 with a^2 by Daniel Weber's rule. This looks hideously constructive and not very useful, though.

The free central groupoid with any number of generators has no idempotent elements by theorem 8 in knuth's paper.

Jose Brox (Oct 24 2024 at 08:17):

I tend to run several searches in parallel. Now I'm looking for more weak-not-strong examples in size 16 and for any example in sizes 18 and 32.

Kevin M (Oct 24 2024 at 08:56):

@Jose Brox I would like to use Mace4 as well, but I tried to do a test run for size 4 and it only found 7 results. I used assign(max_models, -1). which I thought would make it run an exhaustive search. Is there something else I need to set?

Here's the input I used:

assign(domain_size, 4).
assign(max_models, -1).
formulas(theory).
    x = (y * x) * (x * (z * y)).
end_of_list.

Jose Brox (Oct 24 2024 at 09:01):

Kevin M ha dicho:

Jose Brox I would like to use Mace4 as well, but I tried to do a test run for size 4 and it only found 7 results.

Your input looks good, and I get the same result. Do you think there should be more than 7 models? For example in size 8 it finds many more (180).

EDITED: For longer computations, don't forget to also add assign(max_seconds,-1), and perhaps assign(max_megs,N) to N high enough (I usually put 2000).

Kevin M (Oct 24 2024 at 09:11):

Ah okay, to speed up the search it uses a simple heuristic to remove some (but it will not remove all) of the isomorphism in the search.
https://www.cs.unm.edu/~mccune/mace4/mace4.pdf
(section 5.4)

Wikipedia (referencing OEIS https://oeis.org/A283627) says there should be 12 central groupoid of size 4. And my brute for search found 12 weak ones as well. After removing isomorphism, there were only 1 strong and 1 weak, which Mace4 did find. I just wasn't aware it tried to filter out some of this.

Zoltan A. Kocsis (Z.A.K.) (Oct 24 2024 at 09:25):

In general, if Mace4 reports failure with Exhausted, then there really are no examples of a given size. But if it returns all_models, that doesn't mean it enumerated all structures of a given size, it only guarantees that it returned at least one representative from each isomorphism class.

You can pipe its output through | ./get_interps | ./isofilter to keep exactly one representative from each class.

Kevin M (Oct 24 2024 at 11:07):

thormikkel said:

Do we even know if 1485 + idempotency implies 168 for finite magmas?

I haven't made much progress on this.
eqn 168 implies that for all x there exists y such that x = (y ◇ y). (eqn 151 and some others with unknown implication from eqn1485, imply this as well)

So I tried proving at least 1485 + idempotent element -> (x ◇ x) is surjective, and variations on that idea.

The best I've gotten so far is that eqn 1485 implies that if there does exist an idempotent element x, at the very least _that_ element is not the result of another (y ◇ y). Hopefully I encoded that correctly in prover9: ( (x*x)=x ) -> (y=x | (y*y)!=x).

In graph terms that means if there is a good path: x->x->x then there is no good path of the form y -> x -> y.

Bhavik Mehta (Oct 24 2024 at 11:18):

thormikkel said:
Looking at the infinite case, Knuth's paper has an elementary method of constructing free strong central groupoids. Naively following Knuth's procedure, we could produce a free weak central groupoid with any number of generators by the set of all words formed by the generators and the binary operation, replacing any subword on the form (b⋄a)⋄(a⋄(c⋄b)) with a (eq 1845) and ((b⋄c)⋄a)⋄(a⋄b) with a (the dual equation). Additionally we can replace (((a)^2)^2)^2 with a^2 by Daniel Weber's rule. This looks hideously constructive and not very useful, though.

This isn't particularly likely to work - Knuth's procedure succeeds because the three rules he gives form a confluent rewriting system, and so multiplication in the free central groupoid as he gives it is well-defined and satisfies the appropriate law. But we don't have a confluent system for 1485 (I don't think a finite one exists), so the free weak central groupoid seems harder to construct. Note I formalised Knuth's construction in a PR to the equational repo

thormikkel (Oct 24 2024 at 11:20):

Bhavik Mehta said:

thormikkel said:
Looking at the infinite case, Knuth's paper has an elementary method of constructing free strong central groupoids. Naively following Knuth's procedure, we could produce a free weak central groupoid with any number of generators by the set of all words formed by the generators and the binary operation, replacing any subword on the form (b⋄a)⋄(a⋄(c⋄b)) with a (eq 1845) and ((b⋄c)⋄a)⋄(a⋄b) with a (the dual equation). Additionally we can replace (((a)^2)^2)^2 with a^2 by Daniel Weber's rule. This looks hideously constructive and not very useful, though.

This isn't particularly likely to work - Knuth's procedure succeeds because the three rules he gives form a confluent rewriting system, and so multiplication in the free central groupoid as he gives it is well-defined and satisfies the appropriate law. But we don't have a confluent system for 1485 (I don't think a finite one exists), so the free weak central groupoid seems harder to construct. Note I formalised Knuth's construction in a PR to the equational repo

That's interesting! Thank you :big_smile:

Pace Nielsen (Oct 24 2024 at 13:33):

Here is a combinatorial group-theory question related to the translation-invariant functional equation I described above. Does there exist a nontrivial group GG with a subset AGA\subseteq G such that every element of GG is a unique product of elements of AA? [Edited to add: By "product" I mean a product of exactly two elements.]

The only consequence of note I've found so far is the following. If 1=aa1=aa', then 1=aa1=a'a. So to keep 11 as a unique product we must have a=a1=aa=a^{-1}=a'.

Perhaps with a little more work a contradiction can be found, but I haven't yet found one.

Daniel Weber (Oct 24 2024 at 13:36):

Do the elements in the product have to be distinct?

Daniel Weber (Oct 24 2024 at 13:37):

If not, can't you write 11 in multiple ways by concatenating its representation with itself?

Pace Nielsen (Oct 24 2024 at 14:08):

@Daniel Weber Apologies! When I said "product" I should have clarified and said "product of two".

That said, sadly it turns out this is somewhat moot. The functional equation is also satisfied by 168 (the central groupoid law), so even if an example exists, it won't be what we want. [Edited: See a corrected statement below.]

Bernhard Reinke (Oct 24 2024 at 16:49):

Pace Nielsen said:

Here is a combinatorial group-theory question related to the translation-invariant functional equation I described above. Does there exist a nontrivial group GG with a subset AGA\subseteq G such that every element of GG is a unique product of elements of AA? [Edited to add: By "product" I mean a product of exactly two elements.]

The only consequence of note I've found so far is the following. If 1=aa1=aa', then 1=aa1=a'a. So to keep 11 as a unique product we must have a=a1=aa=a^{-1}=a'.

Perhaps with a little more work a contradiction can be found, but I haven't yet found one.

Even if it not relevant to an approach here, I think it is still a nice question. Apparently there are no groups of order smaller or equal 121 satisfying this. If you want a finite group, then its size has to be a square, and you can enumerate small enough groups to encode this into a relatively simple sat problem for each group. Here is my sage snippet:

import z3

def model_group(G):
    l = list(G)
    s = z3.Solver()
    x = [z3.Bool('x{}'.format(i)) for i in range(len(l))]
    for k in range(len(l)):
        pairs = [z3.And(x[l.index(l[k]*l[i]^-1)], x[i]) for i in range(len(l))]
        s.add(z3.AtLeast(*pairs,1))
        s.add(z3.AtMost(*pairs,1))
    return s

for n in range(1,17):
    print('Num Groups', gap.NumberSmallGroups(n*n))
    for m in range(1,1+int(gap.NumberSmallGroups(n*n))):
        print('checking {}, {}'.format(n, m))
        G = SmallPermutationGroup(n*n,m)
        if model_group(G).check() != z3.unsat:
            print(G)

One could also skip abelian groups. If you want to look at the opposite spectrum, apparently S4(7)S_4(7) is the smallest simple group whose order is a square ("only" 11760^2).

Terence Tao (Oct 24 2024 at 16:58):

I suspect a greedy construction on a free nonabelian group might work to solve @Pace Nielsen 's problem, but I haven't thought about it carefully.

Michael Bucko (Oct 24 2024 at 17:07):

@Bernhard Reinke Do I understand correctly? It's a play of a model (currently not in the ml sense), and a structure?
You have a model, which essentially uses the z3 solver with some constraints that you had in mind, and then you're creating a bunch of SPG instances and checking if those instances are compatible with the model (using the z3.unsat criterium)?

Theoretically, the criterium search could be done through some sort of grid search? ie they feel like hyperparameters?

As for the SPG instance, I guess one could construct a bunch of these objects with ml.

Relaxing the conditions related to the (criterium, object (in this case SPG instance)) could be done using some sort of policy optimization?

Terence Tao (Oct 24 2024 at 17:16):

Terence Tao said:

I suspect a greedy construction on a free nonabelian group might work to solve Pace Nielsen 's problem, but I haven't thought about it carefully.

Actually, because of @Pace Nielsen 's observation, we need to have a nontrivial involutive element to start the greedy algorithm, so the free group doesn't work. Instead, we work on the lamplighter group GG of transformations S:F2ZF2ZS: {\mathbf F}_2^{\mathbf Z} \to {\mathbf F}_2^{\mathbf Z} of the form
Sf(n)=f(n+h)+c(n) S f(n) = f(n+h) + c(n)
for some hZh \in {\mathbf Z} and some finitely supported c:ZF2.c: {\mathbf Z} \to {\mathbf F}_2. This lamplighter group is generated by an involution
Ef(n):=f(n)+1n=0 E f(n) := f(n) + 1_{n=0}
(lighting/unlighting the lamp at the lamplighter's location) and a shift map
Tf(n):=f(n+1) T f(n) := f(n+1)
(the lamplighter walks one unit to the right (or left, depending on conventions)). This is a countable solvable group, so we enumerate it in some arbitrary fashion.

Now we perform a greedy construction. At any partial stage one has a finite subset AA of the lamplighter group with the products AAA \cdot A all distinct, starting with A={E}A = \{E\} as the seed. Note that AA avoids the identity 11, but contains the identity in its product set; these properties will be preserved by the algorithm. Now let SS be the first element in the lamplighter group (in the above arbitrary enumeration) not covered already by AAA \cdot A (so in particular x1x \neq 1, let nn be a sufficiently large integer, and adjoin TnT^n and TnST^{-n} S to AA (note these elements are necessarily distinct from each other and from AA for nn large enough; they also avoid adding the identity element 11 to AA). Now S=(Tn)(TnS)S = (T^n) (T^{-n} S) is in the product set, and we have also added TnSTnST^{-n} S T^{-n} S, TnSAT^{-n} S A, ATnSA T^{-n} S, TnSTnT^{-n} S T^n, ATnA T^n, TnAT^n A, T2nT^{2n}, to the product set, and one can check that these are all disjoint from each other and from AAA \cdot A, SS for nn large enough (it is important here that S1S \neq 1 and AA does not contain 11).

Matthew Bolan (Oct 24 2024 at 17:22):

Isn't 1 still expressible in two different ways as ee1ee^{-1} and e1ee^{-1}e ? Perhaps one can fix it by working in the free group on countably many letters, and quotienting out by squares of the generators.

Terence Tao (Oct 24 2024 at 17:33):

Yes, my initial attempt was in error. I think I have now fixed it by working on the lamplighter group instead of the free group. Probably other groups that are sufficiently non-abelian while still containing at least one involution would also work.

Matthew Bolan (Oct 24 2024 at 17:34):

Yes, I believe I've convinced myself that the group I suggested works too, following essentially your original strategy.

Edit: To elaborate on this since apparently such examples are of interest, my group is the free product of countably many copies of F2\mathbb F_2, so the group a1,a2,...a12,a22,...\langle a_1,a_2,...|a_1^2,a_2^2,... \rangle. This group is countable. Beginning with A={a1}A = \{a_1\}, at each stage let xnx_n be the smallest element not yet in A2A^2, and ar,asa_r,a_s be two generators not appearing in any element of AA or xnx_n. Now adjoin to AA the elements asara_sa_r and arasxna_ra_sx_n, so that xnx_n is in A2A^2. Now the effect on A2A^2 is that we have added the sets {asarasar},{xn},{arasxnasar},{arasxnarasxn},asarA,arasxnA,Aarasxn,Aasar,\{a_sa_ra_sa_r\},\{x_n\},\{a_ra_sx_na_sa_r\},\{a_ra_sx_na_ra_sx_n\}, a_sa_rA, a_ra_sx_nA, Aa_ra_sx_n, Aa_sa_r, but these can be seen to be disjoint from each other and A^2 after observing 1∉A1 \not \in A and xn1x_n \not = 1. I think the fact I needed two generators ar,asa_r,a_s is a sign I have added too many inverses, and the essence of the example really lives in Z2FN.Z_2 * F_{\mathbb N}.

Terence Tao (Oct 24 2024 at 17:42):

I am now tempted to revive my previous translation-invariant attempt with such a greedy algorithm. Though the fact that @Pace Nielsen reported that his translation-invariant approach (which is likely equivalent) didn't produce counterexamples does mean that this approach might also not work, but I'd like to see what the obstruction is in the graph language.

Terence Tao (Oct 24 2024 at 17:47):

@Pace Nielsen Did you show that 1485 is immune to the entire translation-invariant approach, or only just to a special type of translation-invariant approach based on a AA=GA \cdot A = G construction?

Terence Tao (Oct 24 2024 at 17:49):

If nothing else, I believe we have at least constructed a new type of central groupoid, namely the lamplighter group GG with law given by SSUV=SUS \diamond SUV = SU whenever SGS \in G and U,VAU, V \in A (this is well-defined if every element of GG has a unique representation of the form AAA \cdot A, which the above greedy construction provides. The corresponding directed graph is the Cayley graph with edges SSUS \to SU, and the paths of length two (all of which are good) are SSUSUVS \to SU \to SUV.

Michael Bucko (Oct 24 2024 at 18:36):

Playing with z3. Getting this -- based on the output, what mistakes am I making? what am I missing?

equational_theories % python3 attack1485-v2.py

Searching for operation tables of size 2x2 that satisfy the following constraints:

1. Central Groupoid Law: S  (S U V) = S U for S  G and U, V  A

2. Magma Equation: x = (y  x)  (x  (z  y))

Found a satisfying operation table for n=2:

Operation Table ():

    0  1

0:  1  1

1:  1  0

Searching for operation tables of size 3x3 that satisfy the following constraints:

1. Central Groupoid Law: S  (S U V) = S U for S  G and U, V  A

2. Magma Equation: x = (y  x)  (x  (z  y))

No satisfying operation table found for n=3 with the given constraints.

Searching for operation tables of size 4x4 that satisfy the following constraints:

1. Central Groupoid Law: S  (S U V) = S U for S  G and U, V  A

2. Magma Equation: x = (y  x)  (x  (z  y))

Found a satisfying operation table for n=4:

Operation Table ():

    0  1  2  3

0:  2  2  3  3

1:  1  1  0  0

2:  1  1  0  0

3:  2  2  3  3

Searching for operation tables of size 5x5 that satisfy the following constraints:

1. Central Groupoid Law: S  (S U V) = S U for S  G and U, V  A

2. Magma Equation: x = (y  x)  (x  (z  y))

No satisfying operation table found for n=5 with the given constraints.

Pace Nielsen (Oct 24 2024 at 18:44):

Terence Tao said:

Pace Nielsen Did you show that 1485 is immune to the entire translation-invariant approach, or only just to a special type of translation-invariant approach based on a AA=GA \cdot A = G construction?

Actually I may have gotten the logic backwards, and 1485 might not be immune. The functional equation for 1485 is

f(f(f(r)r1s1)sf(s)1)=sf(s)1.f(f(f(r)r^{-1}s^{-1})sf(s)^{-1})=sf(s)^{-1}.

Taking u=f(r)r1s1 u =f(r)r^{-1}s^{-1} we get

f(f(u)sf(s)1)=sf(s)1,f(f(u)sf(s)^{-1})=sf(s)^{-1},

which is the functional equation for 168. My error was in thinking that uu is arbitrary (independent of ss) in GG. (If it is, then indeed 168 must always hold.)

What motivated me to consider the AA=GA\cdot A=G construction was taking v:=f(u)sf(s)1v:=f(u)sf(s)^{-1} (which, again, might not a priori be an arbitrary element of GG independent of uu, even if u,sGu,s\in G are arbitrary) and then the functional equation for 168 transforms into

f(f(u)f(v))=f(v).f(f(u)f(v))=f(v).

Terence Tao (Oct 24 2024 at 20:30):

Ok. I am actually cautiously optimistic that one can perform greedy extension here on a sufficiently free group. Basically if one sets f(s) to a novel element c then this creates relatively few 5-cycles in the Cayley graph and I think I can fix all the resultant consistency conditions manually. Will try this later today.

@Michael Bucko : i think you may have omitted the axioms that the multiplication on G is a group. Also that the products of A with itself have to be distinct (though this is implicitly covered by 1485).

Bernhard Reinke (Oct 24 2024 at 21:38):

Bernhard Reinke said:

Even if it not relevant to an approach here, I think it is still a nice question. Apparently there are no groups of order smaller or equal 121 satisfying this. If you want a finite group, then its size has to be a square, and you can enumerate small enough groups to encode this into a relatively simple sat problem for each group.

My code run until 225, there are already too many groups of order 256 to finish in a reasonable time. Still curious whether it is possible for a finite group, even if not relevant.

Terence Tao (Oct 24 2024 at 22:16):

OK, the greedy algorithm creates a rapidly proliferating set of side conditions that doesn't seem to close easily. I can't definitively rule it out as impossible, but it looks challenging. For actual central groupoids, with no bad paths, there wasn't an issue, but the moment there is a bad path, one always has to make sure that it doesn't get surrounded by two good paths and then joined up into a 5-cycle. However, adding more edges to a partially constructed graph by defining a new value for the magma multiplication shift function f(h)f(h) tends to create a bunch of new good paths. So one needs an axiom for the existing good paths to prevent the above scenario from happening. But then that axiopm in turn could get destroyed by adding more edges until another axiom is added... and it rapidly gets bad from there.

I have a vague idea that this might be fixable if one sets up a construction with a very small number of bad paths, which are in turn surrounded by a very small number of good paths, which in turn are somehow prevented from forming 5 cycles, but nothing concrete yet.

Terence Tao (Oct 24 2024 at 22:31):

By the way, it looks quite easy to construct infinite central groupoids greedily (without needing any translation invariant structure). Given a finitely constructed directed graph, with the property that any two points are connected by at most one length two path, and find a pair of points x, y that are not yet connected by such a path, locate a novel vertex c not previously used, and adjoin the edges x → c → y to the graph. This now connects x with y (and makes some other length path connections as well), but does not make any pair of points connected by more than one length two path (the key point is the new vertex c has only one incoming edge and one outgoing edge, so it can't cause much trouble). Iterating this on a countably infinite vertex set, one eventually gets a central groupoid.

I'm going to try the same procedure but starting with a seed directed graph in which a few paths are already labeled "bad" (and the rest labeled "good"). I want to add edges in such a way that I mostly create good paths (basically I want to solve more problems than I create), but occasionally I think I will have to introduce new bad paths to protect old bad paths. I'm hoping I can "cauterize" this by surrounding these new bad paths by several new good paths, but I don't see how to make this procedure converge yet.

Jose Brox (Oct 25 2024 at 00:02):

Jose Brox ha dicho:

I tend to run several searches in parallel. Now I'm looking for more weak-not-strong examples in size 16 and for any example in sizes 18 and 32.

Apparently, Mace4 found 3 million! weak central groupoids which are not strong in size 16. I'm not being able to kill the process, which still continues (I asked for all the models). Let's see if it is writing them to the output file... In the meantime I will re-run the program with a 200 limit.

Zoltan A. Kocsis (Z.A.K.) (Oct 25 2024 at 01:09):

@Jose Brox Interesting! Are you piping the output through get_interps | isofilter? Otherwise, the 3 million number could come from a few isomorphism classes, since there are plenty of ways to permute 16 elements.

Pace Nielsen (Oct 25 2024 at 02:01):

Would ATP's allow us to get a good idea of whether the translation-invariant technique is blocked? Namely, work in the language of (noncommutative) groups with an extra binary symbol \diamond and an extra unary symbol ff. We take as identities the axioms of groups, together with the 1485 law and the connecting identity ab=f(ab1)ba\diamond b = f(ab^{-1})b, and see if the ATP says that the equations we want to avoid are forced to hold.

Daniel Weber (Oct 25 2024 at 03:11):

Using the tptp

fof(gmulassoc, axiom, gmul(X, gmul(Y, Z)) = gmul(gmul(X, Y), Z)).
fof(gmulid, axiom, gmul(X, e) = X).
fof(gmulinv, axiom, gmul(X, inv(X)) = e).
fof(muldef, axiom, mul(X, Y) = gmul(f(gmul(X, inv(Y))), Y)).
fof(eq1485, axiom, X = mul(mul(Y, X), mul(X, mul(Z, Y)))).
fof(eq151, conjecture, X = mul(mul(X, X), mul(X, X))).

It's definitely blocked - Vampire can produce a proof (it's fairly long, though)

Andy Jiang (Oct 25 2024 at 03:14):

Bernhard Reinke said:

Pace Nielsen said:

Here is a combinatorial group-theory question related to the translation-invariant functional equation I described above. Does there exist a nontrivial group GG with a subset AGA\subseteq G such that every element of GG is a unique product of elements of AA? [Edited to add: By "product" I mean a product of exactly two elements.]

The only consequence of note I've found so far is the following. If 1=aa1=aa', then 1=aa1=a'a. So to keep 11 as a unique product we must have a=a1=aa=a^{-1}=a'.

Perhaps with a little more work a contradiction can be found, but I haven't yet found one.

Even if it not relevant to an approach here, I think it is still a nice question. Apparently there are no groups of order smaller or equal 121 satisfying this. If you want a finite group, then its size has to be a square, and you can enumerate small enough groups to encode this into a relatively simple sat problem for each group. Here is my sage snippet:

import z3

def model_group(G):
    l = list(G)
    s = z3.Solver()
    x = [z3.Bool('x{}'.format(i)) for i in range(len(l))]
    for k in range(len(l)):
        pairs = [z3.And(x[l.index(l[k]*l[i]^-1)], x[i]) for i in range(len(l))]
        s.add(z3.AtLeast(*pairs,1))
        s.add(z3.AtMost(*pairs,1))
    return s

for n in range(1,17):
    print('Num Groups', gap.NumberSmallGroups(n*n))
    for m in range(1,1+int(gap.NumberSmallGroups(n*n))):
        print('checking {}, {}'.format(n, m))
        G = SmallPermutationGroup(n*n,m)
        if model_group(G).check() != z3.unsat:
            print(G)

One could also skip abelian groups. If you want to look at the opposite spectrum, apparently S4(7)S_4(7) is the smallest simple group whose order is a square ("only" 11760^2).

Tangent on this question of writing a group as product of elements of A:
OK I'm not sure if this is correct, but for finite groups if you write the sum of the elements of A in the group algebra of G over the rationals as E, then E^2 is the sum of all of the elements of G in the group algebra. so E^2 acts trivially on every simple rational representation except for the trivial one, where it acts as n^2 (where n^2 is the size of the group)

so the trace of E on every nontrivial simple representation is 0 and the trace on trivial rep is n. Hence the trace of E on the group algebra itself is n. Isn't this a contradiction b/c the trace of nontrivial elements of the group on the group algebra is zero?

Daniel Weber (Oct 25 2024 at 03:23):

Daniel Weber said:

Using the tptp

fof(gmulassoc, axiom, gmul(X, gmul(Y, Z)) = gmul(gmul(X, Y), Z)).
fof(gmulid, axiom, gmul(X, e) = X).
fof(gmulinv, axiom, gmul(X, inv(X)) = e).
fof(muldef, axiom, mul(X, Y) = gmul(f(gmul(X, inv(Y))), Y)).
fof(eq1485, axiom, X = mul(mul(Y, X), mul(X, mul(Z, Y)))).
fof(eq151, conjecture, X = mul(mul(X, X), mul(X, X))).

It's definitely blocked - Vampire can produce a proof (it's fairly long, though)

It can also do this for other simple variants:

  • ab=f(a)ba \diamond b = f(a) b
  • ab=f(ab)ba \diamond b = f(ab) b
  • ab=f(ba)ba \diamond b = f(ba) b
  • ab=f(a1b)ba \diamond b = f(a^{-1} b) b
  • ab=f(ba1)ba \diamond b = f(b a^{-1}) b

Zoltan A. Kocsis (Z.A.K.) (Oct 25 2024 at 03:27):

@Daniel Weber Neat! How far can it go? Can it prove that a semigroup examples of the first 3 forms (and e.g. regular semigroup constructions for the last two) are blocked as well?

Zoltan A. Kocsis (Z.A.K.) (Oct 25 2024 at 03:33):

A search for the shortest scheme which Vampire cannot prove blocked might also be interesting.

One could also try to see if there are finite models of the required form of the antecedent at all.

IIRC this is one of the groups where my search which tried various group words over a family of infinite groups as schemes couldn't even generate models of the antecedent. Wish I had time to set something like that up again.

Daniel Weber (Oct 25 2024 at 03:36):

Yes for 1. For 2 and 3 it can't prove it in one minute, but with a left identity 2 is blocked and with a right identity 3 is blocked but with a left identity 2 and 3 are blocked, and have no nontrivial models. With a right identity only 3 looks blocked

Daniel Weber (Oct 25 2024 at 03:38):

For 1 it can actually prove that there are no nontrivial models

Daniel Weber (Oct 25 2024 at 03:39):

For 2 and 3 it can find finite models, even with a right identity

Pace Nielsen (Oct 25 2024 at 04:44):

Daniel Weber said:

Using the tptp

fof(gmulassoc, axiom, gmul(X, gmul(Y, Z)) = gmul(gmul(X, Y), Z)).
fof(gmulid, axiom, gmul(X, e) = X).
fof(gmulinv, axiom, gmul(X, inv(X)) = e).
fof(muldef, axiom, mul(X, Y) = gmul(f(gmul(X, inv(Y))), Y)).
fof(eq1485, axiom, X = mul(mul(Y, X), mul(X, mul(Z, Y)))).
fof(eq151, conjecture, X = mul(mul(X, X), mul(X, X))).

It's definitely blocked - Vampire can produce a proof (it's fairly long, though)

Excellent, and thank you for posting that code! Now I can piece together the syntax enough to try some of my own explorations with Vampire.

There is one thing I'm a little confused about. It says that the termination reason is "Refutation". I would have thought that this means that Vampire refuted the conjecture, rather than proved it. Our actual goal is to get a refutation, so what am I missing?

Daniel Weber (Oct 25 2024 at 04:45):

It negates the conjecture and searches for a contradiction, "Refutation" means it found a contradiction

Michael Bucko (Oct 25 2024 at 04:54):

Now the script including the feedback from @Terence Tao and the tptp definition from @Daniel Weber . That's the output (as always, I'd appreciate any feedback; I'm still learning):

equational_theories % python3 attack1485-v3.py

Searching for group structures of size 2 with the given constraints:

Found a satisfying group structure for n=2:

Operation Table (mul):

    0  1

0:  8 11

1: 14 15

Group Multiplication Table (gmul):

    0  1

0:  0  1

1:  1  3

Inverse Table (inv):

Element : Inverse

  0    :   5

  1    :   6

Function f:

Element : f(Element)

  0    :      7

  1    :      16

Multiplication Table (mul):

    0  1

0:  8 11

1: 14 15

Searching for group structures of size 3 with the given constraints:

Found a satisfying group structure for n=3:

Operation Table (mul):

    0  1  2

0: 19 22 25

1: 28 29 32

2: 35 38 39

Group Multiplication Table (gmul):

    0  1  2

0:  0  1  2

1:  1  3  4

2:  2  5  6

Inverse Table (inv):

Element : Inverse

  0    :   15

  1    :   16

  2    :   17

Function f:

Element : f(Element)

  0    :      18

  1    :      40

  2    :      41

Multiplication Table (mul):

    0  1  2

0: 19 22 25

1: 28 29 32

2: 35 38 39

Searching for group structures of size 4 with the given constraints:

Found a satisfying group structure for n=4:

Operation Table (mul):

    0  1  2  3

0: 45 48 51 54

1: 57 58 61 64

2: 67 70 71 74

3: 77 80 83 84

Group Multiplication Table (gmul):

    0  1  2  3

0:  0  1  2  3

1:  1  4  5  6

2:  2  7  8  9

3:  3 10 11 12

Inverse Table (inv):

Element : Inverse

  0    :   40

  1    :   41

  2    :   42

  3    :   43

Function f:

Element : f(Element)

  0    :      44

  1    :      85

  2    :      86

  3    :      87

Multiplication Table (mul):

    0  1  2  3

0: 45 48 51 54

1: 57 58 61 64

2: 67 70 71 74

3: 77 80 83 84

That's the script:

attack1485-v3.py

Daniel Weber (Oct 25 2024 at 04:59):

You also need to bound the range of the operations - everything should operate on the same set

Michael Bucko (Oct 25 2024 at 05:18):

Yes, completely forgot that. You're right @Daniel Weber :man_facepalming:

Here's the output:

python3 attack1485-v3.py

============================================================

Searching for group structures of size 2 with the given constraints:

No satisfying group structure found for n=2 with the given constraints.

============================================================

Searching for group structures of size 3 with the given constraints:

No satisfying group structure found for n=3 with the given constraints.

============================================================

Searching for group structures of size 4 with the given constraints:

No satisfying group structure found for n=4 with the given constraints.

============================================================

Searching for group structures of size 5 with the given constraints:

No satisfying group structure found for n=5 with the given constraints.

============================================================

Searching for group structures of size 6 with the given constraints:

No satisfying group structure found for n=6 with the given constraints.

============================================================

Searching for group structures of size 7 with the given constraints:

No satisfying group structure found for n=7 with the given constraints.

Here's the improved version:
attack1485-v3.py

Please share your feedback. It doesn't find any solutions, but I can implement better search, new algos, and so on. Any suggestions?

Also, the script could eventually just accept the tptp from the command line.

Terence Tao (Oct 25 2024 at 06:02):

@Andy Jiang - nice use of representation theory here!

In a different direction, I was trying to figure out how to relax the high/low structure and still get viable weak central groupoids. Here is one observation. Suppose one has a directed graph on a vertex set that is divided into high and low classes that obey the following two axioms:

  • Axiom 1 There are no 5-cycles abcdeaa \to b \to c \to d \to e \to a where b,c,db,c,d are low and a,ea,e are high.
  • Axiom 2 Any two vertices x,yx,y are connected by exactly one path xzyx \to z \to y that is not of the form lowlowlowlow \to low \to low. (They are allowed to be connected by arbitrarily many paths of the form lowlowlowlow \to low \to low, though.)

Then this creates a weak central groupoid, by setting z=xyz = x \diamond y in axiom 2.

Right now we only know how to create such weak central groupoids by forbidding highhighhigh \to high edges (which makes Axiom 1 automatic), in which case we can for instance create some order 8 examples. (Genuine central groupoids correspond to the case where there are no lowlowlowlow \to low \to low paths.) I would like to create some infinite examples, and my plan for doing so is to try to first find some interesting directed graphs on high and low vertices that obey the following weaker set of axioms:

  • Axiom 1 There are no 5-cycles abcdeaa \to b \to c \to d \to e \to a where b,c,db,c,d are low and a,ea,e are high.
  • Axiom 2' Any two vertices x,yx,y are connected by at least one path xzyx \to z \to y that is not of the form lowlowlowlow \to low \to low.

I have a vague sense that I can "blow up" any graph of this form, turning each vertex class into an infinite set, and perform a greedy construction on the resulting object, to create a genuine weak central groupoid that could be used for counterexamples.

But I first need some non-trivial examples of directed graphs on high and low vertices that obey Axiom 1 and Axiom 2'. Again, we have examples if we forbid highhighhigh \to high edges, for instance the order 2 graph with one low vertex 0, one high vertex 1, and edges 000 \to 0, 010 \to 1, and 101 \to 0. But to get interesting examples I need a directed graph with at least one highhighhigh \to high edge (and at least one lowlowlowlow \to low \to low path). I played around a little bit with pen and paper but was not able to find such a graph... perhaps an ATP could see if there are any small order examples?

Michael Bucko (Oct 25 2024 at 06:18):

Terence Tao schrieb:

But I first need some non-trivial examples of directed graphs on high and low vertices that obey Axiom 1 and Axiom 2'. Again, we have examples if we forbid highhighhigh \to high edges, for instance the order 2 graph with one low vertex 0, one high vertex 1, and edges 000 \to 0, 010 \to 1, and 101 \to 0. But to get interesting examples I need a directed graph with at least one highhighhigh \to high edge (and at least one lowlowlowlow \to low \to low path). I played around a little bit with pen and paper but was not able to find such a graph... perhaps an ATP could see if there are any small order examples?

I had a look at this and got this:

Found a satisfying directed graph for n=9:

Vertex Classification:

High vertices: [2, 4]

Low vertices: [0, 1, 3, 5, 6, 7, 8]

Edge List:

0  1

0  2

0  3

0  5

0  6

0  7

0  8

1  0

1  2

1  3

1  5

1  6

1  7

1  8

2  0

2  1

2  3

2  4

2  5

2  6

2  7

2  8

3  0

3  1

3  2

3  5

3  6

3  7

3  8

4  2

5  0

5  1

5  2

5  3

5  6

5  7

5  8

6  0

6  1

6  2

6  3

6  5

6  7

6  8

7  0

7  1

7  2

7  3

7  5

7  6

7  8

8  0

8  1

8  2

8  3

8  5

8  6

8  7

Adjacency Matrix:

    0  1  2  3  4  5  6  7  8

0: 1  1  1  1  0  1  1  1  1

1: 1  1  1  1  0  1  1  1  1

2: 1  1  0  1  1  1  1  1  1

3: 1  1  1  1  0  1  1  1  1

4: 0  0  1  0  1  0  0  0  0

5: 1  1  1  1  0  1  1  1  1

6: 1  1  1  1  0  1  1  1  1

7: 1  1  1  1  0  1  1  1  1

8: 1  1  1  1  0  1  1  1  1

But perhaps I am still missing something.

Michael Bucko (Oct 25 2024 at 06:19):

That's the main adjustment, for the reference. The good news is that every time we're experimenting with new approaches or constraints, it's not that much code.

    high_high_edges = [z3.And(edge[x][y], is_high[x], is_high[y])
                       for x, y in itertools.product(elements, repeat=2) if x != y]
    s.add(z3.Or(high_high_edges))

    low_low_low_paths = []
    for x, y, z in itertools.product(elements, repeat=3):
        low_low_low_paths.append(z3.And(edge[x][y], edge[y][z], is_low[x], is_low[y], is_low[z]))
    s.add(z3.Or(low_low_low_paths))

Terence Tao (Oct 25 2024 at 06:30):

Thanks, that was quick! I was able to confirm that this did indeed obey Axiom 1 and Axiom 2'. Need to call it a night for now, but tomorrow I will see if I can use this structure to create an infinite directed graph obeying Axiom 1 and Axiom 2 (the idea is to have nine infinite vertex classes V0,,V8V_0,\dots,V_8, with V2,V4V_2, V_4 being low vertices and the rest being high vertices, and start greedily attaching edges in a manner consistent with the provided directed graph (e.g., no edges from V0V_0 to V4V_4 are permitted) - in order to automatically satisfy Axiom 1 - until Axiom 2 is completely satisfied. Fingers crossed that this works and actually generates useful examples of weak central groupoids...

Terence Tao (Oct 25 2024 at 06:38):

Sorry, I just realized that in order for these examples to have a shot at contradicting 1483, I need one more axiom:

  • Axiom anti-1483. There is at least one modified 5-cycle abcdea \to b \to c \to d \to e, eae \to a (i.e. a 5-cycle with the final edge reversed in orientation) such that b,c,db,c,d are low and a,ea,e are high.

(I haven't thought carefully about what the corresponding axiom would need to be to contradict 166, but I'll think about that tomorrow.) Unfortunately, the graph you provided doesn't obey this axiom. Would it be possible to conduct another search to see if there is a graph that obeys Axioms 1, 2', and anti-1483?

Daniel Weber (Oct 25 2024 at 06:39):

Is this supposed to be aea \to e? Otherwise I think that's just a regular 5 cycle

Michael Bucko (Oct 25 2024 at 07:10):

For the time being, I assumed abcdea \to b \to c \to d \to e, eae \to a and added (anything wrong in this code?)

        modified_cycles = []
        for a, b, c, d, e in itertools.product(elements, repeat=5):
            cond = z3.And(
                is_high[a],
                is_low[b],
                is_low[c],
                is_low[d],
                is_high[e],
                edge[a][b],
                edge[b][c],
                edge[c][d],
                edge[e][d],
                edge[e][a]
            )
            modified_cycles.append(cond)
        if modified_cycles:
            s.add(z3.Or(modified_cycles))
        else:
            s.add(z3.BoolVal(False))

but then I'm getting nothing

python3 attack1485-v5.py

============================================================

Searching for directed graphs of size 5 with the given constraints:

No satisfying directed graph found for n=5 with the given constraints.

============================================================

Searching for directed graphs of size 6 with the given constraints:

No satisfying directed graph found for n=6 with the given constraints.

============================================================

Searching for directed graphs of size 7 with the given constraints:

No satisfying directed graph found for n=7 with the given constraints.

============================================================

Searching for directed graphs of size 8 with the given constraints:

No satisfying directed graph found for n=8 with the given constraints.

============================================================

Jose Brox (Oct 25 2024 at 08:24):

Zoltan A. Kocsis (Z.A.K.) ha dicho:

Jose Brox Interesting! Are you piping the output through get_interps | isofilter? Otherwise, the 3 million number could come from a few isomorphism classes, since there are plenty of ways to permute 16 elements.

Yes, I'm aware of that! But I'm running Mace4 from Windows, with the Prover9/Mace4 GUI, and the isofilter is applied after the output file is generated! But the program is not allowing me to kill the run, possibly it is just writing the 3 million groupoids to the output file... Nevertheless, isofilter would take ages to run over 3 million models, I don't think it is feasible (100-200 models already may take some time). I didn't expect that many examples popping up at the same time.

Jose Brox (Oct 25 2024 at 08:26):

On the other hand, I rerun the program with a 200 models limit, and after some isofiltering, we have 181 (*) nonisomorphic weak central groupoids of size 16 which are not strong (file attached).

(*) Mace4 informs there are 180, then in other message it says 181 :sweat_smile:

Mace4 1485 size 16 181 models isofilter portable.txt

Jose Brox (Oct 25 2024 at 08:26):

I also have 100 models of size 18, but isofilter is still running on them.

Michael Bucko (Oct 25 2024 at 08:27):

Jose Brox schrieb:

Zoltan A. Kocsis (Z.A.K.) ha dicho:

Jose Brox Interesting! Are you piping the output through get_interps | isofilter? Otherwise, the 3 million number could come from a few isomorphism classes, since there are plenty of ways to permute 16 elements.

Yes, I'm aware of that! But I'm running Mace4 from Windows, with the Prover9/Mace4 GUI, and the isofilter is applied after the output file is generated! But the program is not allowing me to kill the run, possibly it is just writing the 3 million groupoids to the output file... Nevertheless, isofilter would take ages to run over 3 million models, I don't think it is feasible (100-200 models already may take some time). I didn't expect that many examples popping up at the same time.

Is there any way to parallelize this? If so, I could help with re-building the code.

Jose Brox (Oct 25 2024 at 08:33):

Michael Bucko ha dicho:

Is there any way to parallelize this? If so, I could help re-building the code.

A priori yes, if we divide a batch of models in 2, then we can produce the isoclasses on each subbatch, then join the results and produce isoclasses again.

But I don't know if we need this. Possibly, a more interesting tool to have right now would be a way to automatically feed the finite magma explorer with a batch of models in the format exported by Mace4 (perhaps this can already be done, but Idk how). I'm assuming that we still don't know the implications from 1485 to be immune to finite models.

Jose Brox (Oct 25 2024 at 08:51):

Jose Brox ha dicho:

Possibly, a more interesting tool to have right now would be a way to automatically feed the finite magma explorer with a batch of models in the format exported by Mace4

I suppose this can be done by just taking the program which runs when hitting the "process" button at the finite magma explorer, and producing a loop with it, but then we need also to read the interesting output, i.e., if any unknown implications are refuted.

thormikkel (Oct 25 2024 at 09:20):

Jose Brox said:

On the other hand, I rerun the program with a 200 models limit, and after some isofiltering, we have 181 (*) nonisomorphic weak central groupoids of size 16 which are not strong (file attached).

(*) Mace4 informs there are 180, then in other message it says 181 :sweat_smile:

Mace4 1485 size 16 181 models isofilter portable.txt

Very interesting! I don't have access to a computer right now so I can't check, but my guess is that if any of these don't map to the "primordial" order 2 magma by homomorphism (or is not otherwise the direct product of a strong central groupoid and the order 2 weak groupoid) they would need to have the property that half of the rows do not come in pairs of 2.

I guess the order 16 strictly weak centroid with only two distinct rows maps to the order 2 magma without any other "factoring"?

Andy Jiang (Oct 25 2024 at 10:53):

Jose Brox said:

On the other hand, I rerun the program with a 200 models limit, and after some isofiltering, we have 181 (*) nonisomorphic weak central groupoids of size 16 which are not strong (file attached).

(*) Mace4 informs there are 180, then in other message it says 181 :sweat_smile:

Mace4 1485 size 16 181 models isofilter portable.txt

If I'm not mistaken all of your size 16 examples surject onto the product of 2 copies of the size 2 weak central groupoid with the equivalence relation defined by the equivalence classes
[0, 3, 4, 5]
[1, 2, 6, 7]
[8, 11, 13, 14]
[9, 10, 12, 15]

Zoltan A. Kocsis (Z.A.K.) (Oct 25 2024 at 11:42):

Jose Brox said:

On the other hand, I rerun the program with a 200 models limit, and after some isofiltering, we have 181 (*) nonisomorphic weak central groupoids of size 16 which are not strong (file attached).

(*) Mace4 informs there are 180, then in other message it says 181 :sweat_smile:

Mace4 1485 size 16 181 models isofilter portable.txt

Neat! Just to have it on record, what is the number of the equation which the constants c1...c3 contradict? The strong c.g. law, I presume, but I don't yet know its number by heart.

One of the annoying things about the Prover9/Mace4 distribution is that it contains no built-in tool to filter out irrelevant structure from the interpretations. This means that some of the magmas among the 180 may still be isomorphic as weak central groupoids, but with non-isomorphic "choice of counterexamples".

Zoltan A. Kocsis (Z.A.K.) (Oct 25 2024 at 12:06):

Not coincidentally, two related (not necessarily high-impact at this stage) tasks that I don't have time to tackle but lurkers might be interested in:

  1. Software: Write a script to filter out irrelevant structure from the output of get_interps. Something like keep.py "f" which reads input in get_interps format from stdin and writes only the functions that are called f into stdout for piping into isofilter.
  2. Math: Is it possible to characterize weak-but-not-strong central groupoids using an equational Horn theory?

Michael Bucko (Oct 25 2024 at 12:18):

Zoltan A. Kocsis (Z.A.K.) schrieb:

  1. Math: Is it possible to characterize weak-but-not-strong central groupoids using an equational Horn theory?

You mean sth along these lines?

Weak centrality locally (works for some, but not all):
weak_central = And(op(a, b) == op(b, a), op(a, c) == op(c, a))

And avoid strong:
strong_central = ForAll([a, b], op(a, b) == op(b, a))

Then imply weak AND not strong centrality
horn_setup = Implies(weak_central, Not(strong_central))

and then simply search?

Michael Bucko (Oct 25 2024 at 12:26):

Zoltan A. Kocsis (Z.A.K.) schrieb:

  1. eads input in get_interps format from stdin and writes only the functions that are called f into stdout for piping into isofilter.

As for the eads, this is a one-liner? (or simply a bash script that runs mace4 with input)

eads get_interps --input=stdin --output=stdout | eads filter --function-name=f | eads isofilter

Zoltan A. Kocsis (Z.A.K.) (Oct 25 2024 at 13:13):

Michael Bucko said:

You mean sth along these lines?

Can't look at this deeply right now (a paper revision deadline managed to sneak up on me), but this looks about right if you want to search for such a characterization. Keep in mind that the one-element model will still be a strong central groupoid, so one has to state something like "if the Horn clause holds, and the structure is strong central, then it is in fact trivial".

Michael Bucko (Oct 25 2024 at 13:13):

Michael Bucko schrieb:

For the time being, I assumed abcdea \to b \to c \to d \to e, eae \to a and added (anything wrong in this code?)

        modified_cycles = []
        for a, b, c, d, e in itertools.product(elements, repeat=5):
            cond = z3.And(
                is_high[a],
                is_low[b],
  ..
..

but then I'm getting nothing

python3 attack1485-v5.py

Ran the updated script for 5h+ and got to

No satisfying group structure found for n=16 with the given constraints.

Terence Tao (Oct 25 2024 at 14:23):

Daniel Weber said:

Is this supposed to be aea \to e? Otherwise I think that's just a regular 5 cycle

Yes, that was a typo, thanks. Now corrected.

It's a pity that no such examples have been located yet, because I think such an example would indeed refute 1485 => 1483. Sketch of proof: suppose we have a "base graph" G0=(V0,E0)G_0 = (V_0,E_0) that obeys Axioms 1, 2', anti-1483. (It can be finite or infinite.) Then we can build a directed graph G=(V,E)G = (V,E) obeying Axioms 1, 2, anti-1483 by the following procedure.

  • Take the vertex set VV to be V0×NV_0 \times {\mathbb N}. Vertices (a,n)(a,n) will be considered low if aa was low in V0V_0, and similarly for high.
  • As the initial seed, take the modified $5$-cycle (a,0)(b,0)(c,0)(d,0)(e,0)(a,0) \to (b,0) \to (c,0) \to (d,0) \to (e,0), (a,0)(e,0)(a,0) \to (e,0) where a,b,c,d,ea,b,c,d,e were the Axiom anti-1483 counterexample for the base graph G0G_0. (One may need to add a few more edges to ensure that every edge here is part of a good path, as per Claims 2, 3 of my initial axiom set.)
  • Recursively, we locate two vertices x=(a,n),y=(b,m)x = (a,n),y = (b,m) in GG such that one cannot currently reach yy from xx by a path of length two that is not lowlowlowlow \to low \to low, and add two edges xzyx \to z \to y to GG to join them, for some z=(c,l)z = (c,l) that has not previously been used, and is such that abca \to b \to c are valid edges in the base graph. One can check that this preserves Axioms 1, 2', anti-1483.
  • This procedure concludes with a graph obeying Axioms 1,2,anti-1483. We then declare all the length two paths that are not lowlowlowlow \to low \to low to be good paths. This then obeys the axioms Claim 1-4 for a weak central groupoid in directed graph form, and anti-1483 can be checked to refute 1483.

Unfortunately, this strategy cannot refute 161. For that, one needs a pattern of the form xyzx \to y \leftrightarrow z where xyzx \to y \to z is bad, but xzxx \to z \to x is good. But in this model, the bad paths are the lowlowlowlow \to low \to low paths, and the two claims are incompatible.

It's possible that one could generalize this strategy by finding a different rule for determining badness than "bad paths are the ones which are "low -> low -> low". My general approach here is to try to find a way to ensure the tricky 5-cycle axiom (Claim 4) that does not directly involve the concept of "good path", but instead only involves classifying (or coloring) the vertices into a small number of categories (colors) such as "low" and "high". Axiom 2 is the simplest attempt to do so, but perhaps other axioms are possible, and then the greedy algorithm on the extension allows any such axiom to be relaxed in that pairs of points can be connected by more than one good path.

The fact though that Axioms 1, 2', anti-1483 have no small countereexamples is concerning though. I wonder in ATP can be used to show that no such graph exists? I guess Axiom 1 may be more restrictive than I first thought.

Jose Brox (Oct 25 2024 at 14:35):

Zoltan A. Kocsis (Z.A.K.) ha dicho:

Neat! Just to have it on record, what is the number of the equation which the constants c1...c3 contradict? The strong c.g. law, I presume, but I don't yet know its number by heart.

Yes, that's the one, with number 168

Zoltan A. Kocsis (Z.A.K.) ha dicho:

This means that some of the magmas among the 180 may still be isomorphic as weak central groupoids, but with non-isomorphic "choice of counterexamples".

While isofiltering I chose the "ignore constants" option, isn't this preventing what you are warning about?

Terence Tao (Oct 25 2024 at 14:50):

So basically the effect of the greedy algorithm construction outlined above is to be able to relax the requirement of "exactly one good path" between any pair to "at least one good path". So, for instance, to refute 1485 -> 1483, it suffices to find a directed graph with a notion of "good" path that obeys the following axioms:

Claim 1' Any two points are connected by at least one good path of length two.
Claim 2 Any directed edge in the graph is the first component of at least one good path.
Claim 3 Any directed edge in the graph is the second component of at least one good path.
Claim 4: given a 5-cycle a → b → c → d → e → a, if the paths b → c → d and d → e → a are good, then the path c → d → e is also good.
Claim anti-1483: There is a modified 5-cycle a → b → c → d → e, a → e, in which the paths b → c → d and d → e → a are good, but the path c → d → e is bad.

[Note: Claim 2 and Claim 3 can be dropped here, so long as we require that the edge a → e in Claim anti-1483 is the initial or final edge of a good path.]

Similarly to refute 161 166, replace Claim anti-1483 by
Claim anti-161166: There is a pattern a → b ↔ c where b → c → b is good but a → b → c is bad.

I was hoping that the additional flexibility to have multiple good paths would create new interesting base graphs to work with, but this doesn't seem to be as likely as I hoped. It does seem that one can generate a large number of new infinite weak central groupoids this way, just not the ones that would refute 1483 or 161 166...

Pace Nielsen (Oct 25 2024 at 15:06):

After running Vampire for 600 CPU s on each unsolved implication, using the translation-invariant ansatz, it shows that the method is blocked for showing 1485 !=> 151, 3456, 3862. It is apparently not blocked for the other implications.

Michael Bucko (Oct 25 2024 at 15:13):

Treating these claims a bit like ml features (well, there're different in a way since I then give them to an unrelated solver). The search seems faster (I guess it should be, but maybe I'm wrong).

So far graphs of size 9 not found.

The addtions (perhaps there's bugs):

        for a, b, c, d, e in itertools.product(elements, repeat=5):
            forbidden_cycle = z3.And(edge[a][b], edge[b][c], edge[c][d], edge[d][e], edge[a][e])
            forbidden_path = z3.And(is_good[b], is_good[d], z3.Not(is_good[c]))
            s.add(z3.Implies(forbidden_cycle, forbidden_path))
        for a, b, c, d, e in itertools.product(elements, repeat=5):
            cond1 = z3.And(edge[a][b], edge[b][c], edge[c][d], edge[d][e], edge[e][a])
            cond2 = z3.And(is_good[b], is_good[d])
            cycle_condition = z3.Implies(z3.And(cond1, cond2), is_good[c])
            s.add(cycle_condition)
        for x, y in itertools.product(elements, repeat=2):
            good_path_exists = []
            for z in elements:
                path = z3.And(edge[x][z], edge[z][y])
                good_path = z3.Or(is_good[x], is_good[z], is_good[y])
                good_path_exists.append(z3.And(path, good_path))
            s.add(z3.Or(good_path_exists))

and so on. I guess such things should just land in the z3-feature-store.

Andy Jiang (Oct 25 2024 at 16:03):

Terence Tao said:

Sorry, I just realized that in order for these examples to have a shot at contradicting 1483, I need one more axiom:

  • Axiom anti-1483. There is at least one modified 5-cycle abcdea \to b \to c \to d \to e, eae \to a (i.e. a 5-cycle with the final edge reversed in orientation) such that b,c,db,c,d are low and a,ea,e are high.

(I haven't thought carefully about what the corresponding axiom would need to be to contradict 166, but I'll think about that tomorrow.) Unfortunately, the graph you provided doesn't obey this axiom. Would it be possible to conduct another search to see if there is a graph that obeys Axioms 1, 2', and anti-1483?

output(1).png
does this work?

Terence Tao (Oct 25 2024 at 16:14):

@Andy Jiang Would you by any chance have this graph in an adjacency matrix form? Oh wait, let me see if GPT can do this automatically for me...

Terence Tao (Oct 25 2024 at 16:17):

Heh, it missed the self-loop at 4, and maybe some others, but I can use this as a starting point to input in my code to check the graph.

adj_matrix = [
    [0, 1, 1, 0, 0, 0, 0],  # Node 0
    [0, 0, 1, 0, 1, 1, 1],  # Node 1
    [1, 0, 0, 0, 0, 0, 0],  # Node 2
    [0, 0, 0, 1, 0, 0, 0],  # Node 3 (self-loop)
    [0, 0, 0, 0, 0, 1, 1],  # Node 4
    [0, 0, 0, 1, 0, 1, 0],  # Node 5 (self-loop)
    [0, 1, 0, 0, 0, 0, 0],  # Node 6
]

Terence Tao (Oct 25 2024 at 16:25):

Actually, GPT made quite a few mistakes. It's amazing that it could get anywhere close at all just from the image, but still it shows these tools are just barely shy of being truly useful mathematical assistants right now. Here is the manually corrected matrix:

[
        [0, 1, 0, 1, 0, 0, 1],  # Node 0
        [0, 0, 1, 0, 0, 1, 1],  # Node 1
        [1, 0, 0, 0, 1, 0, 1],  # Node 2
        [1, 0, 0, 1, 1, 0, 0],  # Node 3 (self-loop)
        [0, 0, 1, 0, 1, 1, 0],  # Node 4 (self-loop)
        [0, 1, 0, 1, 0, 1, 0],  # Node 5 (self-loop)
        [1, 1, 1, 0, 0, 0, 0],  # Node 6
]

(For fairness: I also made an error in my manual transcription.) Now off to test it in Python.

Terence Tao (Oct 25 2024 at 16:30):

Hey, it checks out! The anti-1483 cycle is at 3 → 0 → 1 → 2 → 4, 3 → 4. I think this clinches the anti-implication 1485 -> 1483. I will write up a blueprint argument now.

Zoltan A. Kocsis (Z.A.K.) (Oct 25 2024 at 16:54):

Jose Brox said:

While isofiltering I chose the "ignore constants" option, isn't this preventing what you are warning about?

Yes, I think it should! All good then!

Terence Tao (Oct 25 2024 at 18:15):

Starting on the blueprint now. To refute the other 18 anti-implications from 1485, one would need to construct directed graphs obeying Claims 1', 2, 3, 4, and another anti-axiom. The situation is captured by the Graphiti: lower blue hypotheses are easier to refute, so the "easiest" refutations are 1483 and 166, and I have already listed the anti-166 axiom above. One of the hardest is 3456 x ◇ x = x ◇ ((x ◇ x) ◇ x) (or its dual, 3862, though refuting one automatically refutes the other by duality). This turns out to be a collapsed version of anti-1483:

Claim anti-3456: There is a modified 5-cycle a → b → c → d → ea → e, in which the paths a → b → c and c → d → e are good, but the path b → c → d is bad, with the additional constraints that b = e and that b → c → b is also a good path (in particular c → b also has to be in the graph).

There are several intermediate anti-implications between 1483 and 3456 that one could also encode in graph theoretic language, but it is somewhat tedious to do so. But in any case finding more graphs tailored to these anti-claims would give us more refutations than the single refutation 1485 -> 1483 that we seem to have right now.

EDIT: intermediate in difficulty between anti-1483 and anti-3456 is anti-1481, which only adds one additional condition rather than two:

Claim anti-1481: There is a modified 5-cycle a → b → c → d → ea → e, in which the paths a → b → c and c → d → e are good, but the path b → c → d is bad, with the additional constraint that b = e.

Terence Tao (Oct 25 2024 at 19:02):

A proof of the anti-implication 1485 !-> 1483 will land in the blueprint soon, equational#735 and I have opened equational#737 to formalize it.

Hopefully a similar method can handle the other 18 anti-implications! (They come in dual pairs, so we have at most 9 tasks here, and also some anti-implications are stronger than others, so we may not need to study all 9 separately.)

Observation: the proof strategy here is a combination of the "extend a base model" strategy and the "greedy" strategy. The new observation is due to the freedoms in the greedy strategy when working with an infinite extension, the base model can obey a more relaxed set of axioms than the original axiom set, and so the inability to find useful finite models to the original axiom 1485 is not as much of an obstruction now as it was before.(Roughly speaking, the relaxation is that we drop the requirement that the magma operation be a function: it can now be "multi-valued" in some sense, as encoded by the "is a good path" ternary relation.)

Mario Carneiro (Oct 25 2024 at 20:52):

@Terence Tao Can you elaborate a bit more on "Lemma 11.4 (Reversing the claims)"? The proof is just "reverse the claims" which is a bit <.<

Mario Carneiro (Oct 25 2024 at 20:53):

In particular, I don't see why Claim 4 is not included there. This would seem to imply that claim 4 can be derived just from claims 1-3 but I don't see any way to do this

Mario Carneiro (Oct 25 2024 at 21:35):

Also Theorem 11.6 (1485 does not imply 1483) does not make sense to me, there is an adjacency matrix for a graph on {0,...,7} but it's only 7*7 rather than 8*8... And none of the edges are marked good or bad so I don't know how to evaluate whether it satisfies the rules

Terence Tao (Oct 25 2024 at 22:01):

Thanks! I have fixed these in equational#739. To summarize:

  • I indeed had omitted Claim 4 in the hypothesis of Lemma 11.4, this is now fixed and a more verbose proof supplied.
  • The carrier is indeed {0,...,6} rather than {0,...,7}. The good/bad rule is that the paths contained in {0,1,2} are bad, all others are good. I've updated the text to indicate this.

Mario Carneiro (Oct 25 2024 at 22:05):

by the way, the graph has 3-fold symmetry, I wonder whether the numbering should reflect that? (Specifically, (012)(354) is an order 3 automorphism of the graph.)

Terence Tao (Oct 25 2024 at 23:02):

I'm hoping that other participants may discover more graph examples, and get some more insight as to where they are coming from, so that we don't have to rely on computer assistance to understand the proof (though of course for Lean formalization purposes just supplying a single computer-generated example is perfectly fine). The symmetry is a nice observation, it perhaps hints that we should look at symmetric examples to try to refute other implications as well.

Terence Tao (Oct 25 2024 at 23:26):

With the swapping of 4 and 5 (and also moving the special vertex 6 out of the way) I start seeing more of the structure: one has a triangular prism with the vertical edges bidirectional, the upper triangle oriented in the opposite direction to the lower triangle (and with loops on all the upper vertices), and finally the lower triangle is all conencted to the final vertex 6 by bidirectional edges. It is then possible to visually verify the axioms by "diagram chasing".

One thing this diagram chasing suggests to me is that adding a bunch of bidirectional edges or loops isn't such a bad idea once one has relaxed from Claim 1 to Claim 1'.

Mario Carneiro (Oct 25 2024 at 23:40):

I used the symmetry to construct a human proof of the theorem

Kevin M (Oct 26 2024 at 01:18):

How much freedom do you have in extending that size 7 model?
For example if you choose the 'good' paths such that for all x, (x ◇ x) != 6, then you can also rule out the implication eqn1485 -> eqn166, because eqn166 requires that (x ◇ x) is surjective.

Terence Tao (Oct 26 2024 at 01:39):

I think by adding an ad hoc rule that whenever selecting a path to connect x to itself, one avoids passing through a 6-colored vertex if possible, one can do this. (One has to check that each step of the greedy algorithm doesn't "accidentally" create a bidirectional edge from an existing vertex to a 6-class vertex, but this looks true since at each stage one is hooking up a previously unused vertex.) This should contradict not only 166, but also 151!

I could update the blueprint for this but I don't want to cause further merge conflicts with equational#738 , so hopefully the above sketch of argument is clear enough for now.

Perhaps a similar strategy could be used to refute other implications involving squares? Resolving 3456 for instance would clear out a lot of the implications thanks to duality.

Mario Carneiro (Oct 26 2024 at 02:14):

don't worry about conflicts with me, I'm sure I did a subset of what you did anyway

Mario Carneiro (Oct 26 2024 at 02:16):

I pushed the construction of the base magma and the proof that it is a "relaxed" counterexample. I have not yet done the extension process, I'm hoping that someone will have something good to say here because I've seen some of the proofs done for these kind of enumerative constructions and I'm sure we can do way better than that at a higher level of abstraction

Terence Tao (Oct 26 2024 at 02:55):

Wait, it turns out that one cannot prevent elements in the "6" color class from being squares after all. At some point one has to connect such an element (6,n)(6,n) to itself by a good path through some new element xx - and the moment one does that, xx will then square to (6,n)(6,n) since we have the path x(6,n)xx \to (6,n) \to x which will automatically be classified as good.

So we have not yet contradicted 166 or 151. Still, the idea to be smarter with the greedy algorithm to add more properties to the weak central groupoid being constructed is a promising one.

If we had a base relaxed weak central groupoid with the property that there was a pair x,yx,y with xyxx \to y \to x a good path but yxyy \to x \to y a bad path, this would contradict 151 (and hence 161)...

Andy Jiang (Oct 26 2024 at 03:45):

Terence Tao said:

So basically the effect of the greedy algorithm construction outlined above is to be able to relax the requirement of "exactly one good path" between any pair to "at least one good path". So, for instance, to refute 1485 -> 1483, it suffices to find a directed graph with a notion of "good" path that obeys the following axioms:

Claim 1' Any two points are connected by at least one good path of length two.
Claim 2 Any directed edge in the graph is the first component of at least one good path.
Claim 3 Any directed edge in the graph is the second component of at least one good path.
Claim 4: given a 5-cycle a → b → c → d → e → a, if the paths b → c → d and d → e → a are good, then the path c → d → e is also good.
Claim anti-1483: There is a modified 5-cycle a → b → c → d → e, a → e, in which the paths b → c → d and d → e → a are good, but the path c → d → e is bad.

[Note: Claim 2 and Claim 3 can be dropped here, so long as we require that the edge a → e in Claim anti-1483 is the initial or final edge of a good path.]

Similarly to refute 161, replace Claim anti-1483 by
Claim anti-161: There is a pattern a → b ↔ c where b → c → b is good but a → b → c is bad.

I was hoping that the additional flexibility to have multiple good paths would create new interesting base graphs to work with, but this doesn't seem to be as likely as I hoped. It does seem that one can generate a large number of new infinite weak central groupoids this way, just not the ones that would refute 1483 or 161...

I can try to look for some of these anti-161 relaxed weak central groupoids. Does a and b have to be different in claim anti-161?

Andy Jiang (Oct 26 2024 at 04:05):

if not (a and b can be the same), can someone check the following example? I don't particularly trust my coding skills (even with AI help)
a=0,b=0,c=1
output(2).png
Adjacency Matrix: [[1 1 0 0 1] [1 1 0 1 0] [0 0 1 1 1] [1 0 1 0 0] [0 1 1 0 0]]
Bad paths: (0,0,0),(0,0,1),(0,1,1),(1,0,0),(1,1,0),(1,1,1),(2,2,2)
Edges List:
(0, 0)
(0, 1)
(0, 4)
(1, 0)
(1, 1)
(1, 3)
(2, 2)
(2, 3)
(2, 4)
(3, 0)
(3, 2)
(4, 1)
(4, 2)

Good Paths List:
(0, 0, 4)
(0, 1, 0)
(0, 1, 3)
(0, 4, 1)
(0, 4, 2)
(1, 0, 1)
(1, 0, 4)
(1, 1, 3)
(1, 3, 0)
(1, 3, 2)
(2, 2, 3)
(2, 2, 4)
(2, 3, 0)
(2, 3, 2)
(2, 4, 1)
(2, 4, 2)
(3, 0, 0)
(3, 0, 1)
(3, 0, 4)
(3, 2, 2)
(3, 2, 3)
(3, 2, 4)
(4, 1, 0)
(4, 1, 1)
(4, 1, 3)
(4, 2, 2)
(4, 2, 3)
(4, 2, 4)

Matthew Bolan (Oct 26 2024 at 04:53):

Threw together a quick testing script, which approves of the example.

Matthew Bolan (Oct 26 2024 at 04:55):

It also suggests that the bad path (2,2,2) is not needed (that is, it can be good)

Kevin M (Oct 26 2024 at 05:00):

[EDIT: was riddled with errors, let me try again below]

Andy Jiang (Oct 26 2024 at 05:15):

Terence Tao said:

Wait, it turns out that one cannot prevent elements in the "6" color class from being squares after all. At some point one has to connect such an element (6,n)(6,n) to itself by a good path through some new element xx - and the moment one does that, xx will then square to (6,n)(6,n) since we have the path x(6,n)xx \to (6,n) \to x which will automatically be classified as good.

So we have not yet contradicted 166 or 151. Still, the idea to be smarter with the greedy algorithm to add more properties to the weak central groupoid being constructed is a promising one.

If we had a base relaxed weak central groupoid with the property that there was a pair x,yx,y with xyxx \to y \to x a good path but yxyy \to x \to y a bad path, this would contradict 151 (and hence 161)...

By the way I also looked for the x->y->x good and y->x->y bad combo + relaxed weak central groupoid axioms and didn't find any with <=9 vertices

Matthew Bolan (Oct 26 2024 at 05:17):

I wanted to check Andy's graph / maybe do an overnight run for the other axioms, but I ran into an issue programming the checks for them. I'm guessing

Claim anti-1481: There is a modified 5-cycle a → b → c → d → e, a → e, in which the paths b → c → d and d → e → a are good, but the path c → d → e is bad, with the additional constraint that b = e.

should really ask for paths a → b → c and c → d → e to be good, but b → c → d be bad? This would agree with the example for anti-1483, but if this is the correct formulation then I need to know if it is still vertices b and e which should be identified. I don't think it is possible to satisfy the original phrasing, the existence of the path d → e → a requires an edge e → a, so we in fact have a 5-cycle a → b → c → d → e → a and claim 4 tells us c → d → e is good.

Kevin M (Oct 26 2024 at 05:42):

I checked with prover9, and the subset of elements that are 'squares' satisfy eqn151.
(x = (y*y)) -> ( x = (x*x)*(x*x) ).
Therefore to violate eqn151, we need at least one element that is not a square.

Terence Tao (Oct 26 2024 at 05:50):

@Matthew Bolan Oops, you are right, I had shifted things by one.

Perhaps its best if I explain where these conditions come from. The basic point is that the assertion z = x ◇ y is encoded graph theoretically as "x → z → yis a good path".

Let's see how this works for a failure of 1481: x ≠ (y ◇ x) ◇ (x ◇ (y ◇ x)). Writing z = y ◇ x, w = x ◇ z, the assertions are then that y → z → x is good, x → w → z is good, but z → x → w is not good. So here we have a modified 5-cycle y → z → x → w → z, y → z in which the first and last paths in the y → z → x → w → z chain are good, but the middle one is bad. This is the corrected version of anti-1481.

Matthew Bolan (Oct 26 2024 at 06:05):

With that correction I find that Andy's graph is a new counterexample to 1483, with the 5-cycle in question being 4 1 1 0 1

Terence Tao (Oct 26 2024 at 06:07):

I guess I can do a couple more while on this topic. Let's do anti-2087: x ≠ ((y ◇ x) ◇ x) ◇ (x ◇ x). Writing z = y ◇ x, w = z ◇ x, u = x ◇ x, we have z → x (we in fact have y → z → x as a good path, but y doesn't appear anywhere else, so we can just remove it here), z → w → x and x → u → x good, but w → x → u bad. So now we have a modified 5-cycle z → w → x → u → x, z → x with the outer two paths good and the inner path bad.

Now for anti-2124 x ≠ ((y ◇ y) ◇ x) ◇ (x ◇ x). Write z = y ◇ y, w = z ◇ x, u = x ◇ x, then we have y → z → y, z → w → x, x → u → x as good paths, but w → x → u as a bad path. So the geometry here is a little different, it's a 4-path z → w → x → u → x (with the other paths good and inner path bad) that touches a good path y → z → y at one end.

And to review anti-166 (which I had mistakenly called 161 previously): x ≠ (y ◇ x) ◇ (x ◇ x). Writing z = y ◇ x, w = x ◇ x, we have y → z → x and x → w → x good, but z → x → w bad, so now we have a 4-path y → z → x → w → x in which the outer paths are good but the inner path is bad.

Matthew Bolan (Oct 26 2024 at 06:08):

Oh, and that means andy's graph looks like it satisfies at least anti-1481 as well

Matthew Bolan (Oct 26 2024 at 06:09):

and 2087?

Matthew Bolan (Oct 26 2024 at 06:18):

Unfortunately it doesn't satisfy anti 3456, as for (a,b,c,d,e)= (4 1 1 0 1) (b,c,b)=(1,1,1) is not a good path (and my program finds no other viable modified 5-cycles).

Terence Tao (Oct 26 2024 at 06:22):

Perhaps the reversal of the graph (take the transpose of the adjacency matrix) does better (or equivalently, the existing graph obeys anti-3862 instead of anti-3456).

Anyway, I'm calling it a day here. I'll run my own code to independently check this graph tomorrow.

Andy Jiang (Oct 26 2024 at 06:24):

Matthew Bolan said:

Unfortunately it doesn't satisfy anti 3456, as for (a,b,c,d,e)= (4 1 1 0 1) (b,c,b)=(1,1,1) is not a good path (and my program finds no other viable modified 5-cycles).

If you can help me with translating this condition (anti-3456) I can run the SAT solver on this one as well to look for these relaxed weak central groupoids. (Or if I have time later I can try the translation too)

Matthew Bolan (Oct 26 2024 at 06:26):

Terry already wrote this one out for us above

Claim anti-3456: There is a modified 5-cycle a → b → c → d → ea → e, in which the paths a → b → c and c → d → e are good, but the path b → c → d is bad, with the additional constraints that b = e and that b → c → b is also a good path (in particular c → b also has to be in the graph).

Matthew Bolan (Oct 26 2024 at 06:30):

Also, I think the reversal of the graph does not work for 3456 either, It still supposedly satisfies many of the other desirable properties though.

Matthew Bolan (Oct 26 2024 at 07:00):

Andy's graph also seems to work for 2124. The 4-path is again this 4 1 1 0 1, and it meets the good path 2 4 2.

Andy Jiang (Oct 26 2024 at 07:09):

Matthew Bolan said:

Terry already wrote this one out for us above

Claim anti-3456: There is a modified 5-cycle a → b → c → d → ea → e, in which the paths a → b → c and c → d → e are good, but the path b → c → d is bad, with the additional constraints that b = e and that b → c → b is also a good path (in particular c → b also has to be in the graph).

OK I didn't find any (relaxed weak central groupoids which are anti-3456) with <=9 vertices (tentative on code being correct). Can run it for longer for 10 vertices later if needed

Matthew Bolan (Oct 26 2024 at 07:46):

To summarize before I go to bed, my current understanding is that we win if we find any of the following (though this is not an exhaustive list):

anti-166 : A 4-path y → z → x → w → x in which the outer paths are good but the inner path is bad.
anti-1481: a modified 5-cycle y → z → x → w → zy → z in which the first and last paths in the y → z → x → w → z chain are good, but the middle one is bad.
anti-2087: A modified 5-cycle z → w → x → u → x, z → x with the outer two paths good and the inner path bad.
anti-2124: A 4-path z → w → x → u → x (with the other paths good and inner path bad) that touches a good path y → z → y at one end.
anti-3456: A modified 5-cycle a → b → c → d → ea → e, in which the paths a → b → c and c → d → e are good, but the path b → c → d is bad, with the additional constraints that b = e and that b → c → b is also a good path (in particular c → b also has to be in the graph).

Andy's graph https://leanprover.zulipchat.com/user_uploads/3121/Puzc51dyFLYm6k9_o4ju2piB/output2.png
with adjacency matrix

[[1, 1, 0, 0, 1],
 [1, 1, 0, 1, 0],
 [0, 0, 1, 1, 1],
 [1, 0, 1, 0, 0],
 [0, 1, 1, 0, 0]]

and bad paths

{(0, 0, 0), (0, 0, 1), (0, 1, 1), (1, 0, 0), (1, 1, 0), (1, 1, 1), (2, 2, 2)}

contains examples of 4 of these, namely 166, 1481, 2087, 2124.

The key path is 4 1 1 0 1, which on its own rules out 166, 1481, and 2087, which all basically ask for a modified 5-cycle with various equalities satisfied. 2124 is also ruled out, due to the 2 4 2 path being good. There is no luck with anti-3456, an example is not provided by either this graph or its reversal. I have not checked any others, but I wouldn't be surprised if this works for a few more. I can do the job of translating the axioms and checking them against this graph in the morning if no one else does.

The path (2,2,2) can be made good without changing any of the above.

I think the graph (say the version where (2,2,2) is good) can be checked humanly. Claims 1',2,3 are pretty easy. For claim 4, observe that any 5-cycle containing 3 consecutive vertices in {0,1}\{0,1\} must contain a 4th vertex in {0,1}\{0,1\}, after which it amounts to checking there is no 5-cycle containing any of 1011,1101, 0010,0100 and a vertex not in {0,1}\{0,1\}.

Jose Brox (Oct 26 2024 at 08:43):

(deleted)

Matthew Bolan (Oct 26 2024 at 08:55):

Now that I understand the process needed to convert the laws to this graph language (and since I was having trouble sleeping) I decided to go through the rest of the laws and check them against the graph. I think it only misses 151, 3456, and 3456's dual 3862.

Matthew Bolan (Oct 26 2024 at 09:16):

All except 3511 and 3457 are by essentially this 4 1 1 0 1 path (and the dual laws by 1 0 1 1 3), together with the fact 4 meets the middle of the good paths 2 4 1 and 2 4 2.

For 3457, x ◇ x = x ◇ ((x ◇ x) ◇ y) I let z = x ◇ x, w = z ◇ y, so the statement is that x → z → x, z → w → y are good paths, but x → z → w is not. For this take x=1,z=w=0,y=4.

For 3511, x ◇ y = x ◇ ((x ◇ y) ◇ x), I let z = x ◇ y, w = z ◇ x, so the statement is that x → z → y, z → w → x are good paths, but x → z → w is not. For this take x=z=1,w=0,y=3.

Matthew Bolan (Oct 26 2024 at 09:48):

I also noticed that Terry's formulation of 3456 can be simplified slightly (though it makes it look less like the other laws). There we ask for a modified 5-cycle a → b → c → d → e, a → e, in which the paths a → b → c and c → d → e are good, but the path b → c → d is bad, with the additional constraints that b = e and that b → c → b is also a good path (in particular c → b also has to be in the graph).

I think that if (a,b,c,d,e)=(a,b,c,d,b) satisfies this, then so does (d,b,c,d,b) . Indeed, c → d → e=b → c → b → c is then a 5-cycle with b → c → b and c → d → b both good, so d → b → c is good automatically, and so d can play the role of a.

Andy Jiang (Oct 26 2024 at 12:24):

Matthew Bolan said:

I also noticed that Terry's formulation of 3456 can be simplified slightly (though it makes it look less like the other laws). There we ask for a modified 5-cycle a → b → c → d → e, a → e, in which the paths a → b → c and c → d → e are good, but the path b → c → d is bad, with the additional constraints that b = e and that b → c → b is also a good path (in particular c → b also has to be in the graph).

I think that if (a,b,c,d,e)=(a,b,c,d,b) satisfies this, then so does (d,b,c,d,b) . Indeed, c → d → e=b → c → b → c is then a 5-cycle with b → c → b and c → d → b both good, so d → b → c is good automatically, and so d can play the role of a.

your reformulation looks kind of similar to the formulation of the relaxed model against 151 right? it's a condition on a 3-cycle instead of a 2-cycle for 151

Matthew Bolan (Oct 26 2024 at 15:19):

Yes, I agree they are similar. In fact I believe any example for anti-3456 will also be an example for anti-151. This is because the graph for 3456 will then contain a 5-cycle b → c → b → c → d → b where b → c → b and c → d → b are good and b → c → d is bad, so we must have that c → b → c is bad, otherwise axiom 4 would tell us b → c → d is good. Therefore such a thing has b → c → b good and c → b → c bad, which is the condition to be anti-151.

As a sanity check, on the magma side of things this is saying that given a 1485 magma containing an element x violating 3456 [so x ◇ x ≠ x ◇ ((x ◇ x) ◇ x)] we must have that x violates 151 as well. Suppose it does not, so x = (x ◇ x) ◇ (x ◇ x). Then we get the contradiction (x ◇ x) = ((x ◇ x) ◇ (x ◇ x))((x ◇ x) ◇ ((x ◇ x)◇(x ◇ x))) = (x) ◇ ((x ◇ x) ◇ x), where the first equality is just 1485 with every variable set to (x ◇ x).

Terence Tao (Oct 26 2024 at 15:27):

I can confirm @Matthew Bolan 's calculations using @Andy Jiang 's graph. This graph indeed refutes 3457, 2124, 2087, 3511, which together with duality refutes 16 of the 19 recently open implications (including the previously refuted 1483), leaving only the three anti-implications to the one-variable laws 151, 3456, and 3862 remaining. Since the latter two are dual, one just needs to understand two anti-implications now: 1485 -> 151 and 1485 -> 3456. (Edit: as Matthew just observed, it suffices to refute the former, as 1485 and 3456 imply 151.)

One way to interpret relaxed weak central groupoids is with the "relaxed" multiplication table, in which the x,y entry is listed with all the possible z for which x → z → y is a good path, and reflects all the possible color choices for x ◇ y when performing a greedy extension. In the case of @Andy Jiang 's graph, the relaxed multiplication table is

1    4    4    1    0
3    0    3    1    0
3    4    3,4  2    2
0    0    2    2    0,2
1    1    2    1,2  2

Thus for instance 2 ◇ 2 can be chosen to equal either 3 or 4 at various points in the greedy construction (although sometimes the choice is forced due to previous stages of the construction), but 1 ◇ 0 must always be of type 3. As noted by @Matthew Bolan, one also has the option to permit 2 as a value for 2 ◇ 2 if desired without breaking anything, though it also doesn't seem to add any new refutation power either.

Using this relaxed multiplication table, one can check the violations directly (this is just rewording @Matthew Bolan 's analysis in a different language:

  • Violation of 3457 x ◇ x = x ◇ ((x ◇ x) ◇ y). Take x=1, y=4. 1 ◇ 1 = 0 ≠ 3 = 1 ◇ ( (1 ◇ 1) ◇ 4 ).
  • Violation of 2087: x = ((y ◇ x) ◇ x) ◇ (x ◇ x). Take x=1, y=2. 1 ≠ 3 = ((2 ◇ 1) ◇ 1) ◇ (1 ◇ 1).
  • Violation of 2124: x = ((y ◇ y) ◇ x) ◇ (x ◇ x). Take x=1, y=2, and choose y ◇ y = 4. Then 1 ≠ 3 = ((2 ◇ 2) ◇ 1) ◇ (1 ◇ 1).
  • Violation of 3511: x ◇ y = x ◇ ((x ◇ y) ◇ x). Take x=1, y=3. 1 ◇ 3 = 1 ≠ 3 = 1 ◇ ((1 ◇ 3) ◇ 1).

I'll add these conjectures now and modify the blueprint accordingly (sorry @Mario Carneiro !).

Matthew Bolan (Oct 26 2024 at 15:30):

151 and 1485 look close to implying 3456 to me as well, has this been checked for before?

Matthew Bolan (Oct 26 2024 at 15:45):

I guess when I said 151 there I was thinking about the negation of 151. My earlier argument shows a magma obeying 151 and 1485 obeys 3456. The new question is whether a magma obeying 1485 and 3456 obeys 151.

Matthew Bolan (Oct 26 2024 at 15:54):

Ah, yes. In 1485 x = (yx)(x(zy)) take y=x and z = (xx), to get x = (xx)(x((xx)x)) = (xx)(xx), where the last equality is 3456

Matthew Bolan (Oct 26 2024 at 15:55):

so I think it suffices to falsify only one of the remaining implications.

Terence Tao (Oct 26 2024 at 16:19):

I think I now have some intuition as to what we are doing here; the intuition can be stated without the directed graph language, but for the purpose of formalizing the proofs it seems to be better to keep the graph formalism.

We are constructing an infinite weak central groupoid in which the vertices are divided into five infinite classes: Class 0, Class 1, Class 2, Class 3, and Class 4. (Formally, we can take the carrier to be {0,1,2,3,4}×N\{0,1,2,3,4\} \times {\mathbb N}). The multiplication law will be required to be consistent with the 5 x 5 relaxed multiplication table. For instance, the product xyx \diamond y of a Class 1 object xx and a Class 0 object yy must be a Class 3 object, but the square xxx \diamond x of a Class 2 object xx can either be a Class 3 object or a Class 4 object. The relaxed multiplication table is a good table to use for building weak central groupoids because it has the following property: if one has already chosen values for yxy \diamond x and x(zy)x \diamond (z \diamond y) consistent with the above table, then the table will never prevent you from setting (yx)(x(zy))(y \diamond x) \diamond (x \diamond (z \diamond y)) equal to xx, so the relaxed table is never directly responsible for blocking you from filling out an actual multiplication table to obey 1485. (This is Claim 4, interpreted in magma operation language.) Furthermore, all the entries of the relaxed table are non-empty; there is always at least one choice of class available for any product xyx \diamond y. (This is Claim 1'.)

That said, one might be blocked for a different reason: if one has two objects a,ba, b that can be represented both in the form a=yxa = y \diamond x and b=x(zy)b = x \diamond (z \diamond y) and a=yxa = y' \diamond x' and b=x(zy)b = x' \diamond (z' \diamond y') with xxx \neq x', then one is now in trouble because 1485 would demand that aba \diamond b to be set to equal xx and also to xx'. So one needs a running hypothesis that there is at most one xx that can arise in this fashion. (This is basically Claim 1'' in the blueprint, interpreted back in magma operation language.)

With this in mind, one now runs a greedy algorithm: select a pair xx, yy of objects in which the product xyx \diamond y is not already forced to equal some given object by the above considerations, and then arbitrarily set it equal to any new object that is compatible with the relaxed multiplication table. There is one technical check, which is that this new addition does not generate a violation of the running hypothesis, but in the specific case of constructing 1485 magmas, this is happily the case as long as the new object had no previous appearance in the partially constructed multiplication table. (It is for the purposes of this consistency check that it is best to use the graph theory formalism.) Now one can iterate one's way to generate a complete (and infinite) weak central groupoid that is compatible with the relaxed multiplication table, which is then a good magma to use for refutations.

Terence Tao (Oct 26 2024 at 16:24):

Conjectures now in the github repository. The graphiti of known and open implications from 1485 is now much more civilized! And we are down to 24 open implications!

Terence Tao (Oct 26 2024 at 16:33):

We're going to need another digit of precision on the dashboard. We've achieved six nines!

Matthew Bolan (Oct 26 2024 at 18:17):

I think there is a freeish relaxed weak central groupoid that seems somewhat understandable. The carrier is a union over countable sets CxC_{x} (not necessarily disjoint), with xx some expression in a free magma. Then I want abCxyab \subseteq C_{xy} if aCx,bCya \in C_{x}, b \in C_{y}, together with C(xy)(x(zy))CxC_{(xy)(x(zy))} \subseteq C_x. I think this is easy to satisfy provided you leave yourself enough room in CxC_x to accomodate later products.

Terence Tao (Oct 26 2024 at 18:29):

Interesting. So what would the relaxed operation xyx \diamond y look like - would it be all the expressions zz for which CzCxyC_z \subseteq C_{xy}?

The fact that 1485 is known to not be confluent makes me slightly nervous that this description might not be as tractable as one might hope, but I can certainly imagine that by abstract nonsense there is a canonical relaxed free magma for any given law.

Terence Tao (Oct 26 2024 at 18:39):

Just to extract out what still needs to be done: up to known equivalences, the last remaining task is to refute 1485 -> 151. In the graph language, this is equivalent to finding elements 0,10, 1 in a weak central groupoid such that 0100 \to 1 \to 0 is good but 1011 \to 0 \to 1 is bad (of course this forces 010 \neq 1).

By the greedy construction, it suffices to find a relaxed weak central groupoid with two elements 00, 11 such that 0100 \to 1 \to 0 is good but 1011 \to 0 \to 1 is bad. Using R(x,y,z)R(x,y,z) for the relation that xzyx \to z \to y is a good path (think of this as a relaxation of z=xyz = x \diamond y), we are now asking for a model with two constants 0,10, 1, a binary relation xyx \to y, and a ternary relation RR obeying the following axioms:

  1. If R(x,y,z)R(x,y,z), then xzx \to z and zyz \to y.
  2. If xzx \to z, then there exists yy such that R(x,y,z)R(x,y,z).
  3. If zyz \to y, then there exists xx such that R(x,y,z)R(x,y,z).
  4. For each x,yx,y there exists zz with R(x,y,z)R(x,y,z).
  5. If abcdeaa \to b \to c \to d \to e \to a and R(a,c,b)R(a,c,b), R(c,e,d)R(c,e,d), then R(b,d,c)R(b,d,c).
  6. R(0,0,1)R(0,0,1) is true but R(1,1,0)R(1,1,0) is false.

Perhaps an ATP could resolve this directly?

Mario Carneiro (Oct 26 2024 at 18:42):

I'm looking at the new version of the blueprint proof here and I don't really like that the last line has changed formalisms from everything that came before it

Mario Carneiro (Oct 26 2024 at 18:43):

can this be presented as cycle conditions and good/bad paths? We just went to the trouble of defining this equivalence

Mario Carneiro (Oct 26 2024 at 18:43):

I don't see a huge benefit in working with "multimagmas" over graphs in this situation BTW

Terence Tao (Oct 26 2024 at 18:44):

Fair enough. I'll write an alternate proof using the graph notation now. I guess I was thinking ahead to generalizing to other equations such as 854, in which the relaxed magma formalism seems better suited.

Mario Carneiro (Oct 26 2024 at 18:45):

I suspect that when formalized the two objects are basically indistinguishable

Terence Tao (Oct 26 2024 at 18:46):

Oh I remember why I did this. There were four claims 3457, 2087, 2124, 3511 to refute (that are of course in the original language of magmas), and to do things the previous way I had to then set up corresponding graph theory axioms anti-3457, anti-2087, anti-2124, and anti-3511 and show that they were equivalent to their magma counterparts.

Mario Carneiro (Oct 26 2024 at 18:46):

When you are witnessing a counterexample in a multimagma, I think you need to make the choices for each subexpression?

Matthew Bolan (Oct 26 2024 at 18:47):

Terence Tao said:

Interesting. So what would the relaxed operation xyx \diamond y look like - would it be all the expressions zz for which CzCxyC_z \subseteq C_{xy}?

The fact that 1485 is known to not be confluent makes me slightly nervous that this description might not be as tractable as one might hope, but I can certainly imagine that by abstract nonsense there is a canonical relaxed free magma for any given law.

My hope was that I could tweak this description by using the wiggle room afforded by the fact the CxC_x s are countable to avoid making some element in C0C_0 a square. However, I am nervous about the object's existence now, allow me to check some details and write down an explicit law before speculating further.

Terence Tao (Oct 26 2024 at 18:48):

In principle yes, but actually one only needs to specify a small number of the subexpressions and the base relaxed weak central groupoid does the rest of the work to force a contradiction regardless of the precise values of each subexpression

Mario Carneiro (Oct 26 2024 at 18:48):

I know it's probably obvious, but for formalization it's easiest if someone just says things like x=4, x*y=4, etc

Mario Carneiro (Oct 26 2024 at 18:49):

maybe I'll just punt on these counterexamples for now

Terence Tao (Oct 26 2024 at 18:50):

One can rework the proof a bit to make more products explicit. The minimal thing I need - and only to refute 2124 - is to control a square: (2,0)(2,0)=(4,0)(2,0) \diamond (2,0) = (4,0). But even this could be discarded if needed, because the base relaxed weak central groupoid forces (2,0)(2,0)(2,0) \diamond (2,0) to lie in either {4}×N\{4\} \times {\mathbb N} or {3}×N\{3\} \times {\mathbb N}, and either one can be shown to lead to a contradiction.

Terence Tao (Oct 26 2024 at 18:51):

So actually I think every complete extension of the base relaxed weak central groupoid violates all four laws. So one could just extend the empty graph.

Mario Carneiro (Oct 26 2024 at 18:51):

So supposing that we wanted to prove everything about G0 and then not look at it again, what properties do we need to establish?

Terence Tao (Oct 26 2024 at 18:54):

I can work this out in purely graph theoretic terms. One sec.

Terence Tao (Oct 26 2024 at 19:02):

  • It is a relaxed weak central groupoid (Claims 1', 2, 3, 4)
  • Anti-3457: There exist x,y such that for any z,w with x → z → x and z → w → y both good, x → z → w is bad. (One can take x=1, y=4 in the specific example.)
  • Anti-2087: There exist x,y such that for any z,w,u with y → z → x, z → w → x, and `x → u → x good, w → x → u is bad. (One can take x=1, y=2 in the specific example.)
  • Anti-2124: There exists y such that for any z with y → z → y good, there exists x such that for any w,u with z → w → x and x → u → x good, w → x → uis bad. (We can take y=2, which forces z to equal 3 or 4. If z=4, we take x=1; if instead z=3, we take x=0.)
  • Anti-3511: There exists x,y such that for any z,w with x → z → y and z → w → x good, one has x → z → w bad. (One can take x=1, y=3.)

Mario Carneiro (Oct 26 2024 at 19:03):

So this is something that was confusing me about the switch to multimagmas: It is now a exists-forall statement instead of an exists

Mario Carneiro (Oct 26 2024 at 19:03):

are you sure this is needed?

Terence Tao (Oct 26 2024 at 19:05):

No it is overkill admittedly. One could get away with weaker claims, but then one has to tailor the seed magma for the greedy algorithm to each of the four anti-claims separately.

Mario Carneiro (Oct 26 2024 at 19:06):

I am not sure I follow. Isn't the claim that the 5-point G0 satisfies both versions of all four claims?

Terence Tao (Oct 26 2024 at 19:06):

With the above strong axioms, every extension of G0 will refute all four laws. With the weaker laws (basically in which all the "for all"s are replaced with "there exists", one can create four custom extensions to refute each of the four laws separately.

Terence Tao (Oct 26 2024 at 19:06):

It obeys both the strong version and the weak version.

Terence Tao (Oct 26 2024 at 19:07):

(The multiplication table is almost completely unambiguous for G0, so the two versions don't differ very much there.)

Mario Carneiro (Oct 26 2024 at 19:07):

So the extension process needs to know about these claims and avoid causing problems?

Terence Tao (Oct 26 2024 at 19:09):

Not directly. With the strong versions of the claims, one can just take any extension of G0 and the strong claims will then ensure that every extension will refute all four laws. With the weak version of say anti-3457, one just has to start with a small partial extension of G0 that already refutes 3457 (and the weak version will let you do this), extend that to a complete extension, and that extension will necessarily give a weak central groupoid that refutes 3457, but may possibly obey the other three laws.

Terence Tao (Oct 26 2024 at 19:11):

The weak version of anti-3457 is that there exist x,y,z,w with x → z → x and z → w → y both good, x → z → w is bad. Given such a x,y,z,w, one can construct a partial extension to G0 using the directed graph (x,0) → (z,1) → (x,0) → (w,2) → (y,3) for instance (I choose different numbers in the second coordinate here to avoid unwanted collisions) and then extending greedily.

Mario Carneiro (Oct 26 2024 at 19:12):

So you can't just extend from nothing?

Mario Carneiro (Oct 26 2024 at 19:12):

I'm wondering if in the general infrastructure for doing this style of proof we want to prefer the first form or the second

Terence Tao (Oct 26 2024 at 19:12):

If one only has the weak claims and not the strong claims, then I can't, because I can't guarantee that the products land where I want them to land. But with the strong claims, I can extend from nothing and refute all four of the laws without having to plan an initial non-empty seed for the extension.

Terence Tao (Oct 26 2024 at 19:13):

Hmm. Probably the weak form with a non-empty seed is the more flexible and natural approach. The approach in the blueprint is sort of a shortcut that allowed me to not think about the seed issue

Terence Tao (Oct 26 2024 at 19:14):

So if you like I can rewrite the blueprint to be oriented around the weak version of the claims for G0, then to refute the four laws I create four different seeds (as guided by the weak claims), extend each of them, and use them to refute the four laws separately.

Mario Carneiro (Oct 26 2024 at 19:16):

Did the second coordinate numbers above have to be chosen by a method other than "sufficiently fresh"?

Mario Carneiro (Oct 26 2024 at 19:16):

because isn't that already what the greedy extension method is doing?

Terence Tao (Oct 26 2024 at 19:16):

No, I just needed to make sure that the elements (x,0), (z,1), (w,2), (y,3) were distinct, and this was the quickest way

Mario Carneiro (Oct 26 2024 at 19:17):

how much of this is specific to the law being refuted?

Terence Tao (Oct 26 2024 at 19:17):

The seed has to have the property that any two points are connected by at most one good path. The fewer collisions I have, the easier it is to verify this property

Terence Tao (Oct 26 2024 at 19:17):

It's possible that (x,0), (z,0), (w,0), (y,0) still obeys this axiom even with collisions, but I didn't want to think about this

Mario Carneiro (Oct 26 2024 at 19:18):

By seed you don't mean the base multimagma

Terence Tao (Oct 26 2024 at 19:18):

No. The seed is a partially defined weak central groupoid living in G0 x Nat and will be enlarged into a complete weak central groupoid G with that carrier (equipped with a homomorphism to G0).

Mario Carneiro (Oct 26 2024 at 19:19):

Every extension will have the "at most one" property (it's an invariant of the greedy construction), but we need a nonempty graph to start with to ensure at least one failing of the law in question is preserved

Mario Carneiro (Oct 26 2024 at 19:20):

and adding points to the graph means that claim 1'' has to be checked explicitly

Terence Tao (Oct 26 2024 at 19:20):

One also needs the seed to obey the "at most one" property, but yes that is correct.

Terence Tao (Oct 26 2024 at 19:21):

In past experience, the hard part was always finding an invariant ruleset for the greedy construction that contained the original hypothesis law (in this case, the weak central groupoid law 1485). Once one had it, finding a seed that obeyed the ruleset as well as holding the desired counterexample was just a finite search, usually quite easy.

Terence Tao (Oct 26 2024 at 19:24):

OK I will go ahead and refactor the blueprint. [update: equational#743]

Matthew Bolan (Oct 26 2024 at 19:32):

In my current understanding, the relaxed axioms are that we have a relaxed product G×GP(G){}G \times G \to \mathcal{P}(G)\setminus\{\varnothing\},
such that the set (yx)(x(zy))(yx)(x(zy)) contains xx for all x,y,zGx,y,z \in G . Can someone confirm?

Matthew Bolan (Oct 26 2024 at 19:37):

Ah, I think it is actually that the set abab contains xx for all aa in yxyx and all bb in x(yz)x(yz), in turn for all x,y,zGx,y,z \in G.

Terence Tao (Oct 26 2024 at 19:39):

It's the latter, yes. The intuition is that the choices you make for the subproducts can't automatically make you lose the game for the later product.

Terence Tao (Oct 26 2024 at 19:41):

Roughly speaking, in order for the greedy algorithm to work, one needs to work in a "no regrets" environment in which any foolish choices one makes earlier in life don't end up causing you to end up in a no-win situation later in life.

Mario Carneiro (Oct 26 2024 at 19:43):

On that note, Kyle made an interesting comment on the other stream which suggests that in some situations actually "no regrets" is not a requirement? I have not properly grokked that proof

Terence Tao (Oct 26 2024 at 19:52):

I think one can break up the greedy algorithm into two components:

  1. Compactness theorem: if for any finite set of tasks, one can find a state that obeys those tasks, then one can find a state that obeys all tasks simultaneously. [This theorem does not require any compatibility (or "no regret" guarantee) between states. This requires the state space to be compact, e.g., an infinite product space of finite sets, which is what we have in our problem. One also needs the set of states solving any given task to be closed.]
  2. Finitary greedy algorithm: If from every state one can find a later state that solves a task, and the property of solving a task is preserved by the ordering of states, then every finite set of tasks can be solved by at least one state.

Of course, by combining 1 and 2 one gets the infinite greedy algorithm.

Terence Tao (Oct 26 2024 at 19:57):

[Moved to the other thread.]

Terence Tao (Oct 26 2024 at 20:14):

I guess in our context, a state is a partial function from a (finite or infinite) subset of a free magma to our intended carrier G, and tasks are claims that this partial function is a partial homomorphism, together with various components of the laws one wants to prove or disprove about G that can be computed from this partial homomorphism.

Mario Carneiro (Oct 26 2024 at 20:14):

is that general enough to also cover the asterix construction?

Mario Carneiro (Oct 26 2024 at 20:15):

I would like to have one file which proves greedy works once and for all

Mario Carneiro (Oct 26 2024 at 20:15):

(and it may or may not be a mathlib theorem)

Terence Tao (Oct 26 2024 at 20:17):

Yeah. The one tiny concern I have is that with 1485 one has to carry the directed graph as well as the magma with us in the greedy construction currently, whereas with Asterix one can just work with the magma alone. So one either has to (a) refactor the 1485 analysis to avoid graphs, or (b) find a version of the greedy algorithm that can handle different types of data being greedily constructed (in particular, either just a magma operation, or a magma operation concurrently with a companion graph).

Mario Carneiro (Oct 26 2024 at 20:18):

well, that's why the general version I was talking about in the other thread is a partial order

Mario Carneiro (Oct 26 2024 at 20:18):

this gives you a lot of flexibility to decide what you want to be your state

Mario Carneiro (Oct 26 2024 at 20:19):

but also, there is an advantage to using the bare minimum weakest laws, e.g. a preorder or just a transitive relation, because you usually have to construct it on the spot

Mario Carneiro (Oct 26 2024 at 20:19):

(I'd rather not have to prove such and such is a topological space and such and such is compact!)

Terence Tao (Oct 26 2024 at 20:21):

Yeah so for 1485 I think the states are directed graphs on a free magma together with a partial function into the carrier, that may be required to obey some basic downward closure axioms (if the partial function is defined on a word, it is defined on every subword). Assuming some partial homomorphism tasks, this should then define enough of an operation on the carrier, as well as a directed graph structure on that carrier, which one can then use to formulate other laws and counterexamples. (It will be important to index things by a free magma of words rather than the labels of the carrier because otherwise I think you don't get the compactness.)

Terence Tao (Oct 26 2024 at 20:22):

Ugh actually compactness is a subtle issue here because you technically have an infinite number of choices at each step

Mario Carneiro (Oct 26 2024 at 20:22):

is the idea there that we are doing the greedy construction relative to a counterexample law expressed as a MagmaLaw?

Mario Carneiro (Oct 26 2024 at 20:23):

because we weren't doing free magmas in the original construction IIUC

Mario Carneiro (Oct 26 2024 at 20:23):

it was just G0×NG_0 \times \mathbb{N}

Mario Carneiro (Oct 26 2024 at 20:27):

My branch has been updated to use Andy Jiang 's 5-point multimagma, including the counterexample proofs. Still haven't started on the greedy construction

Terence Tao (Oct 26 2024 at 20:27):

Yeah the problem is this. Given any finite subset FF of G0×NG_0 \times {\mathbb N}, one can trivially put a multiplication table on F×FF \times F that doesn't violate any interesting laws by assigning a completely novel element to each product xyx \diamond y, x,yFx,y \in F, so that no words of length longer than two are actually defined. But one can't take a limit of any such tables to get any complete infinite multiplication table; any novel element you introduce in one table has to eventually show up as a row and column of a future table. The problem here is the non-compactness of (G0×N)(G0×N)×(G0×N)(G_0 \times {\mathbb N})^{(G_0 \times {\mathbb N}) \times (G_0 \times {\mathbb N})}. Now our greedy algorithms don't run into this problem, but I have to think about what is the abstract reason why they don't, there is still some compactness going on that I don't understand.

Terence Tao (Oct 26 2024 at 20:28):

I thought I could fix this by indexing things by words of generators rather than elements of the carrier, but that's not quite the right approach. Thinking.

Mario Carneiro (Oct 26 2024 at 20:29):

But one can't take a limit of any such tables to get any complete infinite multiplication table; any novel element you introduce in one table has to eventually show up as a row and column of a future table.

I don't see the issue here, all elements of the table are tasks so the limit will have solved this

Terence Tao (Oct 26 2024 at 20:31):

The problem is that no limit exists. If for table N you set all products in the N x N table to be larger than N, there is no subsequence (or subnet) of these tables that converges pointwise to an infinite table.

Terence Tao (Oct 26 2024 at 20:31):

OK the missing axiom is not topological, it is that every chain in the partially ordered set has an upper bound

Mario Carneiro (Oct 26 2024 at 20:32):

But it is converging? I'm very confused

Terence Tao (Oct 26 2024 at 20:32):

more precisely every directed set in the partially ordered set (in the complete set of states, not the subset that one runs the greedy algorithm on) has an upper bound. Which by Zorn's lemma is equivalent to the chains having an upper bound I think.

Mario Carneiro (Oct 26 2024 at 20:32):

If I just need to fill out the table with numbers then I can always just pick the next larger number and ensure I visit every cell of the table through some fair scheduling

Mario Carneiro (Oct 26 2024 at 20:33):

This is basically what the cantor pairing function looks like

Terence Tao (Oct 26 2024 at 20:33):

Yeah as long as the tables form a chain. I was thinking of a sequence of unrelated tables, where the N+1 x N+1 table is not an extension of the previous N x N table. These ones don't converge.

Mario Carneiro (Oct 26 2024 at 20:33):

Ah okay

Terence Tao (Oct 26 2024 at 20:33):

So basically the topological notion of compactness is not suitable here.

Terence Tao (Oct 26 2024 at 20:35):

Kyle's example of 4-coloring is one where there is topological compactness: there are only four choices for each task, that's a compact set. So there we can reduce the infinite problem to a finite problem. But here we do not seem to have that sort of compactness. So you can ignore all my stuff about free magmas, this was my attempt to compactify something that probably can't be usefully compactified.

Terence Tao (Oct 26 2024 at 20:49):

Basically there is a Venn diagram of "compactness arguments" and "greedy arguments" to construct examples that do various interesting things. Some arguments can proceed both by compactness and by greediness, but there are other arguments that only can proceed by one but not the other. But it's best to think of them as orthogonal types of constructions.

Mario Carneiro (Oct 26 2024 at 21:39):

Do you know how we can express "there exists a seed extension satisfying the anti-N law and Claim 1'' " as a finite computation on the base multimagma G0?

Terence Tao (Oct 26 2024 at 21:42):

It's more of a semidecidable computation: once one has a candidate seed, one can verify it in a finite computation, but it isn't obvious that such a seed exists, or if it does, that there is a usable (not Busy-Beaver type) bound on its size. Basically we run a naive greedy algorithm: plonk down the minimal number of elements needed to create a counterexample, and see of all the rules in our greedy ruleset are obeyed. If not, we may have to add a few more elements to satisfy them, but this may trigger other rules to be verified. In practice, because we so often have the option to leave the magma operation undefined, it is relatively easy to dodge the scenario of an exponentially growing number of rules to satisfy and get the situation to close, but we don't have any theoretical guarantee that this is always possible for any ruleset. (For instance, if the counterexample is actually a logical consequence of the ruleset, then it can't be possible, of course, but in that case the counterexample would most likely be a logical consequence of the actual equation - or else we chose a bad choice of base magma.)

Mario Carneiro (Oct 26 2024 at 21:43):

Right, I am not asking to determine what the seed is, but rather what are the actual constraints we need to check

Mario Carneiro (Oct 26 2024 at 21:45):

In particular, it would be nice if the underlying type G0×NG_0\times\mathbb{N} does not "leak" from the greedy construction, and any finite type suffices to prove a seed exists

Terence Tao (Oct 26 2024 at 21:46):

Any partially defined finite magma with a homomorphism to the base magma can be used as a seed, and relabeled into a finite subset of G0×NG_0 \times {\mathbb N} by some arbitrary numbering scheme

Terence Tao (Oct 26 2024 at 21:46):

Well for us we have a graph as well as a magma, so the rules are slightly more complicated, but that's the general idea

Mario Carneiro (Oct 26 2024 at 21:48):

In the example you gave earlier, every variable was mapped to a distinct natural number. Is that a representative situation?

Mario Carneiro (Oct 26 2024 at 21:51):

It seems like the theorems like

theorem anti_3457 :  x y z w, IsGood x z x  IsGood z w y  ¬IsGood x z w := ...

are not really sufficient in this setting, we need to add more things in the conjunction or rephrase them somehow

Terence Tao (Oct 26 2024 at 21:52):

Well if the law you want to refute contains a square, such as xxx \diamond x, you want to make both instances of xx map to the same element of G0×NG_0 \times {\mathbb N}. But one can have situations in which one has two variables x,yx, y which you want to assign the same color in G0G_0 to get the refutation, but you have no reason to require them to otherwise be equal. Then I think it usually can't hurt to give them different numbers to force them to be unequal, as this tends to make it easier to check what remaining axioms are left for the seed to obey. However, I know that for some greedy constructions there are rulesets where the conclusion is an equality, e.g., xy=zx \diamond y = z and xy=zx'\diamond y = z implies x=xx = x' (right cancellability), in which case it is not always advantageous to make every single variable in one's seed distinct.

Terence Tao (Oct 26 2024 at 21:54):

It's enough for G0 to obey a law like anti_3457. Then one can use that to build a seed up in G0 × ℕ that obeys that law with the additional properties that x y z w are distinct, and that Claim 1'' is satisfied (the latter is not completely automatic, but is an easy check if you give the x y z w different numbers).

Mario Carneiro (Oct 26 2024 at 21:55):

Well, what I mean is that that theorem is not the beginning and end of all 3457-specific reasoning about this multimagma

Terence Tao (Oct 26 2024 at 21:55):

Yes. There's a small amount of additional argument required for each law. It's easy, but perhaps not fully automated.

Mario Carneiro (Oct 26 2024 at 21:55):

because we also need to verify that there exists a seed extension that can be built from the specific values used in the counterexample

Mario Carneiro (Oct 26 2024 at 21:55):

so the existential quantification is too early as stated

Terence Tao (Oct 26 2024 at 21:56):

One could automate it as a sat solving problem and use that to autogenerate Lean code if we had to scale this out to many more implications than these 4.

Terence Tao (Oct 26 2024 at 21:57):

OK one could state the required condition as that there exists an anti-3457 seed then that also obeys Claim 1''. It's only slightly harder to verify that than the original version of anti_3457

Mario Carneiro (Oct 26 2024 at 21:57):

I'm not talking about automation here, I'm talking about the logical structure of the proof (whether or not the parts are proved by automation - I think all the examples so far are easy enough to do manually)

Mario Carneiro (Oct 26 2024 at 21:58):

There is some "interface" between the greedy construction infrastructure and the specifics of the 1485 multimagma and I'm trying to figure out how to state it in order to minimize the amount of example-specific checking

Terence Tao (Oct 26 2024 at 21:58):

and one can derive it easily from anti_3457 by a manual calculation starting from something like obtain (x,y,z,w) = anti_3457; use (x,0), (y,1), (z,2), (w,3)(or whatever the Lean syntax is)

Mario Carneiro (Oct 26 2024 at 21:59):

right but what is the goal of the theorem starting with that as the proof?

Mario Carneiro (Oct 26 2024 at 22:00):

if the values of x,y etc are not known (as they would be if you did that), then I'm not sure how you are checking the uniqueness restrictions (and other requirements) on those new witness values

Terence Tao (Oct 26 2024 at 22:01):

∃ seed : DirectedGraph (G₀ × ℕ), IsExtension G₀ seed ∧ anti_1537 seed ∧ Claim_1'' seed ∧ Finite seed

Mario Carneiro (Oct 26 2024 at 22:02):

Isn't every seed going to satisfy IsExtension G₀ seed here?

Mario Carneiro (Oct 26 2024 at 22:02):

if we are fixing the homomorphism to be the first projection on G₀ × ℕ

Terence Tao (Oct 26 2024 at 22:03):

The edges of the seed need to map to edges of G_0

Terence Tao (Oct 26 2024 at 22:03):

i.e. the projection map needs to be a graph homomorphism

Mario Carneiro (Oct 26 2024 at 22:04):

Okay so the seed requires me to give

  • A list of edges, whose projection yield edges in G0
  • With the good and bad triples lifted from G0, verify that claim 1'' holds

?

Terence Tao (Oct 26 2024 at 22:05):

And also a witness to whatever law you are trying to disprove

Terence Tao (Oct 26 2024 at 22:05):

and a Finite instance in order to run the greedy algorithm

Mario Carneiro (Oct 26 2024 at 22:05):

Finite will certainly be structurally true

Mario Carneiro (Oct 26 2024 at 22:07):

  • And points x,y,z,w whose projection satisfy anti_3457 and where all the good triples and bad triples are witnessed in the seed

Mario Carneiro (Oct 26 2024 at 22:07):

or is it edges that have to be witnessed in the seed?

Terence Tao (Oct 26 2024 at 22:08):

Both, but the way things are set up it's equivalent. A seed witneeses a good triple with a path iff the two edges of the path are in the seed, and the projection to the base is good there, and similarly for bad.

Mario Carneiro (Oct 26 2024 at 22:09):

like in the case of anti_3457

theorem anti_3457 :  x y z w, IsGood x z x  IsGood z w y  ¬IsGood x z w := ...

we need xz, zx, zw, wy, xz, zw to be in the seed?

Mario Carneiro (Oct 26 2024 at 22:09):

or is it sufficient to just have the bad edges in the seed

Terence Tao (Oct 26 2024 at 22:09):

Yeah so I spoke incorrectly earlier with regards to use (x,0), (y,1), (z,2), (w,3). It's not enough to specify the vertices of the seed, you need the edges too.

Terence Tao (Oct 26 2024 at 22:09):

You need both the good and bad paths to be witnessed

Terence Tao (Oct 26 2024 at 22:10):

But as long as the seed projects down to a graph in the base where the good and bad paths are witnessed there by the projections, you have that.

Terence Tao (Oct 26 2024 at 22:10):

Well okay if two edges project down to the same edge then that statement is not literally true, but it is almost true

Mario Carneiro (Oct 26 2024 at 22:10):

The paths are automatic, I think, once you specify the vertices, since it's a prop (this is a simple graph)

Mario Carneiro (Oct 26 2024 at 22:11):

but you have to prove each path exists in the given list

Terence Tao (Oct 26 2024 at 22:12):

yes so in your example you do need xz, zx, xw, wy, xz, xw to be in the seed and then the paths make sense. It may be helpful to set up an IsBad predicate which isn't simply the negation of IsGood but asserts that the edges are there in the seed, but not forming a good path.

Terence Tao (Oct 26 2024 at 22:13):

(or equivalently: that the edges are in the seed, and their projection to the base is a bad path.)

Matthew Bolan (Oct 26 2024 at 22:14):

I've been trying to greedily construct a graph for 151 "by hand" and in the process think I've figured out the "abstract nonsense" I was dancing around earlier. If we let the carrier be some free magma, we can define a relaxed multiplication xy={xy = \{ all expressions equal to xyxy under 1485 }\}. This works as for all aa in yxyx and all bb in x(yz)x(yz), abab is equal under 1485 to (yx)(x(yz))(yx)(x(yz)), which in turn is equal to xx. My hope to tweak this to make some element obviously not a square seems dashed though, unless I'm missing something obvious. I guess I'll switch to an SAT solver unless someone else is already running one.

Terence Tao (Oct 26 2024 at 22:15):

You mean 1483 or 1485?

Matthew Bolan (Oct 26 2024 at 22:15):

1485, sorry

Matthew Bolan (Oct 26 2024 at 22:17):

So it is like modding out by the relations implied by 1485, but only "entrywise"

Matthew Bolan (Oct 26 2024 at 22:17):

It's also probably possible to make the carrier basically any magma with a little more care.

Terence Tao (Oct 26 2024 at 22:19):

I think this abstract nonsense is taking the free 1485 magma (== the completely free magma, modulo equivalence of words under 1485) and lifting from this quotient back up to the full magma to get a relaxed magma. Abstract nonsense tells us that if there is a counterexample to say 151 for some magma, it will also appear in the free 1485 magma, and hence also in the relaxed magma of all words with the ambiguity of 1485 equivalence. But because the 1485 law isn't confluent and doesn't seem to have an obvious complete rewriting system, we don't understand the "equivalence under 1485" relation well enough to say anything that isn't abstract nonsense here, I think.

Terence Tao (Oct 26 2024 at 22:22):

oh actually maybe you are doing something slightly different than the free magma modulo laws construction, in which you only get to use the law to contract words and not to expand them, or something. Hmm... there's no free lunch though.

Matthew Bolan (Oct 26 2024 at 22:22):

Yes, I think I agree. I merely noticed that most of my greedy attempts wound up building something like this thing if I wasn't greedy enough. I think to get any mileage from this you'd have to pick something much cleverer than the free magma to carry it, at which point you've probably already done most of the work.

Mario Carneiro (Oct 26 2024 at 23:55):

Is there a characterization of the weakest class of magmas for which the multimagma technique applies? In the language of graphs, this is the ones for which Claims 1-3 hold but not necessarily Claim 4 (which is a version of equation 1485)

Mario Carneiro (Oct 26 2024 at 23:56):

working it out I guess the property you really need is (∃ z, x = z ◇ y) <-> (∃ z, x ◇ z = y), but I'm not sure if there is a way to write that purely equationally

Terence Tao (Oct 26 2024 at 23:57):

Well, we only have one working example of this approach so far, and in that example the magma is entangled with a graph. We've just started to explore what it might look like for 854, where we are not initially trying to introduce a graph, but I don't know whether it's going to work.

Terence Tao (Oct 26 2024 at 23:57):

Yeah, that's basically the axiom we're leaning heavily on. Basically means that the magma is "central groupoid-like".

Mario Carneiro (Oct 26 2024 at 23:58):

Well what I mean is that you can do this graph thing as long as you have the above equivalence, which is true for 1485 but 1485 is stronger than that and includes this 5-cycle condition on top

Mario Carneiro (Oct 26 2024 at 23:58):

maybe that should be the definition of an actual "weak central groupoid"?

Terence Tao (Oct 26 2024 at 23:58):

I guess one could call that the "very weak central groupoid axiom".

Terence Tao (Oct 26 2024 at 23:59):

Very Mathlib-esque to try to find the absolute minimal hypotheses required to do any math :grinning_face_with_smiling_eyes:

Mario Carneiro (Oct 27 2024 at 00:00):

well, it's often the best way to simplify an argument once you have the core of the idea

Mario Carneiro (Oct 27 2024 at 00:01):

Also I'm a mathlib maintainer so what do you expect :wink:

Mario Carneiro (Oct 27 2024 at 00:01):

for me proving the theorem is just the beginning of the simplification process

Mario Carneiro (Oct 27 2024 at 00:02):

I proved the generalized greedy theorem in the following form

theorem exists_greedy_chain {α β : Type*} [Preorder α] [Countable β]
    (task : α  β  Prop) (H :  a b,  a', a  a'  task a' b) (a : α) :
     c, IsChain (·  ·) c  ( x  c, a  x)   b,  a  c, task a b := by

Terence Tao (Oct 27 2024 at 00:02):

You could artificially turn this very weak axiom into an equational axiom by introducing a somewhat arbitrary operation that takes an edge in the directed graph and returns a way to move the edge either forward or backward in a way that creates a good path (but typically there will not be a unique choice here).

Mario Carneiro (Oct 27 2024 at 00:02):

the next step is to make a version tailored to 1485 style proofs

Terence Tao (Oct 27 2024 at 00:04):

I think you want to impose non-emptiness on α and conclude non-emptiness of c

Mario Carneiro (Oct 27 2024 at 00:04):

Nonemptiness is implied by the given a : A

Terence Tao (Oct 27 2024 at 00:04):

Ah got it

Mario Carneiro (Oct 27 2024 at 00:04):

but the conclusion could also mention that a is in c

Terence Tao (Oct 27 2024 at 00:05):

At this stage you're not imposing upward closure on task, is that meant for a corollary?

Mario Carneiro (Oct 27 2024 at 00:06):

it's not actually needed for this statement of the theorem, since I have to say that every element is somewhere in the chain anyway

Terence Tao (Oct 27 2024 at 00:06):

Fair enough

Terence Tao (Oct 27 2024 at 00:09):

Regarding "very weak central groupoids": every magma generates a right Cayley graph (connecting xyx \to y if y=xzy = x \diamond z for some zz) and left Cayley graph (connecting xyx \to y if y=zxy = z \diamond x). Here the special thing is that the two graphs are dual to each other. But I guess one could just always have the trio of magma and two Cayley graphs, and add the duality as another equational axiom.

Mario Carneiro (Oct 27 2024 at 00:10):

dual?

Mario Carneiro (Oct 27 2024 at 00:10):

aren't they the same in this setting?

Terence Tao (Oct 27 2024 at 00:10):

Edge reversal.

Mario Carneiro (Oct 27 2024 at 00:10):

oh that dual

Terence Tao (Oct 27 2024 at 00:10):

I guess it's called transpose

Terence Tao (Oct 27 2024 at 00:12):

For most magmas we get the path xRxyLyx \stackrel{R}{\to} x \diamond y \stackrel{L}{\leftarrow} y but for the central groupoid like magmas we get xRxyRyx \stackrel{R}{\to} x \diamond y \stackrel{R}{\to} y

Mario Carneiro (Oct 27 2024 at 00:13):

Terence Tao said:

Very Mathlib-esque to try to find the absolute minimal hypotheses required to do any math :grinning_face_with_smiling_eyes:

Also I think there are good reasons to minimize assumptions in this setting since there is a whole zoo of not obviously related equations we are looking at here and it is good to know on which equations we can apply this technique

Mario Carneiro (Oct 27 2024 at 00:14):

Are the two implications self-dual?

Terence Tao (Oct 27 2024 at 00:14):

Which two?

Mario Carneiro (Oct 27 2024 at 00:14):

in the very weak central groupoid axiom

Mario Carneiro (Oct 27 2024 at 00:15):

in chapter 11 this was proved using the fact that 1485 is self-dual

Terence Tao (Oct 27 2024 at 00:15):

Well they don't imply each other for sure, take the left-absorptive magma xy=xx \diamond y = x for instance. If a magma obeys one implication then its dual will obey the other.

Mario Carneiro (Oct 27 2024 at 00:16):

I don't have a good intuition whether you can still do the greedy extension process when there are two cayley graphs

Mario Carneiro (Oct 27 2024 at 00:17):

Why do we even need claims 2 and 3 in the construction?

Mario Carneiro (Oct 27 2024 at 00:17):

claim 1 seems like the important one for building a magma

Terence Tao (Oct 27 2024 at 00:17):

We're actually kind of lucky here that extending a partial weak central groupoid by adding two new edges doesn't create a bunch of extra side conditions. I suspect it causes trouble in 854 though.

Terence Tao (Oct 27 2024 at 00:18):

Actually claims 2 and 3 might now be dropped as we have baked the counterexamples directly into the seed. They were useful when the seed was empty and I had to conjure up the counterexample somehow from the general properties of the extension.

Mario Carneiro (Oct 27 2024 at 00:20):

I'm trying to figure out why this technique isn't extremely general and proves all counterexamples

Mario Carneiro (Oct 27 2024 at 00:20):

somehow we aren't using any interesting properties of either the law we want to prove or the one we want to refute

Terence Tao (Oct 27 2024 at 00:21):

It's in checking that extending a partial magma by adding one or two elements/edges/good paths doesn't cause a violation of one of the existing axioms. For 1485 we luck out that it is almost a triviality, but this is not usually the case. Um let me try to illustrate with 854 as I had tried this recently. one sec

Terence Tao (Oct 27 2024 at 00:28):

The equation is x = x ◇ ((y ◇ z) ◇ (x ◇ z)). Suppose one has gotten some partial magma obeying this law. One can try to pick a pair a, b for which a ◇ b is not currently defined, and set it equal to some new object c. Easy, right?

Well, you could be blocked. If b is already of the form b = (y ◇ z) ◇ (a ◇ z) then a ◇ b is "spoken for", it has to equal a.

So one can add a new axiom: if (y ◇ z) ◇ (a ◇ z) is already equal to b, then a ◇ b must equal a. Fine, that's just restating 854, you'd expect that.

But this creates secondary consequences. Suppose that a is of the form y ◇ z and b is of the form x ◇ z already. Then this new axiom means that setting a ◇ b = c forces x = x ◇ c. OK, so you add that in. But now if you are forced to have y = y ◇ c and z = x ◇ c then you are now also forced to have x = x ◇ (y ◇ x). This could contradict some existing choices you made in the partial magma, so now you have to impose a new axiom preventing that. But then all the above changes could ruin that axiom unless you impose a further axiom, and so on. We know that the automated approach to close this procedure does not seem to work; I'm hoping that by imposing the additional axioms that the magma is an extension of a well chosen base that somehow these problems go away like they did for 1485, but I don't know if this is going to work yet.

Mario Carneiro (Oct 27 2024 at 00:30):

But then why doesn't this happen for 1485?

Mario Carneiro (Oct 27 2024 at 00:31):

The chapter 11 proof just says it is automatic that claim 4 is preserved

Terence Tao (Oct 27 2024 at 00:32):

Because the most problematic 1485 axiom - Claim 4 - we have managed to make purely a consequence of being an extension of the base G0, which was designed to already obey Claim 4, together with the graph homomorphism axiom (which I didn't give a name, but it is part of the definition of an extension). The graph homomorphism axiom is much simpler than Claim 4 and is really easy to extend greedily because it doesn't couple a whole lot of vertices together.

If we could also "push" most of the difficult part of Equation 854 to the base and leave only some very simple axiom upstairs that is much easier to extend then we'd be in business.

Mario Carneiro (Oct 27 2024 at 00:33):

Right, what is the property of 1485 that means that it is preserved by homomorphisms

Mario Carneiro (Oct 27 2024 at 00:34):

I'm looking at the proof and I don't see anything that is 1485 specific

Terence Tao (Oct 27 2024 at 00:35):

OK, this is a good question. I indeed see that the argument is suspiciously general-looking. It's something to do with the graph, but I can't pinpoint it yet. Thinking.

Mario Carneiro (Oct 27 2024 at 00:36):

maybe I shouldn't have left this for last :stuck_out_tongue:

Terence Tao (Oct 27 2024 at 00:41):

OK, I think I get it. Usually, knowing the left and right Cayley graphs is not enough information to determine the magma operation. But in the case of extensions obeying Claim 1'', it is: xyx \diamond y is the unique zz (if it exists) such that xRzLyx \stackrel{R}{\to} z \stackrel{L}{\leftarrow} y and such that the image of this path in the base graph is a good path. As such, any law, when specialized to an extension obeying Claim 1'', can now be converted to a purely graph-theoretic statement.

In which case, this is in fact a general approach that could in principle work for any magma. One still has to check that the graph theoretic rules one has can be completed to a ruleset that admits greedy extensions, but basically the use of the base magma has transformed the problem, hopefully from an unsolvable one to a solvable one.

Terence Tao (Oct 27 2024 at 00:44):

On the other hand, this sort of construction presupposes that a right image xGx \diamond G and a left image GyG \diamond y are inclined to only intersect once for each fiber of the extension, which is reasonable for central groupoid-like objects (which are horribly non-cancellative on either the left or the right), but would never be satisfied for say groups in which left and right multiplication are bijections. But for 854 we have a chance, the magmas we've seen there are pretty far from being cancellative in either the left or the right.

Mario Carneiro (Oct 27 2024 at 00:44):

Claim 1'' isn't actually saying anything about the graphs or depending on the very weak central groupoid axiom... When caching it out in magma language it's just saying you have a partial magma?

Mario Carneiro (Oct 27 2024 at 00:45):

In my formalization it is literally just a partial function with the type of a magma operation

Terence Tao (Oct 27 2024 at 00:46):

It's preventing one from having two paths xzyx \to z \to y, xzyx \to z' \to y in the graph that both project down to a good path in G0G_0. This is a non-trivial assertion about the graph.

One could alternatively spell it in terms of a partial magma operation function \diamond with the property that a path xzyx \to z \to y in the partial graph projects down to a good path iff xyx \diamond y is defined and equal to zz.

Mario Carneiro (Oct 27 2024 at 00:48):

So my formalization looks something like this so far:

abbrev ExtBase := G × Nat

abbrev Extension := ExtBase G  ExtBase G  Part (ExtBase G)

structure Extension.OK (e : Extension G) where
  finite : Set.Finite {x : ExtBase G × ExtBase G | (e x.1 x.2).Dom}
  consistent :  a b,  c, c  e a b  IsGood a.1 c.1 b.1

Mario Carneiro (Oct 27 2024 at 00:48):

The last property is the one doing the work here, I think

Mario Carneiro (Oct 27 2024 at 00:49):

The "at most one" property is ensured by the use of Part A (which you can think of as a set with at most one element)

Terence Tao (Oct 27 2024 at 00:51):

You're also going to need a directed graph E : Finset (ExtBase × ExtBase) as part of your Extension and IsGood a.1 c.1 b.1 needs to be combined with (a,c) ∈ E ∧ (c,b) ∈ E

Terence Tao (Oct 27 2024 at 00:53):

Without it there aren't going to be any extensions, because e a b has to equal every single element of the fiber above a candidate c.1

Mario Carneiro (Oct 27 2024 at 00:54):

So the idea is that in an extension, you can decide you want to include a -> c as a path, and you can (independently) include a triple a -> b -> c which is good?

Terence Tao (Oct 27 2024 at 00:55):

Yeah, except that they are totally not at all independent because of consistent. Specify the edges and you specify the partial operation / good paths, and vice versa.

Mario Carneiro (Oct 27 2024 at 00:56):

Can I include two paths but not yet commit to it being colored good or bad?

Terence Tao (Oct 27 2024 at 00:56):

(Well, technically the operation doesn't fully describe the edge set unless you also impose Claim 2 and Claim 3. You can put in some "junk" edges that don't connect to anything else if you want, if those extra claims are dropped.)

Terence Tao (Oct 27 2024 at 00:57):

No, because consistent will force you to label these paths using the labeling of the base.

Terence Tao (Oct 27 2024 at 00:58):

Add an edge, and all paths involving that edge are now classified as good or bad automatically, you have no flexibility in this.

Mario Carneiro (Oct 27 2024 at 00:58):

So maybe I just need a Finset (ExtBase × ExtBase) then? And this at most one property is just a property about it and not a piece of data

Terence Tao (Oct 27 2024 at 00:58):

Yeah. And now you've come back to Claim 1''.

Terence Tao (Oct 27 2024 at 00:59):

Both spellings are mathematically fine I think, it's just up to you which one works better in Lean.

Mario Carneiro (Oct 27 2024 at 01:00):

But that seems harder to work with, because then when you add an edge you get a whole bunch of new good and bad triples and have to figure out whether that's okay

Terence Tao (Oct 27 2024 at 01:00):

The beauty is that laws on the base automatically extend to laws on the extension, in particular Claim 4. So you don't need to recheck this every time you add an edge, it comes for free.

Terence Tao (Oct 27 2024 at 01:01):

So long as you preserve the all-important Claim 1''

Mario Carneiro (Oct 27 2024 at 01:01):

Actually, what even is the extension at this point? It's just a subset of the multiplication table domain?

Mario Carneiro (Oct 27 2024 at 01:01):

I think I know what the limit of that is

Mario Carneiro (Oct 27 2024 at 01:02):

Something here feels like cheating

Mario Carneiro (Oct 27 2024 at 01:03):

There is a choice we have to make during the extension process, which is where to make an individual multiplication map to (the natural number coordinate). A Finset (ExtBase x ExtBase) isn't enough to capture that

Terence Tao (Oct 27 2024 at 01:06):

Given E: Finset (ExtBase x ExtBase), the product x ◇ y for x y: ExtBase is the unique point z: ExtBase such that IsGood x.1 z.1 y.1 ∧ (x,z) ∈ E ∧ (z,y) ∈ E, if it exists. It's completely determined by the Finset (and the base G0).

Mario Carneiro (Oct 27 2024 at 01:06):

Terence Tao said:

You're also going to need a directed graph E : Finset (ExtBase × ExtBase) as part of your Extension and IsGood a.1 c.1 b.1 needs to be combined with (a,c) ∈ E ∧ (c,b) ∈ E

Coming back to this, what would happen if we just said that this finset is the same as the domain of the map? So that means that c ∈ a ◇ b means that a ◇ c and c ◇ b are defined

Terence Tao (Oct 27 2024 at 01:08):

That would give a finite set that contains the actual a ◇ b that we want (if it exists), but a lot of "fake" values c as well in which the path a → c → b projects down to a bad path. You can define this, but I don't think it's terribly useful to do so.

Mario Carneiro (Oct 27 2024 at 01:08):

why would that happen?

Mario Carneiro (Oct 27 2024 at 01:09):

I would also assume that if c ∈ a ◇ b then IsGood a.1 c.1 b.1

Terence Tao (Oct 27 2024 at 01:10):

Hang on maybe you have a typo in what you wrote. When you said a ◇ c and c ◇ b are defined I think you meant (a,c) ∈ E ∧ (c,b) ∈ E. That's fine. If you have the map, you can define the graph; if you have the graph, you can define the map. So you can implement the extension just using the graph, or just using the map, or as a bundle of both. Mathematically, it doesn't matter, it's just up to Lean aesthetics.

Mario Carneiro (Oct 27 2024 at 01:11):

No, that's not a typo, I'm saying what if we enforce that E is the same as the domain of the partial function

Mario Carneiro (Oct 27 2024 at 01:11):

since it seems that they are redundant with each other

Mario Carneiro (Oct 27 2024 at 01:12):

Since the partial function induces a set and the set induces a partial function I don't know whether having a consistent pair in this way actually works

Terence Tao (Oct 27 2024 at 01:13):

E and the domain of the partial function are both of type Finset (ExtBase × ExtBase), but there is no reason to expect them to agree. (a,b) ∈ E means that there will one day be a c for which b = a ◇ c (or one day there will be a d for which a = d ◇ b); it doesn't mean that a ◇ b is itself defined.

Terence Tao (Oct 27 2024 at 01:15):

The relation between E and the partial function is that E is the set of all pairs that are either of the form (a, a ◇ b) or (a ◇ b, b) for a, b with a ◇ b defined by the partial function.

Terence Tao (Oct 27 2024 at 01:16):

i.e. the partial function generates a bunch of good paths a → a ◇ b → b, E is the union of those paths, and this can also generate some bad paths too inside E.

Mario Carneiro (Oct 27 2024 at 01:16):

but once we have a compatible pair (a,b) ∈ E, (b, c) ∈ E, we must have a ◇ cdefined to be something?

Terence Tao (Oct 27 2024 at 01:17):

Not necessarily. If IsGood a.1 b.1 c.1 then a ◇ c will be forced to equal b. Otherwise, its definition may still be pending further extension of E.

Terence Tao (Oct 27 2024 at 01:19):

Or if there is another b' with (a,b') ∈ E(b', c) ∈ E and  IsGood a.1 b'.1 c.1 then a ◇ c will equal b' instead. But we may not have yet created enough edges in E for this to happen.

Terence Tao (Oct 27 2024 at 01:19):

That's task (a, c).

Mario Carneiro (Oct 27 2024 at 01:20):

Let's say for generality that we have both a set and a partial function. Are these the constraints:

  • E is finite
  • If c ∈ a ◇ b then (a, c) ∈ E ∧ (c, b) ∈ E
  • If a ◇ b is defined then for all c, c ∈ a ◇ b iff IsGood a.1 c.1 b.1

Terence Tao (Oct 27 2024 at 01:21):

Not quite. It's

  • E is finite
  • For all a b c, c ∈ a ◇ b iff IsGood a.1 c.1 b.1 ∧ (a, c) ∈ E ∧ (c, b) ∈ E

Mario Carneiro (Oct 27 2024 at 01:22):

okay, and then an extension step consists of adding an element to E and then a bunch of things happen and that's okay?

Mario Carneiro (Oct 27 2024 at 01:23):

it seems like with this version you can't control what happens to the partial function

Terence Tao (Oct 27 2024 at 01:23):

Two elements to E (of the form (a,c) and (c,b) where c is new, and it requires a check that Claim 1'' is preserved, but yes.

Terence Tao (Oct 27 2024 at 01:24):

Yeah, the partial function does weird things. But we can largely ignore it until the very end when we want to convert the (now completed) graph back into a weak central groupoid

Mario Carneiro (Oct 27 2024 at 01:24):

But if b is old, then it could have another (b, d) \in E and then the partial function has gained more new elements

Mario Carneiro (Oct 27 2024 at 01:24):

in particular, it might gain an element inconsistent with the value it already has

Terence Tao (Oct 27 2024 at 01:25):

Yeah. As long as we keep Claim 1'' satisfied, these bad things don't happen. The one nontrivial thing is to ensure that Claim 1'' remains true after adding the two edges.

Mario Carneiro (Oct 27 2024 at 01:25):

which sounds like the issue you were talking about for 854

Mario Carneiro (Oct 27 2024 at 01:26):

maybe I'll just try writing it your way and see where things end up and maybe I will be able to ask better questions then

Terence Tao (Oct 27 2024 at 01:26):

Yes but Claim 1'' is a lot simpler than 854. Somehow being able to use an ansatz in which the magma operation is completely described by one graph helps a lot. With 854 we may need two graphs, not sure whether it's going to work yet.

Mario Carneiro (Oct 27 2024 at 01:27):

Terence Tao said:

Not quite. It's

  • E is finite
  • For all a b c, c ∈ a ◇ b iff IsGood a.1 c.1 b.1 ∧ (a, c) ∈ E ∧ (c, b) ∈ E

Oh, also does c ∈ a ◇ b imply (a, b) \in E here?

Terence Tao (Oct 27 2024 at 01:28):

No

Terence Tao (Oct 27 2024 at 01:28):

E is unrelated to the domain of the partial function

Mario Carneiro (Oct 27 2024 at 01:30):

But the tasks are to make sure the partial function is filled up, I guess

Terence Tao (Oct 27 2024 at 01:31):

OK there is something subtle about the 1485 law x = (y ◇ x) ◇ (x ◇ (z ◇ y)) that the other laws don't necessarily have. When you have enough of the graph needed to define both y ◇ x and x ◇ (z ◇ y), i.e., you have both x → y ◇ x → x and x → x ◇ (z ◇ y) → z ◇ y as good paths, then miraculously you automatically have enough of the graph already there to be able to say whether x is equal to (y ◇ x) ◇ (x ◇ (z ◇ y)), because the path y ◇ x → x → x ◇ (z ◇ y) is automatically right there for you to use and you just need to test whether it projects to a good path or a bad path. This is a special property of central groupoid type laws, not enjoyed by other laws.

Terence Tao (Oct 27 2024 at 01:33):

In particular this is probably going to be tricky to replicate for 854.

Terence Tao (Oct 27 2024 at 01:43):

The ansatz c ∈ a ◇ b ↔ IsGood a.1 c.1 b.1 ∧ (a, c) ∈ E ∧ (c, b) ∈ Eis particularly well suited for testing the law 1485 x = (y ◇ x) ◇ (x ◇ (z ◇ y)) law in particular because the (a, c) ∈ E ∧ (c, b) ∈ E components of this ansatz are propagated automatically and are almost tautological to verify (in fact I think they would literally follow from tauto). Only the IsGood a.1 c.1 b.1 component needs verification and this is why it reduces to just verifying something on the base.

Terence Tao (Oct 27 2024 at 01:45):

Thanks for the lengthy discussion, this actually clarified a lot for me why this argument actually works for 1485, and why it will be challenging to transport it to other equations that do not resemble central groupoid laws.

Mario Carneiro (Oct 27 2024 at 01:46):

Assuming the invariant I quoted above, what is the property we obtain on the limit object?

Mario Carneiro (Oct 27 2024 at 01:46):

Obviously the partial function is now total, but I'm not sure what happened to E

Mario Carneiro (Oct 27 2024 at 01:47):

I guess it's still not everything (or else it would be kind of trivial)

Terence Tao (Oct 27 2024 at 01:47):

E is some complicated directed graph which is determining the magma operation.

Mario Carneiro (Oct 27 2024 at 01:48):

Is it sufficient to just forget about it and keep the property about IsGood?

Terence Tao (Oct 27 2024 at 01:48):

A model case to keep in mind is that of a natural central groupoid S×SS \times S with operation (a,b)(c,d)=(b,c)(a,b) \diamond (c,d) = (b,c). Here E(S×S)×(S×S)E \subset (S \times S) \times (S \times S) is the set of pairs ((a,b),(b,c))((a,b), (b,c)) in which the second coordinate of the first element equals the first coordinate of the second.

Terence Tao (Oct 27 2024 at 01:49):

You really want to drop E don't you? :sweat_smile: No if you drop the membership in E then there is no way a ◇ b will be uniquely defined; if c obeys the condition then any other c' with c.1 = c'_1 will also obey the condition.

Mario Carneiro (Oct 27 2024 at 01:50):

No I mean after the construction is done

Mario Carneiro (Oct 27 2024 at 01:50):

In the limit object, the partial function induced by E is now a total function, but we have this complicated condition relating it to a combination of E and IsGood and I don't know what to do with that

Mario Carneiro (Oct 27 2024 at 01:51):

Do we want to say that E coincides with Path?

Terence Tao (Oct 27 2024 at 01:52):

Once you have completed the greedy construction and determined the (now total) operation , it is true that E no longer plays a role, the only thing that is needed now is the implication c ∈ a ◇ b → IsGood a.1 c.1 b.1. This is not an if and only if, because E will be far from total, but it's good enough to prove 1485, due to the miracle I mentioned earlier.

Mario Carneiro (Oct 27 2024 at 01:52):

ah okay, that's what I was hoping

Terence Tao (Oct 27 2024 at 01:52):

Wait what I said is wrong

Mario Carneiro (Oct 27 2024 at 01:52):

:sweat_smile:

Terence Tao (Oct 27 2024 at 01:53):

You need the full statement c ∈ a ◇ b ↔ IsGood a.1 c.1 b.1 ∧ (a, c) ∈ E ∧ (c, b) ∈ E to verify 1485

Mario Carneiro (Oct 27 2024 at 01:54):

And it's fine if E is any set at all as long as it participates in that relation?

Terence Tao (Oct 27 2024 at 01:54):

In particular you need to use this statement in the forward direction to get ((y ◇ x),x) and (x, (x ◇ (z ◇ y)) in E, then together with IsGood (y ◇ x).1 x.1 (x ◇ (z ◇ y)).1 and the statement in the backwards direction you get x = (y ◇ x) ◇ (x ◇ (z ◇ y).

Terence Tao (Oct 27 2024 at 01:54):

Yes

Terence Tao (Oct 27 2024 at 01:55):

We don't have much control on E at the end

Mario Carneiro (Oct 27 2024 at 01:56):

It seems likely that E will be the Path relation of a weak central groupoid in that case

Mario Carneiro (Oct 27 2024 at 01:58):

which is to say, we could define it as ∃ z, x = z ◇ y

Terence Tao (Oct 27 2024 at 01:58):

Yes - the one determined by the operation corresponding to E. As long as you managed to add Claim 2 and Claim 3 to your task lists, this will be the case.

Terence Tao (Oct 27 2024 at 01:58):

Otherwise you might potentially have some "junk" edges in E that are not part of the Path

Mario Carneiro (Oct 27 2024 at 01:59):

oh interesting

Terence Tao (Oct 27 2024 at 01:59):

Though maybe with some additional argument you could show they don't exist (now that you have the 1485 law you can play a few more games)

Terence Tao (Oct 27 2024 at 01:59):

But this is just an intellectual curiosity, I don't think understanding the final state of E is needed to get the anti-implication.

Mario Carneiro (Oct 27 2024 at 02:00):

fair enough

Terence Tao (Oct 27 2024 at 02:02):

Anyway what I realize now is that this approach is simply using a new type of ansatz. In previous approaches we have considered a linear ansatz x ◇ y = a * x + b * y and a translation-invariant ansatz x ◇ y = x * f ( x⁻¹ * y ) which helped simplify certain implication problems significantly. This new ansatz c ∈ a ◇ b ↔ IsGood a.1 c.1 b.1 ∧ (a, c) ∈ E ∧ (c, b) ∈ E is just one more ansatz, particularly well suited for central groupoid type laws.

Matthew Bolan (Oct 27 2024 at 03:37):

I have replicated Andy's finding that no graphs with 9 or fewer vertices falsify 151. I'll try and run higher n. It also found a very small graph apparently falsifying 1483,
with adjacency matrix [[1, 1, 1], [1, 1, 1], [1, 1, 0]], and bad paths {(1, 0, 1), (1, 1, 0), (0, 1, 0), (0, 0, 0), (1, 0, 0), (0, 0, 1), (1, 1, 1), (0, 1, 1)}.

Kevin M (Oct 27 2024 at 03:49):

The squares seem key.

These follow from eqn1485

    (x = (y*y)) -> ( x = (x*x)*(x*x) ).   % square elements satisfy eqn151
    (x = (y*y)) -> ( x*x = x*((x*x)*x) ). % square elements satisfy eqn3456
    (x = (y*y)) -> ( x*x = (x*(x*x))*x ). % square elements satisfy eqn3862

And it appears that an idempotent element cannot 'share' its square.
(Is this 'obvious' for some reason?)

    (x = (x*x) & x = (y*y)) -> (x=y).

So non-square elements need to square to elements that cycle between two values when squared.

Andy Jiang (Oct 27 2024 at 03:53):

I also checked 10 or 11 vertices without luck (again assuming code is not buggy)

Terence Tao (Oct 27 2024 at 03:53):

If one can show that squaring is injective, then on finite 1485 magmas it is surjective, hence by what you just wrote, 151 holds. So now it has the full set (F) (T) (L) (C) of immunities: finite magmas, translation-invariant magmas, linear magmas, and commutative magmas.

Is this squaring injectivity statement something you got from an ATP? I don't see how to prove it.

Matthew Bolan (Oct 27 2024 at 03:55):

That statement assumes x = (x*x), so isn't just injectivity

Terence Tao (Oct 27 2024 at 03:56):

Oh I see, I missed that.

Mario Carneiro (Oct 27 2024 at 03:56):

I managed to prove that the greedy construction produces a weak central groupoid satisfying 1485 :tada:

Michael Bucko (Oct 27 2024 at 04:20):

Mario Carneiro schrieb:

I managed to prove that the greedy construction produces a weak central groupoid satisfying 1485 :tada:

Could you please share the code?
Did you only need compactness + the greedy construction? Did you use directed graphs for modeling the extension process?
How special is 1485? Can we identify more laws that allow the magma operation to be determined by the graph homomorphism condition?

Mario Carneiro (Oct 27 2024 at 04:21):

I pushed it to the branch (equational#738)

Mario Carneiro (Oct 27 2024 at 04:21):

oh whoops it seems it was merged already

Mario Carneiro (Oct 27 2024 at 04:23):

https://github.com/teorth/equational_theories/compare/main...digama0:equational_theories:wcg

Mario Carneiro (Oct 27 2024 at 04:24):

The :tada: is maybe slightly premature, because this isn't the end of the proof

Mario Carneiro (Oct 27 2024 at 04:24):

I still have to show the existence of a suitable seed extension

Andy Jiang (Oct 27 2024 at 05:23):

Terence Tao said:

If one can show that squaring is injective, then on finite 1485 magmas it is surjective, hence by what you just wrote, 151 holds. So now it has the full set (F) (T) (L) (C) of immunities: finite magmas, translation-invariant magmas, linear magmas, and commutative magmas.

Is this squaring injectivity statement something you got from an ATP? I don't see how to prove it.

I tried to look for a weaker counterexample by not requiring every edge to be the start or end of a good path (but still requiring good paths between 2 points + 5-cycle condition) and still couldn't find any examples (n<=8) where there is a point which is not the midpoint of a good path with start=end. Not sure what to make of this though.

Matthew Bolan (Oct 27 2024 at 05:33):

I tried that too, and also found no examples (n<=9). I did find a lot of examples when I started messing with axiom 1 instead, for example it finds a plethora of examples if you don't require good paths beginning / ending at the nth vertex. Perhaps this is for trivial reasons but I thought I was seeing some kind of pattern in the output.

Kevin M (Oct 27 2024 at 05:34):

It may help to build up more details to aid the search.
From my last post we know any counterexample for eqn1485 -> eqn151, or eqn3456, or eqn3862 must have an element 0,1,2 such that:
(0 ◇ 0) = 2
(1 ◇ 1) = 2
(2 ◇ 2) = 1

I also just noticed that while non-squares do not satisfy eqn151 x = (x ◇ x) ◇ (x ◇ x),
they do satisfy the weaker: x ◇ x = ((x ◇ x) ◇ (x ◇ x)) ◇ x and x ◇ x = x ◇ ((x ◇ x) ◇ (x ◇ x))
And therefore we also have:
2 = (1 ◇ 0)
2 = (0 ◇ 1)

Hopefully that will help searches.

Matthew Bolan (Oct 27 2024 at 05:40):

Matthew Bolan said:

I tried that too, and also found no examples (n<=9). I did find a lot of examples when I started messing with axiom 1 instead, for example it finds a plethora of examples if you don't require good paths beginning / ending at the nth vertex. Perhaps this is for trivial reasons but I thought I was seeing some kind of pattern in the output.

Tomorrow I might check to see if the examples I found this way are "coming from a greedy algorithm" somehow, but I am not hopeful. There are quite a few bad paths involving the final vertex, suggesting the program is cheating by hiding all the difficulty there. Perhaps the correct relaxation of this condition is to also require no bad paths involving the unpaired vertex.

Matthew Bolan (Oct 27 2024 at 05:56):

Ok, it still finds such things, so maybe there's a modicum of hope

Andy Jiang (Oct 27 2024 at 05:57):

Matthew Bolan said:

I tried that too, and also found no examples (n<=9). I did find a lot of examples when I started messing with axiom 1 instead, for example it finds a plethora of examples if you don't require good paths beginning / ending at the nth vertex. Perhaps this is for trivial reasons but I thought I was seeing some kind of pattern in the output.

~~I can't even find finite examples of a predicate on triples good(a,b,c)
such that
1) for all a,b, exist c such that good(a,c,b)
2) for all a,b,c,d,e good(a,b,c)+good(b,c,d)+good(d,e,a) => good(c,d,e)
2') for all a,b,c,d,e good(e,a,b)+good(b,c,d)+good(d,e,a) => good(c,d,e)
3) exists a such that for all b : not good(b,a,b)~~

Edit: This is wrong (I had a bug in the previous search so there's lots of examples here)

Andy Jiang (Oct 27 2024 at 05:59):

(deleted)

Matthew Bolan (Oct 27 2024 at 06:01):

Here are the adjacency matrices/bad paths of some such things: output.txt. I added the modified condition somewhat hastily so perhaps there was a bug.

Matthew Bolan said:

Ok, it still finds such things, so maybe there's a modicum of hope

Matthew Bolan (Oct 27 2024 at 06:02):

I'm going to bed. Hopefully one of my overnight runs finds something (Andy if you are doing a large run perhaps we should make sure we aren't overlapping our searches).

Andy Jiang (Oct 27 2024 at 06:04):

Matthew Bolan said:

I'm going to bed. Hopefully one of my overnight runs finds something (Andy if you are doing a large run perhaps we should make sure we aren't overlapping our searches).

I have a search going for size 12 for relaxed weak central groupoid which is anti-3456.

Andy Jiang (Oct 27 2024 at 06:06):

[Edit : some wrong statement was here and removed]

Matthew Bolan (Oct 27 2024 at 06:11):

I have 4 searches for relaxed weak central groupoids that are anti-151/3456 (which are equivalent) running now, for 10,11,12,13. Since my CPU can take it I guess I'll keep the redundant ones going.

Andy Jiang (Oct 27 2024 at 06:33):

Matthew Bolan said:

I have 4 searches for relaxed weak central groupoids that are anti-151/3456 (which are equivalent) running now, for 10,11,12,13. Since my CPU can take it I guess I'll keep the redundant ones going.

that's fine I think, I don't trust my coding that much anyway.

Mario Carneiro (Oct 27 2024 at 07:57):

okay now we can celebrate: equational#744

theorem not_3457 :  (G : Type) (_ : Magma G), Facts G [1485] [3457] := ...
theorem not_3511 :  (G : Type) (_ : Magma G), Facts G [1485] [3511] := ...
theorem not_2087_2124 :  (G : Type) (_ : Magma G), Facts G [1485] [2087, 2124] := ...

Mario Carneiro (Oct 27 2024 at 08:02):

One nice thing about having the formalization is that I can just comment out parts and see if anything breaks: turns out conditions isGood_left and isGood_right are never used

class RelaxedWeakCentralGroupoid (G : Type*) extends Magma G where
  Path : G  G  Prop
  IsGood : G  G  G  Prop
  op_isGood (x y : G) : IsGood x (x  y) y
  isGood_path {x y z : G} : IsGood x y z  Path x y  Path y z
  -- isGood_left {x y : G} : Path x y → ∃ z, IsGood z x y
  -- isGood_right {x y : G} : Path x y → ∃ z, IsGood x y z
  isGood_five {a b c d e : G} : IsGood a b c  IsGood c d e  Path e a  IsGood b c d

Mario Carneiro (Oct 27 2024 at 08:03):

that is to say, claims 2 and 3

Andy Jiang (Oct 27 2024 at 08:07):

Mario Carneiro said:

One nice thing about having the formalization is that I can just comment out parts and see if anything breaks: turns out conditions isGood_left and isGood_right are never used

class RelaxedWeakCentralGroupoid (G : Type*) extends Magma G where
  Path : G  G  Prop
  IsGood : G  G  G  Prop
  op_isGood (x y : G) : IsGood x (x  y) y
  isGood_path {x y z : G} : IsGood x y z  Path x y  Path y z
  -- isGood_left {x y : G} : Path x y → ∃ z, IsGood z x y
  -- isGood_right {x y : G} : Path x y → ∃ z, IsGood x y z
  isGood_five {a b c d e : G} : IsGood a b c  IsGood c d e  Path e a  IsGood b c d

huh this also tracks with Matthew Bolan and my observation that it's hard to find finite counterexamples against 151/3456 even throwing away those axioms

Kevin M (Oct 27 2024 at 08:13):

Found more conditions that hopefully will aid search.

Any counterexample for eqn1485 -> eqn151, or eqn3456, or eqn3862 must have distinct elements 0,1,2 such that:
(summarizing previous results)
(0 ◇ 0) = 2, (0 ◇ 1) = 2, (1 ◇ 0) = 2, (1 ◇ 1) = 2, (2 ◇ 2) = 1,

and new restrictions found with prover9
there must exist at least other distinct elements 3,4
(0 ◇ 2) = 3
(2 ◇ 0) = 4
and
for all x (0 ◇ x) != 1 and (x ◇ 0) != 1
for all x (1 ◇ x) != 0 and (x ◇ 1) != 0

Andy Jiang (Oct 27 2024 at 08:26):

Kevin M said:

Found more conditions that hopefully will aid search.

Any counterexample for eqn1485 -> eqn151, or eqn3456, or eqn3862 must have distinct elements 0,1,2 such that:
(summarizing previous results)
(0 ◇ 0) = 2, (0 ◇ 1) = 2, (1 ◇ 0) = 2, (1 ◇ 1) = 2, (2 ◇ 2) = 1,

and new restrictions found with prover9
there must exist at least other distinct elements 3,4
(0 ◇ 2) = 3
(2 ◇ 0) = 4
and
for all x (0 ◇ x) != 1 and (x ◇ 0) != 1
for all x (1 ◇ x) != 0 and (x ◇ 1) != 0

are the last two conditions coming from the impossibility of edges between 0 and 1?

Kevin M (Oct 27 2024 at 08:37):

Oh, Is there a way to show that with the 5-cycle rule? I haven't translated it into the graph language to check.

Honestly, I've been exploring the consequences of a non-square existing in an eqn1485 magma by treating the ATP like playing a long game of "twenty questions" where the person can only answer True or Unknown. Making many quick guesses and checking. There's probably a better way to do this, but this is not very mentally taxing and I found it fun.

Andy Jiang (Oct 27 2024 at 08:41):

Kevin M said:

Oh, Is there a way to show that with the 5-cycle rule? I haven't translated it into the graph language to check.

Honestly, I've been exploring the consequences of a non-square existing in an eqn1485 magma by treating the ATP like playing a long game of "twenty questions" where the person can only answer True or Unknown. Making many quick guesses and checking. There's probably a better way to do this, but this is not very mentally taxing and I found it fun.

0->2->0->2->1 prevents a 1->0 path I think
and
1->2->0->2->0 prevents a 0->1 path

seems like the graph interpretation is quite powerful at proving these sort of statements

Andy Jiang (Oct 27 2024 at 08:54):

Mario Carneiro said:

One nice thing about having the formalization is that I can just comment out parts and see if anything breaks: turns out conditions isGood_left and isGood_right are never used

class RelaxedWeakCentralGroupoid (G : Type*) extends Magma G where
  Path : G  G  Prop
  IsGood : G  G  G  Prop
  op_isGood (x y : G) : IsGood x (x  y) y
  isGood_path {x y z : G} : IsGood x y z  Path x y  Path y z
  -- isGood_left {x y : G} : Path x y → ∃ z, IsGood z x y
  -- isGood_right {x y : G} : Path x y → ∃ z, IsGood x y z
  isGood_five {a b c d e : G} : IsGood a b c  IsGood c d e  Path e a  IsGood b c d

It could be that a predicate good(a,b,c) satisfying
1) forall a,b, exists c: good(a,c,b)
2) forall a,b,c,d,e,f: good(a,b,f)+good(b,c,d)+good(d,e,a) => good(c,d,e)

formally implies that
(exists e: good(a,b,e)) <=> (exists e: good(e,a,b))

then you could define the edges directly from the good predicate?

Mario Carneiro (Oct 27 2024 at 08:56):

My guess is that it doesn't, but @Terence Tao has a better intuition about that

Kevin M (Oct 27 2024 at 09:10):

It was a long shot, but using prover9 I was hoping to find some implication where the existence of one non-square implies other non-squares, which then need to square to distinct elements (or at least interact with other elements to somehow require new distinct elements), and so on onto infinity ... ruling out finite counter-examples.

But even in the simplest cases I've yet to find any hints about other non-squares. For instance, given a non-square x, is x◇(x◇x) square or non-square? what about the more general case x◇(y◇y) ?
The only case I know is this non-square and square result in a square: x◇((x◇x)◇(x◇x)) = (x◇x).
Maybe someone can try it out with a different ATP or even just different ways of writing the statement to prover9. (I also may have just not hit on the right guesses and thus questions to ask it.)

Andy Jiang (Oct 27 2024 at 09:10):

@Mario Carneiro I asked vampire about this question (whether the conditions on the good(a,b,c) in my comment above imply that we can define edges to satisfy the old axiom 2) and it seems to think it's true, but I haven't checked if I made any mistakes in translation (also I used vampire first time yesterday so...) Also I should probably try to understand the proof....
output1
output2
reducedaxiom1.p
reducedaxiom2.p

Kevin M (Oct 27 2024 at 09:22):

In reducedaxiom1.p the five cycle rule is

% 4. Five-cycle condition:
fof(five_cycle_good_paths, axiom,
    ! [A, B, C, D, E, F] : (
        (good(A,B,F) & good(B, C, D) & good(D, E, A)) => good(C, D, E)
    )
).

Can't the A->B edge also come from the later half of a good path?
So maybe this?

        ((good(A,B,F) | good(F,A,B)) & good(B, C, D) & good(D, E, A)) => good(C, D, E)

Andy Jiang (Oct 27 2024 at 09:30):

I want to say that the above (vampire) proof shows that assuming the axioms I have, good(F,A,B) for some F implies good(A,B,G) for some F

Kevin M (Oct 27 2024 at 09:33):

Ah yes. I missed some of the context, and rereading above it make sense. Sorry about that.

Andy Jiang (Oct 27 2024 at 09:37):

No worries! Let me know if you understand what's going on in the computer proof by the way (or have an alternative proof)

Andy Jiang (Oct 27 2024 at 09:38):

It's funny it reminds me of the octahedral axiom of triangulated categories...

Andy Jiang (Oct 27 2024 at 09:40):

(deleted)

Andy Jiang (Oct 27 2024 at 09:40):

But probably not related at all

Andy Jiang (Oct 27 2024 at 12:12):

Andy Jiang said:

Mario Carneiro said:

One nice thing about having the formalization is that I can just comment out parts and see if anything breaks: turns out conditions isGood_left and isGood_right are never used

class RelaxedWeakCentralGroupoid (G : Type*) extends Magma G where
  Path : G  G  Prop
  IsGood : G  G  G  Prop
  op_isGood (x y : G) : IsGood x (x  y) y
  isGood_path {x y z : G} : IsGood x y z  Path x y  Path y z
  -- isGood_left {x y : G} : Path x y → ∃ z, IsGood z x y
  -- isGood_right {x y : G} : Path x y → ∃ z, IsGood x y z
  isGood_five {a b c d e : G} : IsGood a b c  IsGood c d e  Path e a  IsGood b c d

It could be that a predicate good(a,b,c) satisfying
1) forall a,b, exists c: good(a,c,b)
2) forall a,b,c,d,e,f: good(a,b,f)+good(b,c,d)+good(d,e,a) => good(c,d,e)

formally implies that
(exists e: good(a,b,e)) <=> (exists e: good(e,a,b))

then you could define the edges directly from the good predicate?

I guess actually writing the relaxed central groupoid in this way is like going back to the defining equation of 1485 lol except that you drop the uniqueness of c

Terence Tao (Oct 27 2024 at 14:51):

It is slightly surprising to me that one can get the (exists e: good(a,b,e)) <=> (exists e: good(e,a,b)) axiom "for free"; I tried chasing 5-cycles around a bit to prove this but could not see how, so I'll be interested in seeing some human-readable interpretation of the ATP proofs. In general it seems that the ATP proofs here are actually a bit non-trivial. I have also had trouble verifying Kevin's ATP-generated claim that all squares obey 151; again drawing the graph and staring at 5-cycles didn't seem to help in that case. Perhaps getting a better human understanding of these sorts of relations may provide some further clues (or perhaps some further axioms to help the ATP focus its search).

Terence Tao (Oct 27 2024 at 15:22):

One thought : all the successful examples of relaxed central groupoids had some sort of symmetry in their graph. Andy's first example had a 3-fold rotational symmetry in which one cyclically permuted 0,1,2 and also 3,5,4.
Andy's second example had a vertical reflection symmetry in that if one swapped 0 and 1, and also 3 and 4, one obtained the transpose of the original graph. Matthew's tiny example is self-transpose, and also has a reflection symmetry under swapping 0 and 1.

Perhaps if we impose a symmetry condition on the magma it will make it easier to find? I'm thinking of a reflection symmetry where the key values 0, 1, 2 of the counterexample are on the line of symmetry. Perhaps as a guide for what kind of symmetry to have, consider the natural central groupoid on the four-element {0,1}×{0,1} \{0,1\} \times \{0,1\} with law abcd=bcab \diamond cd = bc. Of course it doesn't violate 151, but we do have a (partial) non-injective homomorphism 0010 \mapsto 01, 1011 \mapsto 01, 2102 \mapsto 10, 3113 \mapsto 11, 4004 \mapsto 00 from Kevin's relations to this natural central groupoid. This suggests perhaps a reflection symmetry in that if one keeps 0,1,2 fixed but swaps 3 and 4, then one should obtain the reversed groupoid (with the transposed graph). One could then try to organize the other elements of the proposed relaxed weak central groupoid either in points on the line of symmetry, or pairs off of the line, and design the multiplication table accordingly by adding reflection constraints (in graph language: aba \to b iff bab^* \to a^*, where aa^* denotes the reflection of aa; in operator language, ab=ca \diamond b = c iff ba=cb^* \diamond a^* = c^*).

Matthew Bolan (Oct 27 2024 at 15:28):

There are probably quite a few more relaxed central groupoids. If our idea is to look for patterns in "interesting" relaxed central groupoids I can try and generate some more, though I'd need to know what we mean by interesting.

Terence Tao (Oct 27 2024 at 15:32):

I guess the simplest thing is to impose a transposition symmetry as I suggested above. So, in addition to the ternary relation RR, one has a unary operator * obeying the involution axiom (x)=x(x^*)^* = x and the reflection axiom R(x,y,z)    R(y,x,z)R(x,y,z) \iff R(y^*,x^*,z^*). One can also impose 0=0,1=1,2=2,3=40^*=0, 1^*=1, 2^*=2, 3^*=4 on the constants in Kevin's laws. My hope is that when dealing with say order 12 examples, this extra symmetry may cut down the number of possibilities to the point where an ATP can actually find interesting examples.

EDIT: Hang on, this may not be the right axiom system: the natural central groupoids don't have this symmetry. (Or rather, they do, but they don't transform 0,1,2,3,4 the way I expected.) I retract the above proposal, need to think about it more.

Matthew Bolan (Oct 27 2024 at 15:39):

I started generating relaxed central groupoids with my definition of "interesting" being "contains a bad path", and it found what looks like a few trivial infinite families of them, where the underlying graph is complete and the bad paths seem to be the paths containing some edge. Useless for 151 of course.

Terence Tao (Oct 27 2024 at 15:39):

Hmm, on comparing with the symmetry (ab)=ba(ab)^* = ba for natural central groupoids, it seems that instead of 0=0,1=1,2=2,3=40^* = 0, 1^* = 1, 2^* = 2, 3^* = 4, perhaps the axiom to impose is 1=21^* = 2. Among other things, this means that 00 (which squares to 2) acquires a reflection 00^* (which squares to 1).

In any case, relaxed natural central groupoids with such a reflection symmetry seem worth looking at as a special subclass.

Terence Tao (Oct 27 2024 at 15:40):

Yeah, this is what happens when you drop Claim 2 and Claim 3, you allow in the trivial situation in which there are some "useless" edges in the graph that only generate bad paths and no good paths.

Matthew Bolan (Oct 27 2024 at 15:41):

ah gotcha. So even though Mario showed axioms 2 and 3 are unneeded, I should probably include them in a search for interesting examples?

Terence Tao (Oct 27 2024 at 15:58):

Yeah; you can always prune those bad edges out afterwards as well, but perhaps it's less distracting for the ATP if you remove them right away.

By the way, I find it a little bit helpful to think of these graphs as road systems (with a lot of one-way streets), where the good paths correspond to "legal turns" and bad paths to "illegal turns". Claim 1' is then the claim that one can get from any point to any other point in two steps with one legal turn. Claims 2, 3 are the claims that you can't get stuck in a dead end where there are no legal options out, and conversely you can always legally enter any given street somehow. Claim 4 is somewhat bizarre in this traffic analogy though.

Matthew Bolan (Oct 27 2024 at 15:59):

Ha, Claim 4 is some obscure roundabout law

Terence Tao (Oct 27 2024 at 16:06):

Yeah it says that in any 5-segment roundabout, the portion of the roundabout which can be legally used must always be connected, and if four of the turns are legal, then so must the fifth. In other words, you can't shut down one turn of the roundabout without also shutting down at least one adjacent turn.

Alex Meiburg (Oct 27 2024 at 21:25):

A roundabout where you're not allowed to go in a full circle is a pretty bad roundabout! :laughter_tears:

thormikkel (Oct 27 2024 at 21:36):

Another way one could refute 151 is to find a relaxed weak central groupoid with the property that there exists some w such that a - w - a is good and b - w - b is good for distinct a,b. This might be harder than the a - c - a good c - a - c bad property but I'm just throwing it out there.

Matthew Bolan (Oct 27 2024 at 21:38):

My current searches include many of Kevin's relations, which contain requirements of the type you mention

Andy Jiang (Oct 27 2024 at 21:52):

thormikkel said:

Another way one could refute 151 is to find a relaxed weak central groupoid with the property that there exists some w such that a - w - a is good and b - w - b is good for distinct a,b. This might be harder than the a - c - a good c - a - c bad property but I'm just throwing it out there.

sorry can't this happen without failing 151? for relaxed weak (instead of just weak) I thought you can have noninjectivity of these pseudosquares--- for example you can take a complete graph and call all paths good. I think nonsurjectivity of squares would be a novel example though for finite graphs (Btw is there a known nonsurjective squaring infinite relaxed weak central groupoid? Maybe it can be constructed by this greedy extension thing)

Matthew Bolan (Oct 27 2024 at 21:54):

If squaring is not surjective, then 151 is false, after all it is the statement that x=(xx)(xx)x = (x \diamond x) \diamond (x \diamond x)

Matthew Bolan (Oct 27 2024 at 21:55):

Oh, I missed that you were asking about the relaxed case

Matthew Bolan (Oct 27 2024 at 21:57):

That would still yield a counterexample to 151 IIUC, so I assume we don't know one.

thormikkel (Oct 27 2024 at 21:58):

Hmm ok, maybe I'm misunderstanding the procedure, but from reading the blueprint I thought you could extend this object by greedy extension to a magma that satisfies 1485 and thus refutes 151, due to the square map being non-injective by construction?

Matthew Bolan (Oct 27 2024 at 22:02):

For any xx, the greedy extension at some point needs to add a good path xyxx \to y \to x, where π(x)π(y)π(x)\pi(x) \to \pi(y) \to \pi(x) is good in the base. Then for xx to not be a square, the path π(y)π(x)π(y)\pi(y) \to \pi(x) \to \pi(y) needs to be bad in the base, as otherwise yxyy\to x \to y is good.

Andy Jiang (Oct 27 2024 at 22:04):

Matthew Bolan said:

That would still yield a counterexample to 151 IIUC, so I assume we don't know one.

ah I see, like if it's not surjective on base it's not surjective on any weak central groupoids you can make out of it either. Thanks!

Matthew Bolan (Oct 27 2024 at 22:06):

It does feel very tempting to try and make an infinite such thing greedily though

Andy Jiang (Oct 27 2024 at 22:06):

thormikkel said:

Hmm ok, maybe I'm misunderstanding the procedure, but from reading the blueprint I thought you could extend this object by greedy extension to a magma that satisfies 1485 and thus refutes 151, due to the square map being non-injective by construction?

are you saying you can start with a seed where the squaring is not injective?

Terence Tao (Oct 27 2024 at 22:13):

Unfortunately you can't make squaring non-injective on a seed unless the base already has a pair a,b with a → b → a good and b → a → b bad, because the seed can't have two good paths with the same endpoints

Andy Jiang (Oct 27 2024 at 22:26):

Terence Tao said:

It is slightly surprising to me that one can get the (exists e: good(a,b,e)) <=> (exists e: good(e,a,b)) axiom "for free"; I tried chasing 5-cycles around a bit to prove this but could not see how, so I'll be interested in seeing some human-readable interpretation of the ATP proofs. In general it seems that the ATP proofs here are actually a bit non-trivial. I have also had trouble verifying Kevin's ATP-generated claim that all squares obey 151; again drawing the graph and staring at 5-cycles didn't seem to help in that case. Perhaps getting a better human understanding of these sorts of relations may provide some further clues (or perhaps some further axioms to help the ATP focus its search).

Let me try to write down the vampire proof in a human-way.
1) suppose 0->1->2 is good, we want to show that 0->1 is the second leg of a good path
2) take a such that 2->a is the first leg of a good path
3) take b such that a->b->0 is good
4) Now b->0->1 is good because 2->a->? is good, a->b->0 is good, and 0->1->2 is good

Andy Jiang (Oct 27 2024 at 22:35):

Now showing that 1->2 is the first leg of some good path assuming good(0->1->2):
1) Take x from above such that x->0->1 is good
2) Take y such that 2->y->x is good
3) Then 1->2->y is good because 0->1->2 good, 2->y->x good and x->0->1 good

Matthew Bolan (Oct 28 2024 at 00:08):

I was thinking about trying to tweak the greedy algorithm to select whether a path is good or bad more carefully. Could we maybe gain more wiggle room by allowing the seed graph to be a multigraph instead? That way good paths/bad paths have a chance to "remember where they came from", based on which particular edges between the relevant vertices they involve, and it seems to me like the same greedy algorithm should go through.

Matthew Bolan (Oct 28 2024 at 00:09):

Maybe it gains nothing though, as the 5-cycle axiom doesn't see the multiple edges

Terence Tao (Oct 28 2024 at 00:12):

There could be a little bit of play in being a bit more intelligent in the greedy construction, though like you I suspect that the most likely outcome is that it gives something equally powerful as the original technique. Can't hurt to try, I guess. (I also briefly thought about some iterative approach where one uses one relaxed weak central groupoid to greedily generate another relaxed weak central groupoid, which one then turns into an actual weak central groupoid, but again this may be a more complicated version of the argument that doesn't actually add any power.)

Bhavik Mehta (Oct 28 2024 at 01:49):

This is very cool! Given the prevalence of idempotence and squares in some of the above discussions, I wanted to throw another question into the mix. There are a couple of different proofs that a finite central groupoid has a square number of elements, some combinatorial/graph theoretic, and others linear algebraic. But there seems to be only one known proof (eg, the one in Knuth's note) that a central groupoid with n^2 elements has n idempotents, which uses spectral decomposition of matrices (in fact, a case which mathlib doesn't yet seem to know directly).
Can we use some of the above ideas to prove the weaker claim that a nonempty finite central groupoid has at least one idempotent? Certainly this is true, but is there a graph theory way of seeing why it is true?

Andy Jiang (Oct 28 2024 at 02:11):

Andy Jiang said:

Now showing that 1->2 is the first leg of some good path assuming good(0->1->2):
1) Take x from above such that x->0->1 is good
2) Take y such that 2->y->x is good
3) Then 1->2->y is good because 0->1->2 good, 2->y->x good and x->0->1 good

By the way I think a direct translation of the proof in the weak (non relaxed) probably works too (if not identical to this proof)

Terence Tao (Oct 28 2024 at 02:11):

This is closely related to the question of whether there can be a subset AA of a finite group GG such that the products AAA \cdot A are all distinct and fill out GG. @Andy Jiang showed earlier in this thread that this is not possible, by basically the same matrix-based argument of Knuth; we don't have an alternate proof of this result either.

Matthew Bolan (Oct 28 2024 at 02:25):

Matthew Bolan said:

I was thinking about trying to tweak the greedy algorithm to select whether a path is good or bad more carefully. Could we maybe gain more wiggle room by allowing the seed graph to be a multigraph instead? That way good paths/bad paths have a chance to "remember where they came from", based on which particular edges between the relevant vertices they involve, and it seems to me like the same greedy algorithm should go through.

Probably fruitless but while waiting for some other searches to finish I started a search for multigraph seeds obeying the following:

Claim 1''': For any two vertices a,b, there exists a third vertex c, such that at least one of the paths a → b → cis good.
Claim 4: In a five cycle a → b → c → d → e → a, if a → b → c and c → d → e are good, then so is b → c → d (Where all edges in the good paths are the ones from the 5-cycle).
Anti-151: There are edges A:0→1 B:1→0 with AB good and BA bad.

Hopefully if I've done things right, the greedy algorithm should still go through given such a base (though you now owe a selection of edges in the base multigraph each step). Initially I thought this didn't have a chance to improve on just regular graphs as the seed, since you can just select a single edge between each pair of vertices and get a graph which seems close to the ones we have been searching for already. However, I don't see why some graph produced that way and containing the offending edges from Anti-151 should satisfy Claim 1', so maybe there's an outside chance.

Matthew Bolan (Oct 28 2024 at 14:05):

Terence Tao said:

Hmm, on comparing with the symmetry (ab)=ba(ab)^* = ba for natural central groupoids, it seems that instead of 0=0,1=1,2=2,3=40^* = 0, 1^* = 1, 2^* = 2, 3^* = 4, perhaps the axiom to impose is 1=21^* = 2. Among other things, this means that 00 (which squares to 2) acquires a reflection 00^* (which squares to 1).

In any case, relaxed natural central groupoids with such a reflection symmetry seem worth looking at as a special subclass.

Imposing this sort of relation seemed to speed up the search by quite a bit, but I suspect because there's some reason why it can't work that it rediscovers quickly each time I add a new vertex. Maybe someone better with ATP can check if there's an obstruction.

Matthew Bolan (Oct 28 2024 at 14:15):

I can't even report negative results today really, the searches are starting to exceed what my machine can do in just a few days. The only search which is turning up anything is one for other 166 examples (of which there are many).

Matthew Bolan (Oct 28 2024 at 14:18):

Here is a 17 vertex anti-166 example, where I thought the adjacency matrix looked interesting.
[1 1 1 0 0 0 1 0 0 0 0 1 0 1 0 1 1]
[1 1 1 1 1 1 0 1 1 1 1 0 1 1 0 0 0]
[0 1 1 0 0 0 1 0 0 0 0 1 0 1 0 1 1]
[0 0 1 0 0 0 0 0 0 0 0 0 0 0 1 1 0]
[1 0 1 0 0 0 0 0 0 0 0 0 0 0 1 1 0]
[1 0 1 0 0 0 0 0 0 0 0 0 0 0 1 1 0]
[0 1 0 0 0 0 0 0 0 0 0 0 0 1 1 0 0]
[1 0 1 0 0 0 0 0 0 0 0 0 0 0 1 0 0]
[1 0 0 0 0 0 0 0 0 0 0 0 0 0 1 0 0]
[1 0 1 0 0 0 0 0 0 0 0 0 0 0 1 1 0]
[1 0 1 0 0 0 0 0 0 0 0 0 0 0 1 1 0]
[0 0 0 0 0 0 0 0 0 0 0 0 0 1 1 0 0]
[1 0 0 0 0 0 0 0 0 0 0 0 0 0 1 1 0]
[1 1 1 0 0 0 0 0 0 0 1 0 1 1 0 1 0]
[0 0 0 1 1 1 1 1 1 1 1 1 1 0 1 0 1]
[1 1 0 0 0 0 1 0 0 0 0 1 0 1 0 0 0]
[0 1 0 0 0 0 0 0 0 0 0 0 0 1 1 0 0]
With the bad paths being
{(13, 1, 2), (15, 0, 13), (1, 1, 13), (2, 2, 2), (13, 13, 13), (15, 13, 1), (1, 2, 2), (1, 13, 2), (15, 0, 1), (15, 13, 13), (0, 0, 1), (0, 2, 1), (0, 13, 1), (15, 0, 0), (1, 1, 0), (0, 0, 13), (13, 0, 0), (0, 15, 1), (13, 1, 1), (0, 2, 13), (0, 13, 13), (13, 13, 0), (15, 0, 15), (0, 15, 13), (13, 15, 0), (13, 1, 13), (13, 0, 15), (13, 13, 15), (2, 2, 1), (2, 13, 1), (13, 2, 15), (2, 15, 1), (15, 1, 1), (2, 2, 13), (2, 13, 13), (1, 13, 1), (1, 13, 13), (2, 15, 13), (15, 1, 13), (0, 0, 0), (0, 15, 0), (1, 1, 2), (0, 0, 15), (1, 0, 0), (13, 0, 2), (15, 0, 2), (0, 2, 15), (13, 2, 2), (13, 1, 0), (13, 13, 2), (1, 0, 15), (0, 1, 1), (2, 1, 1), (2, 15, 0), (2, 1, 13), (14, 14, 14), (1, 13, 0), (0, 1, 13), (2, 2, 15), (1, 2, 15), (0, 0, 2), (1, 13, 15), (0, 2, 2), (1, 0, 2), (1, 1, 1), (13, 13, 1)}

thormikkel (Oct 28 2024 at 15:36):

Since we're kind of hitting a computational roadblock here, is there a more non-constructive approach (ramsey-like) that could work? I.e given a graph of sufficient size and with suitable restrictions, there has to exist a subgraph with the partial weak centroid property and the a-b-a good b-a-b bad condition. I understand that this is very close to abstract nonsense and might be similar to the other extensions proposed here. My concern is that the seed graph generating a counterexample to 151 might be too large to find by computer search.

Andy Jiang (Oct 28 2024 at 16:11):

Bhavik Mehta said:

This is very cool! Given the prevalence of idempotence and squares in some of the above discussions, I wanted to throw another question into the mix. There are a couple of different proofs that a finite central groupoid has a square number of elements, some combinatorial/graph theoretic, and others linear algebraic. But there seems to be only one known proof (eg, the one in Knuth's note) that a central groupoid with n^2 elements has n idempotents, which uses spectral decomposition of matrices (in fact, a case which mathlib doesn't yet seem to know directly).
Can we use some of the above ideas to prove the weaker claim that a nonempty finite central groupoid has at least one idempotent? Certainly this is true, but is there a graph theory way of seeing why it is true?

It's an interesting question. I think that once you relax the central groupoid condition you no longer need to have idempotents necessarily (suitably interpreted). This I think also means that by Terry's construction of models from relaxed models, in infinite models idempotents don't have to exist in central groupoids.

Now I think for weak central groupoids even the 2-element model has no idempotent

thormikkel (Oct 28 2024 at 16:14):

The free central groupoid generated by the procedure in Knuth's paper doesn't have any idempotents. But it would be very interesting to find other examples using graph theoretical language!

Andy Jiang (Oct 28 2024 at 16:26):

(deleted)

Andy Jiang (Oct 28 2024 at 16:27):

(deleted)

Matthew Bolan (Oct 28 2024 at 16:29):

My search for 12 is still running too. A search including a version Terry's symmetry ansatz + Kevin's relations has ruled out graphs of size <= 23 as well, and is still fast enough that it can probably push a little higher (and I've checked that the symmetry ansatz still allows counterexamples to all other known laws).

Matthew Bolan (Oct 28 2024 at 16:30):

Still, that level of speed-up is very suspicious to me, it makes me expect that I've asked for something that is "trivially" unsatisfiable somehow.

Terence Tao (Oct 28 2024 at 16:31):

I do agree that the relaxation of the problem from "find an actual weak central groupoid contradicting 151" to "find a relaxed weak central groupoid contradicting 151", while presumably helpful, doesn't seem to be enough to bring the problem into computationally feasible range, and a second relaxation (or some imposition of additional structure to cut down the combinatorial explosion) is needed.

One vague thought I had was to take an existing (relaxed) weak central groupoid and try to extend it by adding a few more elements, in the spirit of what we tried for 854. In the natural central groupoid case S×SS \times S one can extend it by adding a new element 0000, as well as 0s0s and s0s0 for each idempotent ssss in the natural central groupoid, and enlarge the multiplication table by imposing rules such as 0sx=ssx0s \diamond x = ss \diamond x for each xx in the original natural central groupoid (the full set of rules is a bit complicated though). One could try starting with some weak central groupoid in which squaring is a bijection (perhaps even a natural central groupoid to make all calculations explicit), and try to adjoin an additional "square root" of an existing element to defeat 151, and see what other elements have to be added to the system to make everything consistent (and permitting the weak central groupoid to now be relaxed to make our lives as easy as possible).

Terence Tao (Oct 28 2024 at 16:32):

Matthew Bolan said:

Still, that level of speed-up is very suspicious to me, it makes me expect that I've asked for something that is "trivially" unsatisfiable somehow.

One could run an ATP to see if there is indeed some trivial proof that 1485 + symmetry implies 151. If so, that proof could be enlightening, as it would indicate a specific "symmetry breaking" we need to introduce to solve the problem.

Matthew Bolan (Oct 28 2024 at 18:39):

My general search for 12 vertex graphs that I started about a day and a half ago (no ansatzes) just finished - no solutions.

Matthew Bolan (Oct 28 2024 at 23:10):

Memory error killed the remaining searches :upside_down:

Matthew Bolan (Oct 28 2024 at 23:11):

Oh well, it was looking quite grim anyway, this one definitely needs a new idea.

Andy Jiang (Oct 29 2024 at 03:02):

I'm wondering if one can somehow give an explicit model of the free 1485 magma on one generator. Like maybe we look at all binary trees with at most 10 vertices, ask vampire if two trees are obviously the same in the free magma, and otherwise assume for now (lol) that they are different and see if we can observe some patterns in the equivalence classes of trees? not sure if this can gain any ground.

Terence Tao (Oct 29 2024 at 03:08):

I think this is worth a shot. We could use small examples of weak central groupoids to do a preliminary classification: every tree gives a word map on (for instance) order 2 weak central groupoid which divides all trees into four classes depending on what the word map does (e.g., does it send 0 -> 0 and 1 -> 0, 0 -> 0 and 1 -> 1, etc.)

Matthew Bolan (Oct 29 2024 at 03:10):

168 is confluent right? Does that mean we have an explicit model of the free 168 magma on one generator I can look at?

Matthew Bolan (Oct 29 2024 at 03:24):

Ok it's what I expected (and in the Knuth paper already).

Terence Tao (Oct 29 2024 at 03:50):

So maybe we can initially focus on those trees (i.e., words of one variable x) that are equal to the identity x on the free 168 magma, and also on the order 2 weak central groupoid, and see what relations exist within that class. We know that x and the 151 law (x ◇ x) ◇ (x ◇ x) will live there; what other small trees exist in that class?

Matthew Bolan (Oct 29 2024 at 03:52):

(x ◇ x) ◇ (x ◇ (x ◇ x)) and ((x ◇ x) ◇ x) ◇ (x ◇ x) are the other two in our current database of known implications.

Terence Tao (Oct 29 2024 at 04:45):

And these ones are known to equal x, since 1485 implies both 1426 and 2035.

Andy Jiang (Oct 29 2024 at 06:07):

[There was some bug here so I'm removing in meantime]

Andy Jiang (Oct 29 2024 at 06:26):

I did a trial run of 5 leaves with the following conjectural equivalence classes.
Different numbers are definitely different (by Terry's suggestion of checking central groupoid + size 2 case) and same number+letter is definitely same (by Vampire)
Should warn that
1) high chance of bugs in code
2) vampire was only run for 60s on each pair so it's very possible there's equivalent pairs which are not combined

Class 1a:
A
mul(mul(mul(A,A),A),mul(A,A))
mul(mul(A,A),mul(A,mul(A,A)))

Class 1b:
mul(mul(A,A),mul(A,A))

Class 2:
mul(mul(A,mul(A,A)),mul(A,A))

Class 3:
mul(mul(A,A),mul(mul(A,A),A))

Class 4:
mul(A,mul(A,mul(A,A)))

Class 5a:
mul(mul(A,mul(A,A)),A)

Class 5b:
mul(A,A)
mul(A,mul(mul(A,A),mul(A,A)))
mul(mul(mul(A,A),mul(A,A)),A)

Class 5c:
mul(A,mul(mul(A,A),A))

Class 6:
mul(mul(mul(A,A),A),A)

Class 7:
mul(mul(mul(mul(A,A),A),A),A)

Class 8a:
mul(A,mul(A,A))
mul(A,mul(mul(A,mul(A,A)),A))

Class 8b:
mul(A,mul(A,mul(mul(A,A),A)))

Class 9a:
mul(mul(A,A),A)
mul(mul(A,mul(mul(A,A),A)),A)

Class 9b:
mul(mul(mul(A,mul(A,A)),A),A)

Class 10:
mul(A,mul(A,mul(A,mul(A,A))))

Class 11:
mul(A,mul(mul(mul(A,A),A),A))

Class 12:
mul(mul(A,mul(A,mul(A,A))),A)
Edit: I fixed some bugs but still not sure if this is correct....

Andy Jiang (Oct 29 2024 at 09:23):

Sorry I think my implementation has some bugs, I wouldn't trust the above outputs too much. (I had computed 6,7 leaves but the outputs were definitely wrong (like vampire failed to equate some obviously equal expressions)

thormikkel (Oct 29 2024 at 11:27):

Bhavik Mehta said:

This is very cool! Given the prevalence of idempotence and squares in some of the above discussions, I wanted to throw another question into the mix. There are a couple of different proofs that a finite central groupoid has a square number of elements, some combinatorial/graph theoretic, and others linear algebraic. But there seems to be only one known proof (eg, the one in Knuth's note) that a central groupoid with n^2 elements has n idempotents, which uses spectral decomposition of matrices (in fact, a case which mathlib doesn't yet seem to know directly).
Can we use some of the above ideas to prove the weaker claim that a nonempty finite central groupoid has at least one idempotent? Certainly this is true, but is there a graph theory way of seeing why it is true?

If you're still curious about this I tried to prove this by partitioning the graph into cycles and showing that not allowing cycles with 1 element lead to a contradiction. But before I got anywhere I discovered that someone already proved a stronger result (there are exactly k cycles with 1 element for a k^2 order graph with "Property C") using some variation of the same argument: https://www.sciencedirect.com/science/article/pii/S0097316511000616

Bhavik Mehta (Oct 29 2024 at 12:00):

Ah! Proposition 2 is what I was looking for, the walk counting argument specifically. Thanks!

Terence Tao (Oct 29 2024 at 14:24):

This paper also discusses two interesting ways to convert one central groupoid to another, “switching” and extension. If we can find analagous operations for weak central groupoids it may be much easier to explore rhe space of weak central groupoids to look for 151 counter examples of medium size.

Terence Tao (Oct 29 2024 at 14:46):

By the way, the mod p based walk argument in that paper is very nice; it reminds me of one of the standard proofs of the Sylow theorems.

Terence Tao (Oct 29 2024 at 14:49):

The way "switching" works is that one starts with a central groupoid with edges aba \to b, cdc \to d, and "switches" them to ada \to d, cbc \to b (assuming that the latter edges are not present). It isnot hard to see that, if b,db,d have the same outneighbours (i.e., bxb \to x iff dxd \to x) and a,ca,c have the same inneigbours, the central groupoid axiom is preserved. Thus for instance there are a lot of ways to perform switching on a natural central groupoid. Is there an analogous operation for weak central groupoids? One may also have to switch the notion of a good path as well as the underlying edges. If there is a way (at least in principle) to switch a 151-obeying weak central groupoidinto a non-151-obeying groupoid then we may be able to use repeated switchings starting from some large standard example of a weak central groupoid to try to locate a 151 counterexample.

Kevin M (Oct 29 2024 at 17:28):

Since we need a non-square to separate eqn1485 and eqn151, I've tried to derive more relations that must exist if a non-square exists. I've been able to get some more results (I can clean them into a useful format and share if this would be useful to anyone, but it sounds like the searches have fizzled out). However nothing that appears to _require_ adding more distinct elements. As the searches have shown, we know more than 5 elements are required to have a non-square element (even for a "relaxed" weak groupoid). So either this is hitting the limit of the automatic theorem provers, or there are just many options on how to continue filling out the table, so the ATP cannot give as direct an answer as I'd like.

Matthew Bolan (Oct 29 2024 at 17:30):

Your last relations brought 12 barely into reach for my solver, so maybe these ones can help. I'm not really sure how to use the relations with more quantifiers in them in a search for relaxed magmas though.

Matthew Bolan (Oct 29 2024 at 17:31):

Actually, maybe I should look for a non-relaxed magma violating 166. If it exists it would be nice to have as an example, and that's pretty easy for me to do.

Kevin M (Oct 29 2024 at 17:40):

That being said, (very speculatively) playing with it by hand it feels like there is some fairly "simple" infinite graph growing in there. As if there were a class of elements, with rules like f(n+1) = f(n) ◇ 0 or something. I'm not sure what counts as convincing evidence of this, but there are relations like the following which seem to pick out "classes" of vertices.

(2 * (4 * x)) * x = 4 * x.
(0 * (2 * x)) * x = 2 * x.
(0 * (3 * x)) * x = 3 * x.

So (2 ◇ (4 ◇ x)) ~ 4 ?

Kevin M (Oct 29 2024 at 17:40):

I cannot shake the feeling that there is already some structure in there that leads to an infinite number of vertices, showing that a finite counter example does not exist. But I don't know how to "experiment" with this a bit more systematically to convince myself one way or the other.

Vlad Tsyrklevich (Oct 29 2024 at 18:44):

For what it's worth, using the method from here with the 3 axioms given there, as well as a larger set I generated locally, I am not able to generate a finite implication from 1485=>{151,3456}. This is obviously not a disproof that finite counterexamples do not exist, just a data point.

Andy Jiang (Oct 29 2024 at 23:46):

A little off topic, but what's the easiest way to see the proofs of the known implications/anti-implications. Like the equational explorer tells me 1483 -/-> 151, but is there a way to see what the counterexample was easily?

Terence Tao (Oct 29 2024 at 23:50):

the script in equational#741 will do this. You need to first build a json file of implications but that is available here. I'll add a task to integrate this into EE.

Kevin M (Oct 30 2024 at 06:19):

We've seen finite magmas satisfy eqn 1485 such that there is an element 'a' that satisfies: (a ◇ x) = (y ◇ a).
The minimal non-trivial example has this (I'll call this size 2 magma M2). We've seen it in larger magma as well, but I think we've only seen this when the result is a direct product of multiple M2. Have we even seen it in other cases?

I'm interested if it is possible to have such an element in other cases. Can this exist in infinite magmas? And what are the implications of this?

Andy Jiang (Oct 30 2024 at 07:12):

Interesting question! At least vampire doesn't think it implies every element is a square so it seems like it could be a promising thing to assume. Let me just say some obvious things

Call the element b that is equal to (a * x) or y *a for any x or y
1) a only has a path to b and from b (in graph language)
2) b has paths to and from every other vertex (in graph language)
3) b *b =a directly from 1485

Kevin M (Oct 30 2024 at 07:15):

The reason I ask is because if I consider such an element, I found a way to use this along with the required relations for an eqn 1485 -> eqn 151 counter example, to obtain some relations that appear to force the magma to be infinite.

To keep this fairly self contained, a counter example for eqn 1485 -> eqn 151 must have distinct elements 0,1,2,3,4 such that (in prover9 format):

  0 != x*x.
  0*0 = 2.
  0*1 = 2.
  1*0 = 2.
  1*1 = 2.
  2*2 = 1.

  0*2 = 3.
  2*0 = 4.

  4*2 = 0.
  2*3 = 0.
  1*4 = 2.
  3*1 = 2.

  1*x != 0.
  1*x != 1.
  1*x != 3.
  1*x != 4.

  x*1 != 0.
  x*1 != 1.
  x*1 != 3.
  x*1 != 4.

I could not prove (exists a a*1!=2)., and I could not make useful progress filling out the table further.
So I decided to explore adding

  1*x = 2.
  x*1 = 2.

This results in some useful relations

  (2 * x) * 2 = x.
  2 * (x * 2) = x.
  2*x != x.
  x*2 != x.

And now I can prove (2◇(2◇0)) and ((0◇2)◇2) are new distinct elements that square to 2. This also means they must be non-squares (as squares satisfy eqn 151, and (2◇2) = 1).
So we started with one non-square, and a special element "1" that results in a constant with the magma operation, and we proved more non-squares must exists with the same relationship with that special element.

I'm not sure how to prove this formally, but it appears there may be enough here to define:
f(n+1) = (2 ◇(2◇ f(n))), f(0) = 0
so that all f(n) are distinct non-squares, and therefore there is no finite magma with the initial properties:

  • one non-square element 0
  • and a special element 1 = (0◇0)◇(0◇0) with constant multiplication.

Can anyone see some contradiction this causes? Or an idea of how to flesh out the rest of the table?

Kevin M (Oct 30 2024 at 07:44):

Alternatively, imagine this infinite magma exists. The direct product of this with a "strong" central groupoid would still be a counter example for eqn 1485 -> eqn 151 and yet no longer have an element that operates to a constant, correct?

So I wonder if that "freedom" is what is preventing me from further deriving required relations in the magma table. Is there some way to show that in some sense, "without loss of generality" we can take (1◇x) = (x◇1) = 2 ?

Andy Jiang (Oct 30 2024 at 10:09):

Just a quick note that under your assumption on 1 being multiplication constant 3=0 * 4 and 4 = 3 * 0.
I believe...

Andy Jiang (Oct 30 2024 at 11:18):

Screenshot 2024-10-30 at 7.17.28 PM.png
Here is Kevin's multiplications in table form
(plus maybe the one extra I mentioned which follow from his assumption that (1x=y1))

Andy Jiang (Oct 30 2024 at 11:21):

it's kind of interesting that there is
one element in the table which needs exactly 2 copies of 0 to make minimally, 2 which need 3 copies (which are 3,4); and 3 which need 4 copies (1,5,6). I ran a program which said that if you use 5 copies of 0, with his assumptions you get 4 more elements. Not sure if it's just a coincidence yet though

Andy Jiang (Oct 30 2024 at 12:53):

OK at least my code does not think this pattern continues. Attaching a document which is what my code (calling vampire) think the assumption that
0 * 0=2;
2 * 2=1;
1 * x=y * 1 implies
equivalence.txt
I wouldn't trust it too much though there's still a high chance of bugs

Andy Jiang (Oct 30 2024 at 13:06):

Another interesting (albeit easy) observation about Kevin's hypothetical model is that it cannot have any nontrivial quotients which are "strong" central groupoids because 1x=y1=2 implies that x=(1 * x) * (x * 1) = 0 in a strong central groupoid

Michael Bucko (Oct 30 2024 at 13:09):

Andy Jiang schrieb:

OK at least my code does not think this pattern continues. Attaching a document which is what my code (calling vampire) think the assumption that
0 * 0=2;
2 * 2=1;
1 * x=y * 1 implies
equivalence.txt
I wouldn't trust it too much though there's still a high chance of bugs

Andy, could you please also share the script?

Andy Jiang (Oct 30 2024 at 13:12):

search.py
sure.
[I would love to blame any bugs on Claude but the truth is probably if I wrote it myself it would probably be worse (I did try to proofread it for errors though but definitely not extremely thoroughly)]

Michael Bucko (Oct 30 2024 at 13:22):

Andy Jiang schrieb:

search.py
sure.
[I would love to blame any bugs on Claude but the truth is probably if I wrote it myself it would probably be worse (I did try to proofread it for errors though but definitely not extremely thoroughly)]

It's really great that we're using and making tools! Moreover, I think it'd be great to keep adding microservices around all that. All this knowledge and those tools should not be lost. Well, microservices + tests + evals.

For instance, having access to your script saves me a lot of time. Building reliable SW is still tricky, even with the help from AI. Experience and human oversight are extremely important.

thormikkel (Oct 30 2024 at 17:24):

I'm trying to get a better sense of what types of infinite graphs can encode law 1485 using Tao's axiom system. Something struck me: we're leaning heavily on the 5-cycle axiom (claim 4 in the blueprint), but can we prove that a 5-cycle always exists? And what happens if it doesn't? Doesn't the axoim system then collapse to describing a strong central groupoid?

Matthew Bolan (Oct 30 2024 at 17:24):

yeah the graph has tons of 5-cycles

Terence Tao (Oct 30 2024 at 17:25):

We know that any two points are connected by a 2-path, so in particular any 3-path can be completed to a 5-cycle (and any edge can be completed to a 3-path and thence to a 5-cycle). This already gives a lot of 5-cycles

Andy Jiang (Oct 31 2024 at 05:17):

Screenshot 2024-10-31 at 1.15.40 PM.png
Filled in a few more boxes of Kevin's hypothetical model--seems like there's some patterns here but I can't quite put my finger on it. I think it would be interesting to study the graph of this when you remove the vertices 1 and 2 (which are connected to nothing and everything respectively so there's no harm)

Andy Jiang (Oct 31 2024 at 05:24):

I guess Kevin was saying that 9 and 10 here square to 2?

Andy Jiang (Oct 31 2024 at 05:36):

By the way I'm not sure how easy this is but Vampire thinks that left and right multiplication by 2 are inverses (edit: just saw this was already in Kevin's comment above)

Andy Jiang (Oct 31 2024 at 11:21):

I've translated into graph axioms what we would expect a weak (non-relaxed) central groupoid with an element 1 such that 1 * a = b * 1 = 2 for all a,b once you remove 1 and 2 from the graph, following the above ideas of Kevin. I think the conditions are as follows (but I didn't check them carefully). It consists of a graph G with some directed edges and some subset of length 2 paths which we call good. It also now carries an automorphism of the underlying set of the graph which I'll call T (which comes from right multiplication by 2)

1) If Tb -/>T^-1(a) (no edge), then there is a unique c such that a -> c -> b is good. If Tb -> T^{-1}a, there is no such c.
2) If x -> y, then x -> y -> Ty and T^{-1}x -> x -> y are good
3) x -> Tx
4) (5-cycle axiom) If a -> b, b->c->d good and d->e->a good, then c->d->e is good.
5) If x->y, w->T^{-1}(y), and T^{-1}w -> l -> x is good, then l=w

Now it will not lead to easier finite counterexamples b/c when you add back 1 and 2 it becomes a finite counterexample as before. But might be useful for infinite counterexamples.

Jose Brox (Oct 31 2024 at 12:44):

Negative result: I have checked all the weak-not-strong central groupoids that Mace4 generated (180 of size 16, 200 of size 18) and none refutes any unknown implication.

Andy Jiang (Oct 31 2024 at 12:52):

Jose Brox said:

Negative result: I have checked all the weak-not-strong central groupoids that Mace4 generated (180 of size 16, 200 of size 18) and none refutes any unknown implication.

By the way if you have the full list up to isomorphism; could you share it? I think it'll be nice to just verify that all of them do map to the 2-element model (if only for fun)

Kevin M (Oct 31 2024 at 17:54):

Andy Jiang said:

By the way if you have the full list up to isomorphism; could you share it?

I accidentally left the size 18 one running and it filled up my harddrive. So I have ~ 100GB of size 18.
I'm not sure how to simplify the results "up to isomorphism". The isofilter program appears smart enough to not use 18! permutations to do its checks, but it is still a bit too generic, and takes hours per entry. We need something smarter, but I don't quite know what is sufficient.

Using the adjacency matrix, we should be able to split up into "high" and "low" elements. I have the code for that, and if it would be useful I could at least check that on all the models I have. Once using that to split into four 9x9 blocks, I'm not sure what to do next. The "low" block only has entries with "high" elements, and the "high" block only has entries with the "low" elements. 9! is small enough I can "canonicalize" one of the blocks only considering permutations on that block, but really they are coupled, and I do not know how to proceed there. Since they are "almost" decoupled after the high-low split, maybe there is some trick here, but I'd need someone to explain the algorithm to me.

I need that drive space back, so I may have to delete most of that data soon.

Eric Taucher (Oct 31 2024 at 18:20):

Kevin M said:

I need that drive space back, so I may have to delete most of that data soon.

Maybe you have an old USB drive laying around. :thinking:

Kevin M (Oct 31 2024 at 19:17):

Andy Jiang said:

It also now carries an automorphism of the underlying set of the graph which I'll call T (which comes from right multiplication by 2)

I've played with this a bit more, and I think there is more symmetry here. With the special properties of 2,

    2*(x*2) = x.
    (2*x)*2 = x.
    (x !=y ) <-> (2*x != 2*y).
    (x !=y ) <-> (x*2 != y*2).

The two "new" non-squares that show up by operating on 0 are actually related. The process can be repeated and they all fall on a line:
f(n+1) = 2 ◇ (2 ◇ f(n)).
f(n-1) = (f(n) ◇ 2) ◇ 2.
f(0) = 0.
where f : integer -> magma element.

prover9 is able to give me enough to prove that these all square to 2

    % base case
    0*0 = 2.

    % induction step
    (x*x = 2) -> ( (2*(2*x))*(2*(2*x)) = 2 ).
    (x*x = 2) -> ( ((x*2)*2)*((x*2)*2) = 2 ).

As we can "shift" them around by multiplying by 2, in order to prove all f(n) are distinct it would be sufficient to show f(k) != f(0) for all k>0.
However I have not figured out a path to this.
Trying some explicitly:

    % f(1), is  distinct from f(0), (and 1, which also squares to 2)
    2*(2*0) != 0.
    2*(2*0) != 1.

    % f(2) is distinct from f(0), (and 1, which also squares to 2)
    2*(2*(2*(2*0))) != 0.
    2*(2*(2*(2*0))) != 1.

    % f(3) is distinct from f(0), (and 1, which also squares to 2)
    2*(2*(2*(2*(2*(2*0))))) != 0.
    2*(2*(2*(2*(2*(2*0))))) != 1.

    % failed attempt at generalizing
    % too strong? or incorrect phrasing?
    %( x*x=2  &  2*(2*x)=0 ) -> ( x=(0*2)*2 ). %--unknown
    %( x*x=2  &  x!=(0*2)*2 ) -> ( 2*(2*x)!=0 ). %--unknown
    %( x*x=2  &  x!=1 ) -> ( 2*(2*x) != 1 ). %--unknown

Kevin M (Oct 31 2024 at 19:21):

Anyway, the point is that since all f(n) are non-squares that square to 2, and this is all we initially assumed about 0, I think they are all "equivalent", in that we could have chosen any of them to be "0" (shifting the line).

Andy Jiang (Oct 31 2024 at 22:44):

@Kevin M ah ok no worries about the finite models up to isomorphism thing. Probably it's not essential as long as we know they all map to the 2 element model as Terry pointed out.

Andy Jiang (Oct 31 2024 at 23:32):

Also I think whenever Tx has an edge to T^-1(x), then x squares to 2 because
x * x = (T^-1(x) * 2) * (2 * T(x)) = 2 and I think this is iff because x * Tx = (T^-1 x * 2) * (2 * x) and T^-1 x always has edge to x

Andy Jiang (Oct 31 2024 at 23:42):

I don't get your unknown statement btw don't they follow from the T automorphism

Andy Jiang (Oct 31 2024 at 23:49):

If x squares to 2,

T^2(x) * T^2(x) = 2 is equivalent to T^3(x) has an edge to T(x) by above. But I don't immediately see why. Let me try to see how vampire proves it.

Andy Jiang (Nov 01 2024 at 00:21):

Proof of x squares to 2 implies T^2(x) squares to 2 by vampire is kind of complicated. I made it more human readable as follows

Lemma 1
Suppose x squares to 2, and x has an edge to y then
(y * x) * 2 = (y * x) * (x * x) = x
So y * x = 2 * x
Corollary: Tx * x = T^-1 x

Lemma 1.5 if y has an edge to x and x squares to 2 then

2* (x * y) = x so x * y = Tx

Lemma 2
for any x
(T^2 x) * (T^-1 x) = (Tx * 2) * (2 * x) = 2

Lemma 3
If y has an edge to T^2x then x has an edge to y because

2 * ((T^-1x) * y) = ((T^2 x) * (T^-1 x)) * ((T^-1x) * y) = T^-1 x

Lemma 4
Tx = x * (Tx * y) for all y (because y has an edge to 2 and x = 2 * Tx)

Lemma 5
if x squares to 2 then
((T^2 x) * (T^2x)) * x
= 2 * x
Since by lemma 3, x->((T^2 x) * (T^2x)) then apply lemma 1

Lemma 6
Tx * (T^2 x * T^2 x) = T^2 x by lemma 4

Hence, suppose x squares to 2, then
(T^2 x) * (T^2 x)
=(Tx * (T^2 x * T^2 x))* ((T^2 x * T^2 x) * x) by 1485 and x -> Tx
= (T^2 x)* (T^-1 x) by lemma 6,5
= 2 by lemma 2

Andy Jiang (Nov 01 2024 at 02:22):

@Kevin M I asked Vampire whether T^2 (right multiplication by 2 twice) is automorphism of the magma and after 53sec it said yes. Seems important!
1485_auto
I mean the important thing is that T^2 is a homomorphism

Kevin M (Nov 01 2024 at 04:53):

Nice! Its great to know for sure that "shifts along the line" actually are automorphisms. It looks like 1,2 are unique, while all other elements should really be indexed by an integer?

I've been indexing so f(n) ~ f(0) = 0, but as you noted, maybe it would be better to track every "shift" by 2. So maybe it should be
g(n+1) = ( 2 ◇ g(n) )
g(n-1) = ( g(n) ◇ 2 )
g(0) = 0.
All "even" elements are "equivalent", non-squares that square to 2.
All "odd" elements are "equivalent". I cannot seem to prove if these are squares or non-squares. They are not idempotents though.

Matthew Bolan (Nov 01 2024 at 05:17):

Aughh, missed a case and the program I wrote as a sanity check works for the first 10 or so steps of the greedy algorithm so I was bamboozled. Will report back if I repair it.

Matthew Bolan (Nov 01 2024 at 05:42):

I maybe can repair it if we have a seed partial magma that 1. violates 151 and 2. has some kind of control over squares?

Matthew Bolan (Nov 01 2024 at 05:43):

The greedy algorithm it describes still gets far enough to have fooled my sanity check program (and therefore me), so perhaps it can be repaired or can give us some clue?

Matthew Bolan (Nov 01 2024 at 06:25):

It is difficult to write up formally something which does not work. I introduced 3 brand new elements to try and take care of aba \diamond b by making it a part of its own brand new 5-cycle of good paths, with some tweaking when bab \diamond a was already defined. I forgot a case with squares, and it collapsed. The more I look at it the less hopeful I get, given how essential we know squares are.

Matthew Bolan (Nov 01 2024 at 06:27):

And in the process I've guess I've learned that these greedy algorithms for this one can survive a very long time before they actually hit a contradiction, so computer sanity checks should be run until the table is very large...

Matthew Bolan (Nov 01 2024 at 06:42):

So far most of my (probably naive) attempts at a greedy algorithm have taken the form of gluing on at the two vertices I wish to target a second graph where the good path exists, and making all/most paths that cross between the graphs bad paths. A frequent obstacle to this is that paths between these two vertices live in both graphs at once, and force relations.

Kevin M (Nov 01 2024 at 06:57):

Something that is bugging me:
given the automorphism, what can we say about the product: g(n) ◇ g(m)?
To preserve the symmetry, this should be equal to T^n (g(0)◇ g(m-n)). (EDIT: I forgot to require an even number of shifts, and so what follows is flawed.)

As an example: 0 = 4 ◇ 3 = g(1) ◇ g(-1) = T ( g(0) ◇ g(-2) ) = T (0 ◇ ((0 ◇ 2) ◇ 2)) = 2 ◇ (0 ◇ ((0 ◇ 2) ◇ 2)).
Double check that relation is 0 ... yep!

So these products are known if we can get h(k) = g(0) ◇ g(k) for all k.
I went to calculate some of these to see if I could see a pattern, but unlike g(n), the h(k) are not unique!?
In particular:
h(1) = ( 0 ◇ g(1) ) = 3 = g(-1)
h(-2) = ( 0 ◇ g(-2) ) = 3 = g(-1)

So g(n) ◇ g(m) can map into another g(k) ? But not always!? (Since h(0) = g(0) ◇ g(0) = 2 != g(k) for any k.)
This seems super weird to me.

Andy Jiang (Nov 01 2024 at 07:00):

@Kevin M
I'm not sure that every element can be translated to 0 (i.e. I don't think everything is of the form g(n) for some n), I'm not sure if you're assuming this. Assuming that g(n+1) = 2 * g(n) the automorphism shifts the n by even number not all numbers I'm also not sure if that's relevant here

Kevin M (Nov 01 2024 at 07:08):

I'm not assuming that every element can be translated to 0. (For example, if this is actually infinite, then there is no n such that g(n) = 2, likewise for g(n)=1.)

I'm not sure were things like 3◇4 fall though. I'm quite confused, as my initial expectations where all g(n)◇g(m) would be distinct from all g(k) like g(0)◇g(0)=2 is.

Andy Jiang (Nov 01 2024 at 07:09):

well I think 3 * 4 is not in the translation class but 4 * 3 is, as you say

Andy Jiang (Nov 01 2024 at 07:12):

probably the nonuniqueness is because the graph is like 9 -> 4 -> 0 -> 3 -> 10 and 3 has an edge to 4 so the good path from 0 to 10 and to 4 are both through 3. not sure if that's any more intuition than what you said though...

Kevin M (Nov 01 2024 at 07:13):

Wait a second. I just realized we can look at the square of 3 now.
3 ◇ 3 = g(-1) ◇ g(-1) = T^-1 ( g(0) ◇ g(0) ) = T^-1 ( 2 ) = 2 ◇ 2 = 1

What? Did I do that wrong? Because we can prove 3 ◇ 3 != 1.
Does the proof of the automorphism actually lead us to a contradiction?

Andy Jiang (Nov 01 2024 at 07:13):

the automorphism shifts by 2 not 1 right?

Andy Jiang (Nov 01 2024 at 07:13):

like you have to multiply by 2 twice

Kevin M (Nov 01 2024 at 07:14):

Oh, duh. Thanks!

Kevin M (Nov 01 2024 at 07:14):

Oh, I see above I made that mistake before, and just lucked out.

Andy Jiang (Nov 01 2024 at 07:14):

anyway my intuition is that models of the form you describe actually exist (but my intuition is only because as I'm reading some of Vampire's proofs I realize this vampire is really quite powerful so I doubt there can be any short contradictions we can find lol)

Kevin M (Nov 01 2024 at 07:17):

Just in case, can you rerun that automorphism proof with 0!=2, to make sure it is not collapsing to the trivial model?

Andy Jiang (Nov 01 2024 at 07:21):

I'm not sure how that would help. It's finding a proof that exists 1 : 1 * x = 2 => automorphism. If you add more assumptions it will not stop this proof (imagine there are no nontrivial models, then your assumptions will be inconsistent but it will still imply the statement)

Andy Jiang (Nov 01 2024 at 07:22):

but I did run on much better hardware a search for 2 minutes for contradiction with your assumptions from 0!=2 and I didn't find in two minutes. But I can run it longer later

Andy Jiang (Nov 01 2024 at 07:22):

for comparison the proof of automorphism on that machine was ~1sec

Jose Brox (Nov 01 2024 at 09:12):

Andy Jiang ha dicho:

By the way if you have the full list up to isomorphism; could you share it? I think it'll be nice to just verify that all of them do map to the 2-element model (if only for fun)

I believe isofilter in Windows GUI has some bug (for the size 18 models, its says they are all isomorphic, but does the computation in a fraction of a second, and then returns an empty list...), so I will just upload here the full list. The size 16 models are isofiltered, if the program run correctly.
Mace4 1485 size 18 200 models portable.txt
Mace4 1485 size 16 181 models isofilter portable.txt

Andy Jiang (Nov 01 2024 at 10:05):

This is a crazy result and probably a bug somewhere or some idiosyncrasies of vampire, but I'm getting empirically that under Kevin's assumptions of 1 * x = 2 for all x,
if a * a = 2, then a * T^k (a) = 2 if k is congruent to 0,4,7 mod 10 where T(u) = (u * 2) lol (I think those are all which are true in general)

Andy Jiang (Nov 01 2024 at 10:12):

I think the proofs that vampire comes up with for T^(2k)(a)=a => a=1 (the noninteresting square root of 2) are based on playing with the above equalities this periodicity. So potentially if your period is a multiple of 10, vampire can't use the same proofs?? I have no idea to be honest...

Edit: Vampire does seem to so far struggle with ruling out T^10(a) = a (although it's not extremely convincing because it only rules out the previous ones at long time limit so I'll run it for longer soon)

Andy Jiang (Nov 01 2024 at 12:06):

https://teorth.github.io/equational_theories/fme/?magma=%5B%0A%20%20%20%20%5B2%2C2%2C3%2C3%2C3%2C3%2C2%2C3%2C2%2C2%2C2%2C3%2C3%2C2%2C3%2C2%2C2%2C3%2C2%2C2%2C3%2C2%2C3%2C3%2C3%2C3%2C3%2C2%2C2%2C3%2C2%2C2%5D%2C%0A%20%20%20%20%5B2%2C2%2C2%2C2%2C2%2C2%2C2%2C2%2C2%2C2%2C2%2C2%2C2%2C2%2C2%2C2%2C2%2C2%2C2%2C2%2C2%2C2%2C2%2C2%2C2%2C2%2C2%2C2%2C2%2C2%2C2%2C2%5D%2C%0A%20%20%20%20%5B4%2C2%2C1%2C0%2C6%2C3%2C8%2C5%2C10%2C7%2C11%2C9%2C23%2C28%2C22%2C31%2C26%2C27%2C24%2C29%2C25%2C30%2C12%2C13%2C17%2C18%2C14%2C19%2C15%2C21%2C20%2C16%5D%2C%0A%20%20%20%20%5B4%2C2%2C5%2C12%2C17%2C3%2C2%2C5%2C20%2C7%2C11%2C14%2C23%2C4%2C22%2C11%2C26%2C3%2C24%2C29%2C25%2C7%2C12%2C23%2C17%2C22%2C14%2C24%2C29%2C25%2C20%2C26%5D%2C%0A%20%20%20%20%5B4%2C2%2C0%2C0%2C13%2C3%2C8%2C12%2C19%2C7%2C2%2C18%2C23%2C28%2C22%2C8%2C7%2C27%2C24%2C4%2C3%2C30%2C12%2C13%2C23%2C18%2C22%2C19%2C28%2C27%2C24%2C30%5D%2C%0A%20%20%20%20%5B2%2C2%2C7%2C7%2C2%2C2%2C2%2C7%2C7%2C7%2C2%2C7%2C2%2C2%2C7%2C2%2C7%2C2%2C7%2C2%2C2%2C7%2C7%2C2%2C2%2C7%2C7%2C7%2C2%2C2%2C7%2C7%5D%2C%0A%20%20%20%20%5B4%2C2%2C4%2C4%2C4%2C2%2C2%2C4%2C4%2C2%2C2%2C2%2C4%2C4%2C2%2C2%2C2%2C2%2C4%2C4%2C2%2C2%2C4%2C4%2C4%2C2%2C2%2C4%2C4%2C2%2C4%2C2%5D%2C%0A%20%20%20%20%5B2%2C2%2C9%2C18%2C21%2C3%2C8%2C14%2C16%2C7%2C11%2C9%2C3%2C8%2C22%2C31%2C26%2C27%2C7%2C11%2C25%2C30%2C22%2C27%2C25%2C18%2C14%2C30%2C31%2C21%2C26%2C16%5D%2C%0A%20%20%20%20%5B4%2C2%2C6%2C13%2C6%2C3%2C8%2C17%2C15%2C2%2C11%2C21%2C23%2C28%2C3%2C31%2C11%2C27%2C4%2C29%2C25%2C8%2C23%2C13%2C17%2C27%2C25%2C28%2C15%2C21%2C29%2C31%5D%2C%0A%20%20%20%20%5B2%2C2%2C11%2C2%2C11%2C2%2C2%2C11%2C11%2C2%2C11%2C11%2C2%2C2%2C2%2C11%2C11%2C2%2C2%2C11%2C11%2C2%2C2%2C2%2C11%2C2%2C11%2C2%2C11%2C11%2C11%2C11%5D%2C%0A%20%20%20%20%5B2%2C2%2C8%2C8%2C8%2C2%2C8%2C2%2C8%2C2%2C2%2C8%2C2%2C8%2C2%2C8%2C2%2C8%2C2%2C2%2C2%2C8%2C2%2C8%2C2%2C8%2C2%2C8%2C8%2C8%2C2%2C8%5D%2C%0A%20%20%20%20%5B4%2C2%2C10%2C19%2C15%2C2%2C8%2C20%2C10%2C7%2C11%2C16%2C4%2C28%2C7%2C31%2C26%2C8%2C24%2C29%2C11%2C30%2C24%2C28%2C29%2C30%2C26%2C19%2C15%2C31%2C20%2C16%5D%2C%0A%20%20%20%20%5B2%2C2%2C22%2C22%2C3%2C3%2C2%2C22%2C7%2C7%2C2%2C22%2C3%2C2%2C22%2C2%2C7%2C3%2C7%2C2%2C3%2C7%2C22%2C3%2C3%2C22%2C22%2C7%2C2%2C3%2C7%2C7%5D%2C%0A%20%20%20%20%5B4%2C2%2C23%2C23%2C23%2C3%2C2%2C23%2C4%2C2%2C2%2C3%2C23%2C4%2C3%2C2%2C2%2C3%2C4%2C4%2C3%2C2%2C23%2C23%2C23%2C3%2C3%2C4%2C4%2C3%2C4%2C2%5D%2C%0A%20%20%20%20%5B2%2C2%2C26%2C7%2C11%2C2%2C2%2C26%2C26%2C7%2C11%2C26%2C2%2C2%2C7%2C11%2C26%2C2%2C7%2C11%2C11%2C7%2C7%2C2%2C11%2C7%2C26%2C7%2C11%2C11%2C26%2C26%5D%2C%0A%20%20%20%20%5B4%2C2%2C28%2C28%2C28%2C2%2C8%2C4%2C28%2C2%2C2%2C8%2C4%2C28%2C2%2C8%2C2%2C8%2C4%2C4%2C2%2C8%2C4%2C28%2C4%2C8%2C2%2C28%2C28%2C8%2C4%2C8%5D%2C%0A%20%20%20%20%5B2%2C2%2C31%2C8%2C31%2C2%2C8%2C11%2C31%2C2%2C11%2C31%2C2%2C8%2C2%2C31%2C11%2C8%2C2%2C11%2C11%2C8%2C2%2C8%2C11%2C8%2C11%2C8%2C31%2C31%2C11%2C31%5D%2C%0A%20%20%20%20%5B4%2C2%2C24%2C24%2C4%2C2%2C2%2C24%2C24%2C7%2C2%2C7%2C4%2C4%2C7%2C2%2C7%2C2%2C24%2C4%2C2%2C7%2C24%2C4%2C4%2C7%2C7%2C24%2C4%2C2%2C24%2C7%5D%2C%0A%20%20%20%20%5B2%2C2%2C25%2C3%2C25%2C3%2C2%2C25%2C11%2C2%2C11%2C25%2C3%2C2%2C3%2C11%2C11%2C3%2C2%2C11%2C25%2C2%2C3%2C3%2C25%2C3%2C25%2C2%2C11%2C25%2C11%2C11%5D%2C%0A%20%20%20%20%5B2%2C2%2C27%2C27%2C27%2C3%2C8%2C3%2C8%2C2%2C2%2C27%2C3%2C8%2C3%2C8%2C2%2C27%2C2%2C2%2C3%2C8%2C3%2C27%2C3%2C27%2C3%2C8%2C8%2C27%2C2%2C8%5D%2C%0A%20%20%20%20%5B2%2C2%2C30%2C30%2C8%2C2%2C8%2C7%2C30%2C7%2C2%2C30%2C2%2C8%2C7%2C8%2C7%2C8%2C7%2C2%2C2%2C30%2C7%2C8%2C2%2C30%2C7%2C30%2C8%2C8%2C7%2C30%5D%2C%0A%20%20%20%20%5B4%2C2%2C29%2C4%2C29%2C2%2C2%2C29%2C29%2C2%2C11%2C11%2C4%2C4%2C2%2C11%2C11%2C2%2C4%2C29%2C11%2C2%2C4%2C4%2C29%2C2%2C11%2C4%2C29%2C11%2C29%2C11%5D%2C%0A%20%20%20%20%5B2%2C2%2C14%2C22%2C25%2C3%2C2%2C14%2C26%2C7%2C11%2C14%2C3%2C2%2C22%2C11%2C26%2C3%2C7%2C11%2C25%2C7%2C22%2C3%2C25%2C22%2C14%2C7%2C11%2C25%2C26%2C26%5D%2C%0A%20%20%20%20%5B4%2C2%2C12%2C12%2C23%2C3%2C2%2C12%2C24%2C7%2C2%2C22%2C23%2C4%2C22%2C2%2C7%2C3%2C24%2C4%2C3%2C7%2C12%2C23%2C23%2C22%2C22%2C24%2C4%2C3%2C24%2C7%5D%2C%0A%20%20%20%20%5B2%2C2%2C18%2C18%2C27%2C3%2C8%2C22%2C30%2C7%2C2%2C18%2C3%2C8%2C22%2C8%2C7%2C27%2C7%2C2%2C3%2C30%2C22%2C27%2C3%2C18%2C22%2C30%2C8%2C27%2C7%2C30%5D%2C%0A%20%20%20%20%5B4%2C2%2C20%2C24%2C29%2C2%2C2%2C20%2C20%2C7%2C11%2C26%2C4%2C4%2C7%2C11%2C26%2C2%2C24%2C29%2C11%2C7%2C24%2C4%2C29%2C7%2C26%2C24%2C29%2C11%2C20%2C26%5D%2C%0A%20%20%20%20%5B2%2C2%2C16%2C30%2C31%2C2%2C8%2C26%2C16%2C7%2C11%2C16%2C2%2C8%2C7%2C31%2C26%2C8%2C7%2C11%2C11%2C30%2C7%2C8%2C11%2C30%2C26%2C30%2C31%2C31%2C26%2C16%5D%2C%0A%20%20%20%20%5B4%2C2%2C17%2C23%2C17%2C3%2C2%2C17%2C29%2C2%2C11%2C25%2C23%2C4%2C3%2C11%2C11%2C3%2C4%2C29%2C25%2C2%2C23%2C23%2C17%2C3%2C25%2C4%2C29%2C25%2C29%2C11%5D%2C%0A%20%20%20%20%5B4%2C2%2C13%2C13%2C13%2C3%2C8%2C23%2C28%2C2%2C2%2C27%2C23%2C28%2C3%2C8%2C2%2C27%2C4%2C4%2C3%2C8%2C23%2C13%2C23%2C27%2C3%2C28%2C28%2C27%2C4%2C8%5D%2C%0A%20%20%20%20%5B4%2C2%2C19%2C19%2C28%2C2%2C8%2C24%2C19%2C7%2C2%2C30%2C4%2C28%2C7%2C8%2C7%2C8%2C24%2C4%2C2%2C30%2C24%2C28%2C4%2C30%2C7%2C19%2C28%2C8%2C24%2C30%5D%2C%0A%20%20%20%20%5B2%2C2%2C21%2C27%2C21%2C3%2C8%2C25%2C31%2C2%2C11%2C21%2C3%2C8%2C3%2C31%2C11%2C27%2C2%2C11%2C25%2C8%2C3%2C27%2C25%2C27%2C25%2C8%2C31%2C21%2C11%2C31%5D%2C%0A%20%20%20%20%5B4%2C2%2C15%2C28%2C15%2C2%2C8%2C29%2C15%2C2%2C11%2C31%2C4%2C28%2C2%2C31%2C11%2C8%2C4%2C29%2C11%2C8%2C4%2C28%2C29%2C8%2C11%2C28%2C15%2C31%2C29%2C31%5D%0A%20%20%5D

Zoltan A. Kocsis (Z.A.K.) (Nov 01 2024 at 12:09):

@Andy Jiang This is very very nice! I guess this closes the book on the one-time final boss, 1485 :)

Andy Jiang (Nov 01 2024 at 12:10):

Found with Mace4 using Kevin's assumptions + T^10(0)=0

Michael Bucko (Nov 01 2024 at 12:11):

@Andy Jiang Could you please share that Mace4 script? I'd love to have a look.

Andy Jiang (Nov 01 2024 at 12:23):

just

x=(yx)(x(zy)).
0*0=2.
(1*x)=2.
(x*1)=2.
((((((((((02)2)2)2)2)2)2)2)2)2) = 0.

actually even without the last line of T^10(0)=0 condition it can be found pretty fast (aka instantly if you specify the size=32 ...)....in hindsight we should've ran the search on Kevin's conditions on finite models immediately but I guess we didn't believe it could have finite models...

thormikkel (Nov 01 2024 at 12:33):

I'm very (pleasantly) surprised that this was resolved with a finite counterexample! That's some great detective work by you two :smiley:

This also negates the conjecture that every finite weak central groupoid has a homomorphism to the order 2 groupoid right? Apparently these objects have more degrees of freedom than was thought originally.

Vlad Tsyrklevich (Nov 01 2024 at 12:36):

@Andy Jiang Do you plan to PR this? Obviously feel free, but I'd be happy to do so if you don't have a dev env set-up

Andy Jiang (Nov 01 2024 at 12:37):

If you can do it, it would be great!

Andy Jiang (Nov 01 2024 at 12:44):

thormikkel said:

I'm very (pleasantly) surprised that this was resolved with a finite counterexample! That's some great detective work by you two :smiley:

This also negates the conjecture that every finite weak central groupoid has a homomorphism to the order 2 groupoid right? Apparently these objects have more degrees of freedom than was thought originally.

yup! although it does raise the question whether all finite models are of size 2*n^2. Also I'm personally glad that the counterexample can be found by mace4 without much compute so anyone can reproduce the process easily just verify the example. Thanks again to @Kevin M !

Terence Tao (Nov 01 2024 at 12:51):

Wow, this is amazing! I've been distracted and not been following the latest progress too closely. Can someone summarize the path that led to this finite example?

Vlad Tsyrklevich (Nov 01 2024 at 12:55):

Individual build time for the new refutation is ~9s, which seems fine.

Vlad Tsyrklevich (Nov 01 2024 at 12:57):

I noticed that this also refutes 1485=>{1427,1429,2087,2124} which we have disproofs for in WeakCentralGroupoids.lean; however, this example is explicitly finite so it also helps out on solving the finite graph. I will add this manually as I assume there are no objections. This is also a good reminder that I should search to find if there are finite counter examples that we have upstream that we are not recording as finite right now, due to having previously developed a non-finite counterexample.

Terence Tao (Nov 01 2024 at 12:58):

I wonder what the graph of this order 32 counterexample looks like. For the order 8 and 16 examples, computing the square A^2 of the adjacency matrix (i.e., counting how many paths of length two connect any two points) was helpful in revealing additional structure (in those cases, it was a homomorphism to the order 2 magma). I'm curious if something similar happens here.

Terence Tao (Nov 01 2024 at 12:59):

Ironically it's probably a good thing I wasn't involved in 1485 the last few days, I would most likely have been pushing infinite magma approaches and maybe would have dissuaded people from looking too much at finite examples. That's another side benefit of having a very broad base of collaborators - a lot less "groupthink".

Vlad Tsyrklevich (Nov 01 2024 at 13:02):

equational#777

Jose Brox (Nov 01 2024 at 13:05):

Terence Tao ha dicho:

maybe would have dissuaded people from looking too much at finite examples

I've been running long Mace4 runs just in case for these anti-implications, so perhaps they would have eventually hit some day...

Andy Jiang (Nov 01 2024 at 13:05):

Terence Tao said:

Wow, this is amazing! I've been distracted and not been following the latest progress too closely. Can someone summarize the path that led to this finite example?

I can summarize my part and hopefully @Kevin M can do the same because I think the main idea is due to him. The path is kind of convoluted in hindsight but here goes.

Kevin suggested to look for models in which there is an element 1 such that 1 * x = y * 1 = 2 for all x and y. He observed that left and right multiplication by 2 are mutual inverses and asked whether it can be shown that this condition implies there are infinitely many square roots of 2 (noting that you can move from one square root of 2 to another by multiplication by 2 twice)

I followed Kevin's computations and played around with filling in the multiplications starting with any element other than 1 which squares to 2 (as suggested by Kevin). But I only gained a faint sense of structure by doing this but no concrete results.

After this I think most of my understanding came from just trying to ask vampire to prove a bunch of random statement from Kevin's assumptions (lol). From doing this I noticed that if a * a = 2, then a * T^k (a) = 2 where T(a) = a * 2 whenever k is 0,4,7 mod 10 which was really weird to me and it seems like this is how vampire is able to show that if T^2(a)=a or T^4(a)=a then the model cannot be nontrivial (which was observed by Kevin). But it occured to me that this argument probably breaks if T^10(a)=a. Which suggested there was a chance as finite counterexamples.
(Edit: I forgot another promising lead was that T^2 is actually an automorphism of the magma with the multiplication structure!)

At this point I just tried some ways to find finite counterexamples, and ultimately Mace4 worked with Kevin's assumptions (1 * x = y * 1 = 2) together with T^10(0)=0. But I also realized after finding the example that if one were to guess at 32 (the next 2*n^2) we could have just used Kevin's assumptions and found the example even without all the theoretical understanding. But it was fun anyways.

Edit: I want to again emphasize that for me I relied on Vampire a lot during this--basically running it for a while and whenever it doesn't derive counterexample when I add assumptions my working assumption would be it could be satisfiable. I think this approach sounds kind of wild at first to trust this ATP so much but I felt like Vampire was really really strong at deriving contradictions with first order logic at least at this particular task.

Michael Bucko (Nov 01 2024 at 13:32):

Andy Jiang schrieb:

Terence Tao said:

Wow, this is amazing! I've been distracted and not been following the latest progress too closely. Can someone summarize the path that led to this finite example?

I can summarize my part and hopefully Kevin M can do the same because I think the main idea is due to him. The path is kind of convoluted in hindsight but here goes.

Kevin suggested to look for models in which there is an element 1 such that 1 * x = y * 1 = 2 for all x and y. He observed that left and right multiplication by 2 are mutual inverses and asked whether it can be shown that this condition implies there are infinitely many square roots of 2 (noting that you can move from one square root of 2 to another by multiplication by 2 twice)

I followed Kevin's computations and played around with filling in the multiplications starting with any element other than 1 which squares to 2 (as suggested by Kevin). But I only gained a faint sense of structure by doing this but no concrete results.

After this I think most of my understanding came from just trying to ask vampire to prove a bunch of random statement from Kevin's assumptions (lol). From doing this I noticed that if a * a = 2, then a * T^k (a) = 2 where T(a) = a * 2 whenever k is 0,4,7 mod 10 which was really weird to me and it seems like this is how vampire is able to show that if T^2(a)=a or T^4(a)=a then the model cannot be finite (which was observed by Kevin). But it occured to me that this argument probably breaks if T^10(a)=a. Which suggested there was a chance as finite counterexamples.
(Edit: I forgot another promising lead was that T^2 is actually an automorphism of the magma with the multiplication structure!)

At this point I just tried some ways to find finite counterexamples, and ultimately Mace4 worked with Kevin's assumptions (1 * x = y * 1 = 2) together with T^10(0)=0. But I also realized after finding the example that if one were to guess at 32 (the next 2*n^2) we could have just used Kevin's assumptions and found the example even without all the theoretical understanding. But it was fun anyways.

Edit: I want to again emphasize that for me I relied on Vampire a lot during this--basically running it for a while and whenever it doesn't derive counterexample when I add assumptions my working assumption would be it could be satisfiable. I think this approach sounds kind of wild at first to trust this ATP so much but I felt like Vampire was really really strong at deriving contradictions with first order logic at least at this particular task.

So amazing! Congratulations! Thank you for sharing this!

Andy Jiang (Nov 01 2024 at 13:37):

sharing two visualizations (Claude wrote the code to generate this and I didn't really check it so use at your own risk)
1) adjacency table without vertices 1 and 2 (2 is connected bidirectionally to everything and 1 only to 2)
2) the graph without 1 and 2 as well
Screenshot 2024-11-01 at 9.35.35 PM.png
Figure.png

Kevin M (Nov 01 2024 at 20:50):

Andy Jiang said:

Edit: I want to again emphasize that for me I relied on Vampire a lot during this--basically running it for a while and whenever it doesn't derive counterexample when I add assumptions my working assumption would be it could be satisfiable. I think this approach sounds kind of wild at first to trust this ATP so much but I felt like Vampire was really really strong at deriving contradictions with first order logic at least at this particular task.

Yes, I am amazed at the power of these provers. Not just power, but speed. The speed was fast enough to allow an "interactive" style of exploration.

I'd like to add to @Andy Jiang 's summary by going back a bit further and noting that a key piece for this line of exploration was the realization that squares satisfied the implications we were trying to disprove, so we absolutely needed a non-square element. I think this may have been discovered multiple times along the way, but I ran into it while querying prover9 with related questions trying to understand if squaring had to be surjective.

Since we needed a non-square, I again explored this by "querying" prover9. It turned out that an idempotent element cannot "share" its square, and this ended up forcing the table to have at least 5 distinct elements. Playing with the implications of this, there appeared to be some structure that would recursively give me new elements, but I couldn't quite put it together. Prover9 could somehow prove 1◇ x and x ◇ 1 was not almost any element I tried, except 2. Somehow I realized we already had finite model examples with an element 1◇ x = constant, and that a direct product of such a magma with a normal central groupoid would obscure this -- so this was a barrier to exploring further with queries unless I could somehow understand, control, or remove that direct product. (Which I did not know how to do.)

It seemed the easiest way forward was just to guess that prover9 couldn't rule out 1◇ x = 2 because it was possible, and see where that lead. Amazingly/fortuitously it led quickly to some recursive relations I was looking for: f(n+1) = (2 ◇ (2 ◇ f(n))), f(0) = 0, all squared to 2 and therefore had to be non-squares. Weirdly, I found we could prove f(1) != 0, f(2) != 0, f(3) != 0, but prover9 refused all my attempts at some recursive argument that all f(n) were distinct. Now we know why.

Andy had the idea to try to understand the consequence of 1◇ x = 2 more directly, and it turned out the interesting relations with 2 held without the special element 0. So this was more general, and culminated in him discovering this could be made explicit as an automorphism.

It sounds like exploring further on which f(n) could be proven distinct and how they multiply with other elements, then led to the rest. Andy figured that out quite rapidly, so I'm still digesting the final details.

Thanks @Terence Tao for choosing a problem that both allows interesting math proofs but also, as you put it, "experimental mathematics". This has been great fun.

Bruno Le Floch (Nov 01 2024 at 22:54):

Here is a very simple description of the 32-element counterexample. It can be partitioned into six subsets Mi,i=0,,5M_i,i=0,\dots,5 with (5i)\binom{5}{i} elements which appear exactly 3i3^i times in the multiplication table:
M0={1}M_0=\{1\}, M1={0,5,6,9,10}M_1=\{0,5,6,9,10\}, M2=[12,21]M_2=[12,21], M3=[22,31]M_3=[22,31], M4={3,4,7,8,11}M_4=\{3,4,7,8,11\} and M5={2}M_5=\{2\}. These are the same multiplicities as what would happen in the direct product of five copies of the 2-element magma (satisfying 1485). In addition, the product of elements in MiM_i and MjM_j lies in MkM_k with max(5i,5j)kmin(5,10ij)\max(5-i,5-j)\leq k\leq \min(5,10-i-j). This suggests labeling the elements of MM by subsets of Z/5Z\mathbb{Z}/5\mathbb{Z}, and the product having to do with set intersections.

Concretely the map from ad-hoc indices used in the adjacency table of Andy Jiang to subsets can be taken as follows, with each element of MiM_i (listed above) being mapped to a subset with ii elements.

{0 -> {0}, 1 -> {}, 2 -> {0, 1, 2, 3, 4}, 3 -> {0, 2, 3, 4},
 4 -> {0, 1, 2, 3}, 5 -> {2}, 6 -> {3}, 7 -> {0, 1, 2, 4},
 8 -> {0, 1, 3, 4}, 9 -> {4}, 10 -> {1}, 11 -> {1, 2, 3, 4},
 12 -> {0, 2}, 13 -> {0, 3}, 14 -> {2, 4}, 15 -> {1, 3}, 16 -> {1, 4},
  17 -> {2, 3}, 18 -> {0, 4}, 19 -> {0, 1}, 20 -> {1, 2},
 21 -> {3, 4}, 22 -> {0, 2, 4}, 23 -> {0, 2, 3}, 24 -> {0, 1, 2},
 25 -> {2, 3, 4}, 26 -> {1, 2, 4}, 27 -> {0, 3, 4}, 28 -> {0, 1, 3},
 29 -> {1, 2, 3}, 30 -> {0, 1, 4}, 31 -> {1, 3, 4}}

For xMx\in M, or equivalently xZ/5Zx\subset\mathbb{Z}/5\mathbb{Z}, we can define xx^\complement as its complement as a subset of Z/5Z\mathbb{Z}/5\mathbb{Z}, and x+1x+1 and x1x-1 as the result of shifting all elements by one unit modulo 5. Then the operation is

xy=((x+1)(y1)).x\diamond y = ((x+1)\cap(y-1))^\complement .

Let me check that it obeys equation 1485. Importantly, shifts commute with the complement so we have

(x(zy))=(x+1)(z(y2))(x\diamond(z\diamond y)) = (x+1)^\complement\cup(z\cap(y-2))

(yx)(x(zy))=((y+2)x)(x((z1)(y3)))=x((y+2)(z1)(y3)).(y\diamond x)\diamond (x\diamond(z\diamond y)) = ((y+2)\cap x) \cup (x\cap((z-1)\cap(y-3))^\complement) = x \cap ((y+2) \cup (z-1)^\complement\cup(y-3)^\complement) .

Since y+2=y3y+2=y-3, the union (y+2)(y3)=Z/5Z(y+2) \cup (y-3)^\complement = \mathbb{Z}/5\mathbb{Z}, so the whole set is xx.

I feel like there should be a generalization to subsets of Z/nZ\mathbb{Z}/n\mathbb{Z}, but the naive generalization xy=((x+k)(yl))x\diamond y = ((x+k)\cap(y-l))^\complement for integers k,lk,l seems to only obey 1485 if klk\equiv l modulo nn and 5k05k\equiv 0 modulo nn, so we just obtain direct products of the 32-element magma, and also the special case n=1n=1 that reduces to the 2-element magma based on NAND. Maybe a way to make progress would be to relabel in terms of subsets the elements of the 4, 8, and 16 element magmas that were found so far.

Andy Jiang (Nov 01 2024 at 23:30):

Amazing! It also explains why multiplication by 2 twice gives automorphism as predicted by theory. Btw right multiplication by 2 once seems to be some kind of anti automorphism right? I couldn't figure out how to write it formally, do you know how?
Edit (I guess it relates it to a twisted NOR gate or something like that)

Andy Jiang (Nov 01 2024 at 23:39):

Especially amazing that written this way it's clearly a generalized NAND gate

Bruno Le Floch (Nov 02 2024 at 01:16):

I also failed to make sense of an anti-automorphism property. Shifts xx+kx\mapsto x+k are automorphisms (even powers of right multiplication by 2). Right multiplication by 2 is x(x+1)x\mapsto (x+1)^\complement, so really the question is whether "complement" is an anti-automorphism. Complement maps intersections to unions, which changes the operation completely. A related simpler question is in what sense NOT is an anti-automorphism of {true,false} equipped with the NAND operation (namely the 2-element magma obeying equation 1485).

Andy Jiang (Nov 02 2024 at 01:19):

Seems like a general procedure for modifying magmas (in this case product of NAND gates) where for each variable which appears more than once you look at how it occurs.y appears on left hand side as both a left, left and right, right, right so we count it as a -2 and a +3. x occurs on left hand size as a left,right and right, left so 0 for both and on right side as a singleton so 0 as well. when you have a modulus in which all the numbers for a given variable is same, it seems like you can twist?

Andy Jiang (Nov 02 2024 at 01:20):

Bruno Le Floch said:

I also failed to make sense of an anti-automorphism property. Shifts xx+kx\mapsto x+k are automorphisms (even powers of right multiplication by 2). Right multiplication by 2 is x(x+1)x\mapsto (x+1)^\complement, so really the question is whether "complement" is an anti-automorphism. Complement maps intersections to unions, which changes the operation completely. A related simpler question is in what sense NOT is an anti-automorphism of {true,false} equipped with the NAND operation (namely the 2-element magma obeying equation 1485).

right I think it's not an anti-automorphism but rather it is isomorphism with NOR?

Kevin M (Nov 02 2024 at 01:24):

oh, in circuits NOR and NAND are a kind of "inverse" of each other

Bruno Le Floch (Nov 02 2024 at 01:25):

Andy Jiang said:

Seems like a general procedure for modifying magmas (in this case product of NAND gates) where for each variable which appears more than once you look at how it occurs.y appears on left hand side as both a left, left and right, right, right so we count it as a -2 and a +3. x occurs on left hand size as a left,right and right, left so 0 for both and on right side as a singleton so 0 as well. when you have a modulus in which all the numbers for a given variable is same, it seems like you can twist?

Good point. Actually twisting can be slightly more general, you don't have to weight left/right with $+1$ and $-1$. Here the shifts have to be opposite because $x$ appears as left-right (and right-left) and without operation, but for other equations it is not necessarily the case.

Kevin M (Nov 02 2024 at 01:34):

I don't know what the correct word for the NOR and NAND relationship, but its kind of an "involution", where applied twice you get back to the original:
https://en.wikipedia.org/wiki/NAND_logic#/media/File:NOR_from_NAND.svg
https://en.wikipedia.org/wiki/NOR_logic#/media/File:NAND_from_NOR.svg
f(x,y) -> g(x,y) = f(f(f(x,x),f(y,y)), f(f(x,x),f(y,y)))
g(x,y) -> f(x,y) = g(g(g(x,x),g(y,y)), g(g(x,x),g(y,y)))

Kevin M (Nov 02 2024 at 01:42):

EDIT: rephrasing, but still not really useful

The 2 element magma can be viewed as NOR or NAND depending on labeling (and multiplying by the equivalent of the '2' here does that permutation). So these are isomorphic, but not a symmetry of the table, and applying the permutation twice is just the identity permutation -- therefore in this small case the "shift twice" symmetry is trivial.

But in the 32x32 case, applying just a single shift gives us a "dual" operation that is just a permutation, while applying it twice shows an automorphism. Its a 'special' permutation in that it is kind of the sqrt of an automorphism. So not just a generic relabelling, but still not sure how to capture what the "dual" means.

Terence Tao (Nov 02 2024 at 05:37):

Here is one way to think about @Bruno Le Floch interpretation of the model. If MM is a weak central groupoid that happens to have an automorphism T:MMT: M \to M of order 55, then one can "twist" the weak central groupoid to create a new weak central groupoid MTM^T that has the same carrier as MM, but with a new magma operation xTy:=(Tx)(T1y)x \diamond^T y := (Tx) \diamond (T^{-1} y). One can check that (yTx)T(xT(zTy))=(T2yx)(x(T2zT3y)(y \diamond^T x) \diamond^T (x \diamond^T (z \diamond^T y)) = (T^2 y \diamond x) \diamond (x \diamond (T^{-2} z \diamond T^{-3} y) and so the weak central groupoid property is preserved since T2=T3T^2 = T^{-3}.

If one takes the fifth Cartesian power M25M_2^{5} of the order 2 weak central groupoid M2M_2, this acquires an order 5 automorphism (the cyclic shift) and so on twisting this fifth power by this shift one obtains the order 32 example.

The order 5 of the automorphism is also related to the 5-cycle in the graph theory interpretation of the weak central groupoid law.

Bruno Le Floch (Nov 02 2024 at 09:21):

Terence Tao said:

[…] new magma operation xTy:=(Tx)(T1y)x \diamond^T y := (Tx) \diamond (T^{-1} y) […]

For more general equations, twisting can involve two (not necessarily invertible) homomorphisms T,UT,U instead of T,T1T,T^{-1}. The condition is that the composition of TT and UU corresponding to each position of a given variable is the same. For instance, for an equation ((xx)x=xx((x\diamond x)\diamond x=x\diamond x we can take TT=TU=T=UTT=TU=T=U, such as T=U ⁣:(x,y)(x,x)T=U\colon (x,y)\mapsto (x,x) on M×MM\times M.

The procedure maps an equation to a semigroup with two generators, and I think it can be an organizing principle: if an equation implies another one, then all twistings T,UT,U of the first equation have to be valid twistings of the second, so that there has to be a (doubly-pointed) semigroup morphism from the second semigroup to the first semigroup (doubly-pointed in the sense of keeping T,UT,U) if I didn't mess up the direction. In this language, equation 1485 has (additive) semigroup (Z/5Z,+,1,1)(\mathbb{Z}/5\mathbb{Z},+,1,-1), while equation 151 has (additive) semigroup (Z/2Z,+,1,1)(\mathbb{Z}/2\mathbb{Z},+,1,1), which have no doubly-pointed semigroup morphisms.

Equation 854 semigroup T=1,U3=U2\langle T = 1, U^3=U^2\rangle. Equation 413 semigroup T=1,U4=1\langle T = 1, U^4=1\rangle. I think that this means 854 does not imply 413, which until recently was open (maybe still is, I don't really understand where to look). This should also give rather concrete counterexamples by picking specifically U ⁣:M2M2U\colon M^2\to M^2 with U(x,y)=(x,x)U(x,y)=(x,x), but this seems too simple so I am a bit suspicious of my logic or calculations there.

Andy Jiang (Nov 02 2024 at 10:15):

I don't really get the logic at the end. I think you know that the universal 413 model will fail to be 413 after the twisting but how do you know that a model coming from a 854 will fail? what does this argument say in the case the first equation is like x=y?

Andy Jiang (Nov 02 2024 at 10:44):

Sorry maybe now I see what you mean, like do the twist on the universal 854 model and if it implies 413 it would also imply like some twisted version of 413 which presumably will be way too stringent like it will probably have more free variables

Amir Livne Bar-on (Nov 02 2024 at 11:05):

I'm trying to write code for this analysis, and trying to understand why 854 and 413 don't result in the trivial semigroup. For example, 413 starts with x=x(x) x = x \diamond (x \diamond \ldots) . Why doesn't this imply I=T=TU I = T = TU ?

Amir Livne Bar-on (Nov 02 2024 at 11:48):

My code is here: relations.html
(I'll add a version to the repo once I figure out reduction to a minimal set of relations)

Bruno Le Floch (Nov 02 2024 at 12:17):

Amir Livne Bar-on said:

I'm trying to write code for this analysis, and trying to understand why 854 and 413 don't result in the trivial semigroup. For example, 413 starts with x=x(x) x = x \diamond (x \diamond \ldots) . Why doesn't this imply I=T=TU I = T = TU ?

I messed up indeed, the semigroups for 854 and 413 are trivial. In fact, most semigroups are trivial so it is not such a good organizing principle, it's a bit specific to 1485 and a few other equations.

Bruno Le Floch (Nov 02 2024 at 12:19):

Andy Jiang said:

Sorry maybe now I see what you mean, like do the twist on the universal 854 model and if it implies 413 it would also imply like some twisted version of 413 which presumably will be way too stringent like it will probably have more free variables

I think we can have a lot of control of the twist, by finding TT, UU of the form (x1,,xk)(xi1,,xik)(x_1,\dots,x_k)\mapsto(x_{i_1},\dots,x_{i_k}) on MkM^k, and then presumably we don't need to invoke the universal models, but simply any magma with more than one element.

Andy Jiang (Nov 02 2024 at 13:09):

sorry I agree the universal model thing is not really that relevant. I guess we're in agreement that the argument will be of the form that if 854 => 413 it actually implies some stronger form of 413 (presumably the one that does have this extra symmetry).

Amir Livne Bar-on (Nov 02 2024 at 13:16):

Here is a version with some very ugly code to reduce relation: relations.html

Of the open equations, most have a trivial semigroup. But 1323 and 2744 have non-trivial and different ones, and also 917 has a non-trivial semigroup (it's an unresolved implication from 1729 that has a trivial one).

Amir Livne Bar-on (Nov 02 2024 at 13:45):

On second thought, these three semigoups are also trivial, it's just not proved by shortlex reduction

Amir Livne Bar-on (Nov 02 2024 at 13:50):

(deleted)

Andy Jiang (Nov 02 2024 at 14:23):

Amir Livne Bar-on said:

On second thought, these three semigoups are also trivial, it's just not proved by shortlex reduction

I don't see why for example 1516 has trivial symmetry can you explain?

Amir Livne Bar-on (Nov 02 2024 at 14:26):

We have 1=TU=TUU 1 = TU = TUU . Reducing TUU gets us U=1U=1, and T=1T=1 follows

Andy Jiang (Nov 02 2024 at 14:27):

but they don't have to be invertible right if I'm understanding Bruno correctly

Andy Jiang (Nov 02 2024 at 14:28):

oh sorry

Andy Jiang (Nov 02 2024 at 14:28):

couldn't think lol

Amir Livne Bar-on (Nov 02 2024 at 14:28):

Right, but TUU=TUU=UTUU=TU \cdot U =U, and then 1=TU=T 1=TU=T

Andy Jiang (Nov 02 2024 at 14:30):

I guess this must be related to one of Terry's immunities somehow...

Andy Jiang (Nov 02 2024 at 14:32):

you checked 1729 as well?

Andy Jiang (Nov 02 2024 at 14:32):

I continue to not be able to think so I still can't see why that one is trivial...

Andy Jiang (Nov 02 2024 at 14:37):

ok I managed...RLR=1,RR=RLL=LL=RL implies R=RLRR=RLRL=L, so RRR=1 and RR=RRR=1 so R=1

Amir Livne Bar-on (Nov 02 2024 at 14:47):

It looks like it should be possible to turn the difficulty of this problem to the implication question, but I couldn't figure out how

Andy Jiang (Nov 02 2024 at 14:47):

I don't follow what you mean sorry. like use the implication graph to determine the symmetry?

Amir Livne Bar-on (Nov 02 2024 at 14:51):

To reduce the question of whether one law implies another to comparing semigoups define by relations (which is a non-computable problem)

Andy Jiang (Nov 02 2024 at 14:56):

oh you mean like b/c word problem is noncomputable or something? therefore implication is? but like word problem if it does imply search will find it, you just never know when to halt?

Amir Livne Bar-on (Nov 02 2024 at 14:57):

Implication is also like this, by the completeness theorem

Amir Livne Bar-on (Nov 02 2024 at 14:57):

So yes

Terence Tao (Nov 02 2024 at 15:06):

Perhaps this semigroup method would have provided an alternate resolution of 1117->2441, which was one of the final fifteen equations.

Terence Tao (Nov 02 2024 at 15:17):

I've updated the blueprint at equational#779 to add the twist-based proof of 1485!->151.

Amir Livne Bar-on (Nov 02 2024 at 15:26):

Is this contingent on finding appropriate homomorphisms on a model of 1117? Or are this guaranteed to exist on the free magma or something?

Terence Tao (Nov 02 2024 at 15:29):

As long as one has some non-trivial model MM one can raise it to a large Cartesian power MXM^{X}, where XX is a carrier for some model of the semigroup (I think this always exists by Cayley's theorem for semigroups, or Yoneda's lemma if you want to be fancy). The functions on XX give magma endomorphisms on MXM^{X} by pullback, so one can twist by any semigroup. (In the 1485 example, the carrier XX was Z/5Z{\mathbb Z}/5{\mathbb Z}, and the semigroup {T=U1;T5=1}\{ T = U^{-1}; T^5 = 1 \} was represented by Ti=i+1,Ui=i1Ti = i+1, Ui = i-1.)

Kevin Buzzard (Nov 02 2024 at 15:59):

So am I right in saying that you now have a counterexample to one of the implications which you have explained conceptually with a comprehensible-to-humans construction but which was initially discovered by a clever combination of machine search and human guidance? And now we also understand much better the role of 5 in all this?

Amir Livne Bar-on (Nov 02 2024 at 16:41):

Equations 713 and 1447 (of "large greedy" fame) have non-trivial LFT semigroups these are trivial:

  • 713: 1=UT=U3,T=T2=TU \langle 1 = UT = U^3, T = T^2 = TU \rangle
  • 1447: 1=T2=TU=U3 \langle 1 = T^2 = TU = U^3 \rangle

Bernhard Reinke (Nov 02 2024 at 16:45):

I think these are also trivial. In both cases they are actually groups, and you can see that the order of the generators is 1.

Bernhard Reinke (Nov 02 2024 at 16:47):

1 = U^3 implies that U is invertible, then you get that T is its inverse, and then you can compute the order

Terence Tao (Nov 02 2024 at 16:47):

I think the best chances for having a non-trivial semigroup is when there are many variables with few repeats. So the method is unlikely to be too useful for long two-variable equations, though it is still worth trying just in case.

Amir Livne Bar-on (Nov 02 2024 at 16:51):

Thanks for the correction

Matthew Bolan (Nov 02 2024 at 16:53):

Maybe I should just automate the check and see if there are any interesting implications from it.

Matthew Bolan (Nov 02 2024 at 16:54):

It would be interesting at least to know equations with unexpectedly large twisting semigroups.

Matthew Bolan (Nov 02 2024 at 20:31):

A point I missed until my program started telling me false things: If TT and UU are not surjective homomorphisms, then the "for all" quantifier in the twisted version may no longer be a "for all", so we can't always run the argument in this case.

Matthew Bolan (Nov 02 2024 at 20:34):

Also, Amir, I missed that you were already working on automating this. I don't want to duplicate work, so if you're still trying then I'll happily call it quits for the day.

Terence Tao (Nov 02 2024 at 20:36):

Semi-automation may be a more feasible goal than automation - get some heuristics to identify promising equations with potentially large semigroups, but then do some human verification afterwards.

Amir Livne Bar-on (Nov 02 2024 at 20:38):

I've stopped working on it: my goal was to make it easier to check individual cases manually, and I added the simplification only in order for the UI not to show obviously redundant relations. Checking many equations requires interfacing with an ATP, which is not similar at all.

Matthew Bolan (Nov 02 2024 at 20:39):

:thumbs_up: . I'll continue trying to interface with ATP.

Matthew Bolan (Nov 02 2024 at 20:42):

Matthew Bolan said:

A point I missed until my program started telling me false things: If TT and UU are not surjective homomorphisms, then the "for all" quantifier in the twisted version may no longer be a "for all", so we can't always run the argument in this case.

Another silly point is that the equation x=yx=y naively seems like its twisting semigroup should be the free semigroup on T,UT,U, but since the only models of x=yx=y are trivial, it is in fact just the trivial semigroup. I guess this is why Terry said "non-trivial model" in the above.

Andy Jiang (Nov 03 2024 at 00:15):

@Terence Tao Is the failure of semigroup symmetry to solve implications implied by or imply any of the existing immunities?

Terence Tao (Nov 03 2024 at 00:30):

I guess it's related. If the semigroups have a finite model, and the equation has a finite model, then this method should be able to generate finite models, so implications that are immune to finite models are likely immune to the semigroup method. Except that there are some semigroups that only have infinite models, so it's not perfect immunity.

Andy Jiang (Nov 03 2024 at 00:32):

ah ok! I was more thinking like the failure of these ansatz in translation models or models in noncommutative groups/rings (sorry I didn't follow too closely before) could maybe be related to the failure of symmetry here but maybe there's no implication

Terence Tao (Nov 03 2024 at 00:33):

Hmm. Actually if the equation has even a single non-trivial linear model, then this construction will produce another linear model, so if the implication is immune to linear models, it would also be immune to this one.

Terence Tao (Nov 03 2024 at 00:34):

Oh wait, not necessarily

Terence Tao (Nov 03 2024 at 00:34):

because maybe you have to start with a nonlinear model to get a counterexample

Andy Jiang (Nov 03 2024 at 00:36):

ah ok makes sense! so I guess there's no direct implications

Matthew Bolan (Nov 03 2024 at 06:13):

I couldn't work on the twisting monoid as much as I would have liked, so some work remains. However, I managed to get set up enough to answer most implication questions (I've not found a confluent rewriting system for the monoids associated to around 50 equations, but most of them have many variables). I chose to assume that TT and UU themselves are surjective, though in the future this maybe can be relaxed to just asking for those expressions which appear among the relations to be surjective.

Remembering the markings, 110 distinct monoids arise among the equations for which I found confluent rewriting systems. The largest finite one seems to be order 8, though some interesting infinite ones arise as well. A little over half of the magmas we consider have trivial twisting monoids.

Among single variable equations the situation is especially simple, equation 1 is free on T,UT,U, equation 151 is (Z/2Z,+,1,1)(\mathbb Z / 2 \mathbb Z,+,1,1), and all others are trivial.

Among 3 variable equations, equations 898, 2116, 2316 (and their duals) have one of order 7, and I notice that the only non-trivial law they are implied by is 895, which there has been some discussion of already given its relation to Tarski's axiom. There is only a single equation with 3 or fewer variables my program did not find a rewriting system for, which was 3520.

Tomorrow I will cross reference the implications this thing suggests against the ones we know as a sanity check, and then hopefully share some pretty pictures of Hasse diagrams. If there is a list of implications not yet formalized I can check against that too, perhaps we can save someone some work.

Terence Tao (Nov 03 2024 at 06:18):

The implications not formalized are at https://github.com/teorth/equational_theories/blob/main/equational_theories/Conjectures.lean .
Could you explain a bit more the connection between the twisting monoid and confluent rewriting systems? I didn't realize they were related.

Matthew Bolan (Nov 03 2024 at 06:21):

Well, to figure out if I'm looking at a monoid homomorphism, I need to determine if the relations in the source monoid get sent to something trivial in the target. I found it difficult to do this with my usual tools (when I initially offered to do this I assumed sage or gap would take care of it, but I ran into enough issues that I gave up), so eventually I just treated it as a string rewriting problem. As far as I can tell, the only connection to the confluent rewriting systems used in the original project is that we are both solving a string rewriting problem.

Matthew Bolan (Nov 03 2024 at 06:53):

Terence Tao said:

The implications not formalized are at https://github.com/teorth/equational_theories/blob/main/equational_theories/Conjectures.lean .
Could you explain a bit more the connection between the twisting monoid and confluent rewriting systems? I didn't realize they were related.

Nothing assuming no bugs :big_frown:. As expected I suppose.

Matthew Bolan (Nov 03 2024 at 07:42):

Your earlier comment about this giving an alternate resolution to 1117-> 2441 is correct. For 1117, the twisting monoid is (Z,+,1,1)(\mathbb Z,+,1,-1), whereas for 2441 it is trivial. If we ask for just TUTU and UTUT to be surjective and not TT and UU, the original counterexample to this implication can be viewed as already having been of this form! On the other hand, if I try and require surjectivity of TT and UU it is not so easy to find a 1117 magma for which the functional equation in T,UT,U needed to satisfy 2441 is violated, so probably my earlier assumption lost something.

I checked the other equations listed in the outstanding equations threads and they all seemed like they won't work.

Terence Tao (Nov 03 2024 at 16:52):

Well, it's nice to see that we can at least unify the 1117 analysis and 1485 analysis under a single method. The twisting technique is a funny one; it seems to only work rarely, but when it does, it can take out equations that almost all other techniques can't touch.

Matthew Bolan (Nov 03 2024 at 17:24):

I'm going to remove any requirements for surjectivity and just rely on hand verification if anything new turns up.

Bruno Le Floch (Nov 05 2024 at 07:15):

We have a conjecture that the cardinal of M|M| is either a square or twice a square. Here is another conjecture that I tested on all known explicit magmas obeying 1485: for a given xx, all non-empty (Lx)1(z)(L_x)^{-1}(z) and (Rx)1(z)(R_x)^{-1}(z) have the same cardinal (which thus divides M|M|).

A good hope for the second conjecture would be to find a map f ⁣:y(y)f\colon y\mapsto (y\diamond\dots)\diamond\dots (possibly involving yy more than once and more operations \diamond) such that xy=xyx\diamond y=x\diamond y' implies xf(y)=xf(y)x\diamond f(y)=x\diamond f(y') and some bijectivity condition, but I won't have time to think more about that today. I also suppose the two conjectures are related.

Jose Brox (Nov 05 2024 at 13:23):

Terence Tao ha dicho:

In terms of matrices, the adjacency matrix has the block form [[A,B], [C,0]] where BC = AB = CA = CB = J, where J is the all ones block. The square of the matrix is then [[J+A^2, J], [J, J]].

Again running out of time, but wanted to share some thoughts regarding the cardinality of weak central groupoids, building on @Terence Tao's matricial description. Please, check if this holds up!

Call M:=[A,B;C,0]M:=[A,B;C,0] with AA square of size kk, MM square of size nn, and write m:=nkm:=n-k (we do not consider the trivial, strong case in which B,CB,C are empty). Since M2=[J+A2,J;J,J]M^2 = [J+A^2,J;J,J], on the one hand for the trace we have t(M2)=t(Jk+A2)+t(Jm)=k+t(A2)+m=n+t(A2)t(M^2) = t(J_k+A^2)+t(J_m) = k+t(A^2)+m = n+t(A^2), while on the other hand t(M2)=t(M)22σ(M),t(M^2) = t(M)^2 - 2\sigma(M), where σ\sigma is the sum of products of two distinct eigenvalues, and t(M)=t(A)t(M)=t(A). Putting all together we get
2(σ(A)σ(M))=n.2(\sigma(A)-\sigma(M)) = n.
Since A,MA,M are binary matrices this already implies that nn must be an even number.

Moreover, since σ(M)\sigma(M) is the sum (modulo a sign) of the principal 22-minors of MM, by inspecting the matrix structure we find that
σ(M)=σ(A)i,jbijcji=σ(A)k,\sigma(M) = \sigma(A) - \sum_{i,j} b_{ij}c_{ji} = \sigma(A) - k,
since BC=JkBC=J_k. Therefore n=2kn = 2k and m=km=k! Therefore if I'm not mistaken, the numbers of low and high vertices must be equal.

Hence we have AB=CA=BC=CB=JkAB=CA=BC=CB=J_k, in particular B,CB,C commute, and in consequence we get JB=BCB=BJJB = BCB = BJ, JC=CJJC = CJ, CJ=C(AB)=(CA)B=JBCJ = C(AB) = (CA)B = JB, AJ=A(BC)=(AB)C=JCAJ = A(BC) = (AB)C = JC, and JA=BJJA = BJ, what implies that
MJ=[AJ,BJ;CJ,0]=[AJ,AJ;AJ,0]=[JA,JA;JA,0]=JM.MJ = [AJ, BJ; CJ, 0] = [AJ, AJ; AJ, 0] = [JA,JA;JA,0]= JM.
I think this should imply that the left and right cosets xGxG, GyGy of the weak central groupoid GG have the same number of elements for every pair of elements (x,y)(x,y).

I have to leave it here for now, I hope this is correct and useful!

Andy Jiang (Nov 05 2024 at 13:34):

Could be wrong but I thought this high low grouping can't be done on all weak central groupoids. Like I think the counterexample to everything is a square is not of this form

Matthew Bolan (Nov 05 2024 at 16:12):

Since the approach's utility seems to be petering out and my "day job" should not be ignored much longer, I'm not going to attempt any nice visualizations of the twisting monoids. That being said, I did manage to get a handle on all of the twisting monoids and the maps between them: twisting_monoids.txt. Each line of the file is a presentation of a twisting monoid, followed by a list of all equations with this twisting monoid, followed by a list of all other equations for which it produces valid twists. To prove an anti-implication a->b, one wants a to appear in the first list, and b to not appear in the second. Not included are the 1496 singleton laws, as the technique doesn't really apply to these. The relations I list should be confluent and terminating rewriting systems, so you can check equality in these monoids by just naively applying the relations until you cannot anymore. I had to handle laws 3537 and 3649 specially, as for these the Knuth Bendix algorithm doesn't find such a rewriting system in the shortlex order, but it does in the shortrevlex order.

Matthew Bolan (Nov 05 2024 at 16:33):

Also, it is not enough just to find a line with a in the first list and b not in the second, you need to actually try and find a magma to twist up, and there may be non-trivial conditions on this magma (See for example my messages in the Austin pairs channel. There it looks like one should be able to prove 1486↛38621486 \not \to 3862 by twisting, but examining equation 38773877 reveals that one can't prove 1486↛38771486 \not \to 3877 by twisting. As 387738623877 \to 3862, if one wishes to disprove 148638621486 \to 3862 by twisting, one needs to begin with a magma that is at least a witness to 1486↛38771486 \not \to 3877 ).

Jose Brox (Nov 06 2024 at 10:10):

Andy Jiang ha dicho:

Could be wrong but I thought this high low grouping can't be done on all weak central groupoids. Like I think the counterexample to everything is a square is not of this form

Oh my, I was under the impression that this matrix condition was equivalent to Lemmas 11.2 etc. from the blueprint, but of course you are right! We know that M^2 has no zeros, and possibly some other combinatorial facts, but we don't have that nice a structure.

What a pity, and thank you Andy!

Amir Livne Bar-on (Nov 12 2024 at 01:56):

Mario Carneiro said:

I proved the generalized greedy theorem in the following form

theorem exists_greedy_chain {α β : Type*} [Preorder α] [Countable β]
    (task : α  β  Prop) (H :  a b,  a', a  a'  task a' b) (a : α) :
     c, IsChain (·  ·) c  ( x  c, a  x)   b,  a  c, task a b := by

Trying to use this theorem now. In my case β\beta is countable too, but I wonder why this is a requirement. Can't we do well-founded (transfinite) induction? This adds the limit case to the analysis, which would be handled with a union of the relevant chains. (And for this we would also need to require in the induction that the partial chains constructed contain one another)

Johan Commelin (Nov 12 2024 at 05:17):

But the countable case should be sufficient in practice, right?

Amir Livne Bar-on (Nov 12 2024 at 05:41):

Yes, so far none of our constructions needed anything beyond that. And it's hard to believe that will ever change, since there's always some trick, like using direct sums instead of direct products.

Johan Commelin (Nov 12 2024 at 05:51):

If you have a counterexample, then you can always find a countable one.

Johan Commelin (Nov 12 2024 at 05:51):

Take the sub-magma generated by the handful of elements that contradict the other law.

Mario Carneiro (Nov 12 2024 at 07:00):

Amir Livne Bar-on said:

Mario Carneiro said:

I proved the generalized greedy theorem in the following form

theorem exists_greedy_chain {α β : Type*} [Preorder α] [Countable β]
    (task : α  β  Prop) (H :  a b,  a', a  a'  task a' b) (a : α) :
     c, IsChain (·  ·) c  ( x  c, a  x)   b,  a  c, task a b := by

Trying to use this theorem now. In my case β\beta is countable too, but I wonder why this is a requirement. Can't we do well-founded (transfinite) induction? This adds the limit case to the analysis, which would be handled with a union of the relevant chains. (And for this we would also need to require in the induction that the partial chains constructed contain one another)

You kind of answered the question yourself. This theorem is also true for larger cardinals, but you need to assert the existence of upper bounds for every cardinal strictly less than the target bound. It changes the statement of the theorem.

Mario Carneiro (Nov 12 2024 at 07:02):

This additional side condition is not trivial, because it's definitely false if you replace the < with <= in the above claim. So this would add a bunch of unnecessary cardinality reasoning to applications

Mario Carneiro (Nov 12 2024 at 07:03):

Plus as has been observed several times throughout this project, all the action happens over countable sets because the question itself is fundamentally finitely generated

Amir Livne Bar-on (Nov 12 2024 at 16:47):

Another question (asking only because it's interesting, didn't come up in my proof):
The greedy construction is phrased as an existential theorem, which doesn't allow accessing the function itself for any purposes other than proving theorems about it. If it were an explicit function we could also ask Lean to compute and print its values for example. Is there a reason we chose to give it as a theorem?

Mario Carneiro (Nov 12 2024 at 16:50):

There are a lot of things you would want to change to make this a computable function

Mario Carneiro (Nov 12 2024 at 16:50):

it would be exponential time or worse anyway

Mario Carneiro (Nov 12 2024 at 16:53):

It is an existential statement, but it also depends on two other existential statements - The assumption H, and the Countable β assumption, which in the proof is inverted to "there exists a surjection Nat -> Option β"

Mario Carneiro (Nov 12 2024 at 16:56):

Also, even once you are done the thing you get out is a chain, which is a Set α and hence has no computational content. So you would have to change that too to get something useful, and/or turn the ∀ b, ∃ a ∈ c, task a b part into a function

Mario Carneiro (Nov 12 2024 at 16:58):

Put another way: this is a variant of Zorn's lemma, which is noncomputable by nature (it's equivalent to the axiom of choice). It has "arbitrary choices abound" written all over it

Amir Livne Bar-on (Nov 12 2024 at 17:02):

This makes sense, it will be a pain to make all ingredients computable. People have been calculating the greedy construction in some cases but makes sense for this theorem to not require that.

Amir Livne Bar-on (Nov 12 2024 at 17:05):

(in my mind the greedy construction is more similar to Kőnig's lemma, which is somewhat more amenable to computation than Zorn's lemma)

Matthew Bolan (Nov 15 2024 at 13:37):

Pace's question about the existence of a nontrivial group GG with a subset AA such that every element of GG is a unique product of elements of AA has been asked before (at least for finite groups) on MathOverflow: https://mathoverflow.net/questions/384292/is-each-squared-finite-group-trivial .


Last updated: May 02 2025 at 03:31 UTC