logicaffeine_proof/
unify.rs

1//! First-order unification for proof search.
2//!
3//! Implements Robinson's Unification Algorithm with occurs check. This is the
4//! core pattern-matching engine that enables logical reasoning.
5//!
6//! ## What Unification Does
7//!
8//! Unification finds a substitution that makes two terms identical:
9//!
10//! | Pattern | Target | Substitution |
11//! |---------|--------|--------------|
12//! | `Mortal(x)` | `Mortal(Socrates)` | `{x ↦ Socrates}` |
13//! | `Add(Succ(n), 0)` | `Add(Succ(Zero), 0)` | `{n ↦ Zero}` |
14//! | `f(x, x)` | `f(a, b)` | Fails (x can't be both a and b) |
15//!
16//! ## Occurs Check
17//!
18//! The occurs check prevents infinite terms. `x = f(x)` has no finite solution
19//! since it would require `x = f(f(f(...)))`. We reject such unifications.
20//!
21//! ## Alpha-Equivalence
22//!
23//! Bound variable names are arbitrary: `∃e P(e) ≡ ∃x P(x)`. We handle this
24//! by substituting fresh constants for bound variables when unifying
25//! quantified expressions.
26//!
27//! # See Also
28//!
29//! * [`beta_reduce`] - Normalizes lambda applications before unification
30//! * [`ProofTerm`] - The term representation being unified
31//! * [`ProofExpr`] - The expression representation for higher-order unification
32
33use std::collections::HashMap;
34use std::sync::atomic::{AtomicUsize, Ordering};
35
36use crate::error::{ProofError, ProofResult};
37use crate::{MatchArm, ProofExpr, ProofTerm};
38
39/// A substitution mapping variable names to terms.
40///
41/// Substitutions are the output of unification. Applying a substitution
42/// to both sides of a unification problem yields identical terms.
43///
44/// # Example
45///
46/// After unifying `Mortal(x)` with `Mortal(Socrates)`:
47///
48/// ```text
49/// { "x" ↦ Constant("Socrates") }
50/// ```
51///
52/// # Operations
53///
54/// * [`unify_terms`] - Creates substitutions from term unification
55/// * [`apply_subst_to_term`] - Applies substitution to a term
56/// * [`compose_substitutions`] - Combines two substitutions
57pub type Substitution = HashMap<String, ProofTerm>;
58
59/// Substitution for expression-level meta-variables (holes).
60///
61/// Maps hole names to their solutions, typically lambda abstractions
62/// inferred during higher-order pattern unification.
63///
64/// # Example
65///
66/// Solving `?P(x) = Even(x)` yields:
67///
68/// ```text
69/// { "P" ↦ Lambda { variable: "x", body: Even(x) } }
70/// ```
71///
72/// # See Also
73///
74/// * [`unify_pattern`] - Creates expression substitutions via Miller pattern unification
75pub type ExprSubstitution = HashMap<String, ProofExpr>;
76
77// =============================================================================
78// ALPHA-EQUIVALENCE SUPPORT
79// =============================================================================
80
81/// Global counter for generating fresh constants during alpha-renaming.
82static ALPHA_COUNTER: AtomicUsize = AtomicUsize::new(0);
83
84/// Generate a fresh constant for alpha-renaming.
85/// Uses a prefix that cannot appear in user input to avoid collisions.
86fn fresh_alpha_constant() -> ProofTerm {
87    let id = ALPHA_COUNTER.fetch_add(1, Ordering::SeqCst);
88    ProofTerm::Constant(format!("#α{}", id))
89}
90
91// =============================================================================
92// BETA-REDUCTION
93// =============================================================================
94
95/// Check if an expression is a constructor form (safe for fix unfolding).
96/// Only constructors are safe guards against non-termination.
97fn is_constructor_form(expr: &ProofExpr) -> bool {
98    matches!(expr, ProofExpr::Ctor { .. })
99}
100
101/// Beta-reduce an expression to Weak Head Normal Form (WHNF).
102///
103/// Normalizes lambda applications and match expressions, enabling unification
104/// to work on structurally equivalent terms.
105///
106/// # Reduction Rules
107///
108/// | Redex | Reduction |
109/// |-------|-----------|
110/// | `(λx. body)(arg)` | `body[x := arg]` (beta) |
111/// | `(fix f. body) (Ctor ...)` | `body[f := fix f. body] (Ctor ...)` (fix unfolding) |
112/// | `match (Ctor args) { ... }` | Selected arm with args substituted (iota) |
113///
114/// # Weak Head Normal Form
115///
116/// WHNF stops at the outermost constructor or lambda. It does not reduce
117/// under binders, making it efficient for unification purposes.
118///
119/// # Termination
120///
121/// Fix unfolding only occurs when the argument is a constructor (`Ctor`),
122/// ensuring termination for well-typed terms.
123///
124/// # Example
125///
126/// ```text
127/// beta_reduce((λx. P(x))(Socrates))
128/// → P(Socrates)
129///
130/// beta_reduce(match Zero { Zero => True, Succ(k) => False })
131/// → True
132/// ```
133///
134/// # See Also
135///
136/// * [`unify_exprs`] - Calls beta_reduce before comparing expressions
137/// * [`certifier::certify`](crate::certifier::certify) - Uses normalized terms for kernel conversion
138pub fn beta_reduce(expr: &ProofExpr) -> ProofExpr {
139    match expr {
140        // Beta-reduction and Fix unfolding
141        ProofExpr::App(func, arg) => {
142            // First, reduce both function and argument
143            let func_reduced = beta_reduce(func);
144            let arg_reduced = beta_reduce(arg);
145
146            match func_reduced {
147                // Beta-reduction: (λx. body)(arg) → body[x := arg]
148                ProofExpr::Lambda { variable, body } => {
149                    let result = substitute_expr_for_var(&body, &variable, &arg_reduced);
150                    // Recursively reduce the result (handle nested redexes)
151                    beta_reduce(&result)
152                }
153
154                // Fix unfolding: (fix f. body) arg → body[f := fix f. body] arg
155                // Only unfold when arg is a constructor (guard against non-termination)
156                ProofExpr::Fixpoint { ref name, ref body } if is_constructor_form(&arg_reduced) => {
157                    // Substitute fix for f in body
158                    let fix_expr = ProofExpr::Fixpoint {
159                        name: name.clone(),
160                        body: body.clone(),
161                    };
162                    let unfolded = substitute_expr_for_var(body, name, &fix_expr);
163                    // Apply unfolded body to arg and reduce
164                    let applied = ProofExpr::App(Box::new(unfolded), Box::new(arg_reduced));
165                    beta_reduce(&applied)
166                }
167
168                _ => {
169                    // No reduction possible, return normalized application
170                    ProofExpr::App(Box::new(func_reduced), Box::new(arg_reduced))
171                }
172            }
173        }
174
175        // Reduce inside binary connectives
176        ProofExpr::And(l, r) => ProofExpr::And(
177            Box::new(beta_reduce(l)),
178            Box::new(beta_reduce(r)),
179        ),
180        ProofExpr::Or(l, r) => ProofExpr::Or(
181            Box::new(beta_reduce(l)),
182            Box::new(beta_reduce(r)),
183        ),
184        ProofExpr::Implies(l, r) => ProofExpr::Implies(
185            Box::new(beta_reduce(l)),
186            Box::new(beta_reduce(r)),
187        ),
188        ProofExpr::Iff(l, r) => ProofExpr::Iff(
189            Box::new(beta_reduce(l)),
190            Box::new(beta_reduce(r)),
191        ),
192        ProofExpr::Not(inner) => ProofExpr::Not(Box::new(beta_reduce(inner))),
193
194        // Reduce inside quantifiers
195        ProofExpr::ForAll { variable, body } => ProofExpr::ForAll {
196            variable: variable.clone(),
197            body: Box::new(beta_reduce(body)),
198        },
199        ProofExpr::Exists { variable, body } => ProofExpr::Exists {
200            variable: variable.clone(),
201            body: Box::new(beta_reduce(body)),
202        },
203
204        // Reduce inside lambda bodies (but not the lambda itself - that's a value)
205        ProofExpr::Lambda { variable, body } => ProofExpr::Lambda {
206            variable: variable.clone(),
207            body: Box::new(beta_reduce(body)),
208        },
209
210        // Reduce inside modal and temporal operators
211        ProofExpr::Modal { domain, force, flavor, body } => ProofExpr::Modal {
212            domain: domain.clone(),
213            force: *force,
214            flavor: flavor.clone(),
215            body: Box::new(beta_reduce(body)),
216        },
217        ProofExpr::Temporal { operator, body } => ProofExpr::Temporal {
218            operator: operator.clone(),
219            body: Box::new(beta_reduce(body)),
220        },
221
222        // Reduce inside Ctor arguments
223        ProofExpr::Ctor { name, args } => ProofExpr::Ctor {
224            name: name.clone(),
225            args: args.iter().map(beta_reduce).collect(),
226        },
227
228        // Iota reduction: match (Ctor args) with arms → selected arm body
229        ProofExpr::Match { scrutinee, arms } => {
230            let reduced_scrutinee = beta_reduce(scrutinee);
231
232            // Try iota reduction: if scrutinee is a Ctor, select matching arm
233            if let ProofExpr::Ctor { name: ctor_name, args: ctor_args } = &reduced_scrutinee {
234                for arm in arms {
235                    if &arm.ctor == ctor_name {
236                        // Found matching arm - substitute constructor args for bindings
237                        let mut result = arm.body.clone();
238                        for (binding, arg) in arm.bindings.iter().zip(ctor_args.iter()) {
239                            result = substitute_expr_for_var(&result, binding, arg);
240                        }
241                        // Continue reducing the result
242                        return beta_reduce(&result);
243                    }
244                }
245            }
246
247            // No iota reduction possible - just reduce subexpressions
248            ProofExpr::Match {
249                scrutinee: Box::new(reduced_scrutinee),
250                arms: arms.iter().map(|arm| MatchArm {
251                    ctor: arm.ctor.clone(),
252                    bindings: arm.bindings.clone(),
253                    body: beta_reduce(&arm.body),
254                }).collect(),
255            }
256        }
257
258        // Reduce inside Fixpoint
259        ProofExpr::Fixpoint { name, body } => ProofExpr::Fixpoint {
260            name: name.clone(),
261            body: Box::new(beta_reduce(body)),
262        },
263
264        // Atomic expressions don't reduce
265        ProofExpr::Predicate { .. }
266        | ProofExpr::Identity(_, _)
267        | ProofExpr::Atom(_)
268        | ProofExpr::NeoEvent { .. }
269        | ProofExpr::TypedVar { .. }
270        | ProofExpr::Unsupported(_)
271        | ProofExpr::Hole(_)
272        | ProofExpr::Term(_) => expr.clone(),
273    }
274}
275
276/// Substitute an expression for a variable name in another expression.
277///
278/// Used for beta-reduction: (λx. body)(arg) → body[x := arg]
279/// Handles variable capture by not substituting inside shadowing binders.
280fn substitute_expr_for_var(body: &ProofExpr, var: &str, replacement: &ProofExpr) -> ProofExpr {
281    match body {
282        ProofExpr::Predicate { name, args, world } => ProofExpr::Predicate {
283            name: name.clone(),
284            args: args.iter().map(|t| substitute_term_for_var(t, var, replacement)).collect(),
285            world: world.clone(),
286        },
287
288        ProofExpr::Identity(l, r) => ProofExpr::Identity(
289            substitute_term_for_var(l, var, replacement),
290            substitute_term_for_var(r, var, replacement),
291        ),
292
293        ProofExpr::Atom(a) => {
294            // If the atom matches the variable, replace it
295            if a == var {
296                replacement.clone()
297            } else {
298                ProofExpr::Atom(a.clone())
299            }
300        }
301
302        ProofExpr::And(l, r) => ProofExpr::And(
303            Box::new(substitute_expr_for_var(l, var, replacement)),
304            Box::new(substitute_expr_for_var(r, var, replacement)),
305        ),
306        ProofExpr::Or(l, r) => ProofExpr::Or(
307            Box::new(substitute_expr_for_var(l, var, replacement)),
308            Box::new(substitute_expr_for_var(r, var, replacement)),
309        ),
310        ProofExpr::Implies(l, r) => ProofExpr::Implies(
311            Box::new(substitute_expr_for_var(l, var, replacement)),
312            Box::new(substitute_expr_for_var(r, var, replacement)),
313        ),
314        ProofExpr::Iff(l, r) => ProofExpr::Iff(
315            Box::new(substitute_expr_for_var(l, var, replacement)),
316            Box::new(substitute_expr_for_var(r, var, replacement)),
317        ),
318        ProofExpr::Not(inner) => ProofExpr::Not(
319            Box::new(substitute_expr_for_var(inner, var, replacement))
320        ),
321
322        // Quantifiers: don't substitute if the variable is shadowed
323        ProofExpr::ForAll { variable, body: inner } => {
324            if variable == var {
325                // Variable is shadowed, don't substitute in body
326                body.clone()
327            } else {
328                ProofExpr::ForAll {
329                    variable: variable.clone(),
330                    body: Box::new(substitute_expr_for_var(inner, var, replacement)),
331                }
332            }
333        }
334        ProofExpr::Exists { variable, body: inner } => {
335            if variable == var {
336                body.clone()
337            } else {
338                ProofExpr::Exists {
339                    variable: variable.clone(),
340                    body: Box::new(substitute_expr_for_var(inner, var, replacement)),
341                }
342            }
343        }
344
345        // Lambda: don't substitute if the variable is shadowed
346        ProofExpr::Lambda { variable, body: inner } => {
347            if variable == var {
348                body.clone()
349            } else {
350                ProofExpr::Lambda {
351                    variable: variable.clone(),
352                    body: Box::new(substitute_expr_for_var(inner, var, replacement)),
353                }
354            }
355        }
356
357        ProofExpr::App(f, a) => ProofExpr::App(
358            Box::new(substitute_expr_for_var(f, var, replacement)),
359            Box::new(substitute_expr_for_var(a, var, replacement)),
360        ),
361
362        ProofExpr::Modal { domain, force, flavor, body: inner } => ProofExpr::Modal {
363            domain: domain.clone(),
364            force: *force,
365            flavor: flavor.clone(),
366            body: Box::new(substitute_expr_for_var(inner, var, replacement)),
367        },
368
369        ProofExpr::Temporal { operator, body: inner } => ProofExpr::Temporal {
370            operator: operator.clone(),
371            body: Box::new(substitute_expr_for_var(inner, var, replacement)),
372        },
373
374        ProofExpr::NeoEvent { event_var, verb, roles } => {
375            // Substitute in roles, but not the event_var itself
376            ProofExpr::NeoEvent {
377                event_var: event_var.clone(),
378                verb: verb.clone(),
379                roles: roles.iter().map(|(r, t)| {
380                    (r.clone(), substitute_term_for_var(t, var, replacement))
381                }).collect(),
382            }
383        }
384
385        ProofExpr::Ctor { name, args } => ProofExpr::Ctor {
386            name: name.clone(),
387            args: args.iter().map(|a| substitute_expr_for_var(a, var, replacement)).collect(),
388        },
389
390        ProofExpr::Match { scrutinee, arms } => ProofExpr::Match {
391            scrutinee: Box::new(substitute_expr_for_var(scrutinee, var, replacement)),
392            arms: arms.iter().map(|arm| {
393                // Don't substitute if var is bound in this arm
394                if arm.bindings.contains(&var.to_string()) {
395                    arm.clone()
396                } else {
397                    MatchArm {
398                        ctor: arm.ctor.clone(),
399                        bindings: arm.bindings.clone(),
400                        body: substitute_expr_for_var(&arm.body, var, replacement),
401                    }
402                }
403            }).collect(),
404        },
405
406        ProofExpr::Fixpoint { name, body: inner } => {
407            if name == var {
408                body.clone()
409            } else {
410                ProofExpr::Fixpoint {
411                    name: name.clone(),
412                    body: Box::new(substitute_expr_for_var(inner, var, replacement)),
413                }
414            }
415        }
416
417        ProofExpr::TypedVar { .. } | ProofExpr::Unsupported(_) => body.clone(),
418
419        // Holes are meta-variables - don't substitute into them
420        ProofExpr::Hole(_) => body.clone(),
421
422        // Terms: substitute into the inner term
423        ProofExpr::Term(t) => ProofExpr::Term(substitute_term_for_var(t, var, replacement)),
424    }
425}
426
427/// Substitute an expression for a variable in a term.
428///
429/// When a variable in a term matches, convert the replacement expression to a term.
430fn substitute_term_for_var(term: &ProofTerm, var: &str, replacement: &ProofExpr) -> ProofTerm {
431    match term {
432        ProofTerm::Variable(v) if v == var => {
433            // Variable matches, convert replacement to term
434            expr_to_term(replacement)
435        }
436        // BoundVarRef also participates in substitution (for instantiating quantified formulas)
437        ProofTerm::BoundVarRef(v) if v == var => {
438            expr_to_term(replacement)
439        }
440        ProofTerm::Variable(_) | ProofTerm::Constant(_) | ProofTerm::BoundVarRef(_) => term.clone(),
441        ProofTerm::Function(name, args) => ProofTerm::Function(
442            name.clone(),
443            args.iter().map(|a| substitute_term_for_var(a, var, replacement)).collect(),
444        ),
445        ProofTerm::Group(terms) => ProofTerm::Group(
446            terms.iter().map(|t| substitute_term_for_var(t, var, replacement)).collect(),
447        ),
448    }
449}
450
451/// Convert a simple ProofExpr to ProofTerm.
452///
453/// Used during beta-reduction when substituting an expression argument
454/// into a predicate's term position.
455fn expr_to_term(expr: &ProofExpr) -> ProofTerm {
456    match expr {
457        // Atoms become constants
458        ProofExpr::Atom(s) => ProofTerm::Constant(s.clone()),
459
460        // Zero-arity predicates become constants (e.g., "John" as a predicate name)
461        ProofExpr::Predicate { name, args, .. } if args.is_empty() => {
462            ProofTerm::Constant(name.clone())
463        }
464
465        // Predicates with args become functions
466        ProofExpr::Predicate { name, args, .. } => {
467            ProofTerm::Function(name.clone(), args.clone())
468        }
469
470        // Constructors become functions
471        ProofExpr::Ctor { name, args } => {
472            ProofTerm::Function(name.clone(), args.iter().map(expr_to_term).collect())
473        }
474
475        // TypedVar becomes a variable
476        ProofExpr::TypedVar { name, .. } => ProofTerm::Variable(name.clone()),
477
478        // Term is already a term - extract it directly
479        ProofExpr::Term(t) => t.clone(),
480
481        // Fallback: stringify the expression (covers Hole, etc.)
482        _ => ProofTerm::Constant(format!("{}", expr)),
483    }
484}
485
486// =============================================================================
487// TERM-LEVEL UNIFICATION
488// =============================================================================
489
490/// Unify two terms, returning the Most General Unifier (MGU).
491///
492/// The MGU is the smallest substitution that makes both terms identical.
493/// Uses Robinson's algorithm with occurs check.
494///
495/// # Arguments
496///
497/// * `t1` - The first term (often a pattern with variables)
498/// * `t2` - The second term (often a ground term or another pattern)
499///
500/// # Returns
501///
502/// * `Ok(subst)` - A substitution that unifies the terms
503/// * `Err(OccursCheck)` - If unification would create an infinite term
504/// * `Err(SymbolMismatch)` - If function/constant names differ
505/// * `Err(ArityMismatch)` - If argument counts differ
506///
507/// # Example
508///
509/// ```
510/// use logicaffeine_proof::{ProofTerm, unify::unify_terms};
511///
512/// let pattern = ProofTerm::Function(
513///     "Mortal".into(),
514///     vec![ProofTerm::Variable("x".into())]
515/// );
516/// let target = ProofTerm::Function(
517///     "Mortal".into(),
518///     vec![ProofTerm::Constant("Socrates".into())]
519/// );
520///
521/// let subst = unify_terms(&pattern, &target).unwrap();
522/// assert_eq!(
523///     subst.get("x"),
524///     Some(&ProofTerm::Constant("Socrates".into()))
525/// );
526/// ```
527///
528/// # See Also
529///
530/// * [`unify_exprs`] - Unifies expressions (handles quantifiers, connectives)
531/// * [`apply_subst_to_term`] - Applies the resulting substitution
532pub fn unify_terms(t1: &ProofTerm, t2: &ProofTerm) -> ProofResult<Substitution> {
533    let mut subst = Substitution::new();
534    unify_terms_with_subst(t1, t2, &mut subst)?;
535    Ok(subst)
536}
537
538/// Internal unification with accumulating substitution.
539fn unify_terms_with_subst(
540    t1: &ProofTerm,
541    t2: &ProofTerm,
542    subst: &mut Substitution,
543) -> ProofResult<()> {
544    // Apply current substitution to both terms first
545    let t1 = apply_subst_to_term(t1, subst);
546    let t2 = apply_subst_to_term(t2, subst);
547
548    match (&t1, &t2) {
549        // Identical terms unify trivially
550        (ProofTerm::Constant(c1), ProofTerm::Constant(c2)) if c1 == c2 => Ok(()),
551
552        // Different constants cannot unify
553        (ProofTerm::Constant(c1), ProofTerm::Constant(c2)) => {
554            Err(ProofError::SymbolMismatch {
555                left: c1.clone(),
556                right: c2.clone(),
557            })
558        }
559
560        // Variable on the left: bind it to the right term
561        (ProofTerm::Variable(v), t) => {
562            // Check if they're the same variable
563            if let ProofTerm::Variable(v2) = t {
564                if v == v2 {
565                    return Ok(());
566                }
567            }
568            // Occurs check: prevent infinite types
569            if occurs(v, t) {
570                return Err(ProofError::OccursCheck {
571                    variable: v.clone(),
572                    term: t.clone(),
573                });
574            }
575            subst.insert(v.clone(), t.clone());
576            Ok(())
577        }
578
579        // Variable on the right: bind it to the left term
580        (t, ProofTerm::Variable(v)) => {
581            // Occurs check
582            if occurs(v, t) {
583                return Err(ProofError::OccursCheck {
584                    variable: v.clone(),
585                    term: t.clone(),
586                });
587            }
588            subst.insert(v.clone(), t.clone());
589            Ok(())
590        }
591
592        // BoundVarRef on the left: treat like a variable for unification
593        // This allows matching quantified formulas like ∀x P(x) against P(butler)
594        (ProofTerm::BoundVarRef(v), t) => {
595            if let ProofTerm::BoundVarRef(v2) = t {
596                if v == v2 {
597                    return Ok(());
598                }
599            }
600            if occurs(v, t) {
601                return Err(ProofError::OccursCheck {
602                    variable: v.clone(),
603                    term: t.clone(),
604                });
605            }
606            subst.insert(v.clone(), t.clone());
607            Ok(())
608        }
609
610        // BoundVarRef on the right: bind it to the left term
611        (t, ProofTerm::BoundVarRef(v)) => {
612            if occurs(v, t) {
613                return Err(ProofError::OccursCheck {
614                    variable: v.clone(),
615                    term: t.clone(),
616                });
617            }
618            subst.insert(v.clone(), t.clone());
619            Ok(())
620        }
621
622        // Function unification: same name and arity, unify arguments pairwise
623        (ProofTerm::Function(f1, args1), ProofTerm::Function(f2, args2)) => {
624            if f1 != f2 {
625                return Err(ProofError::SymbolMismatch {
626                    left: f1.clone(),
627                    right: f2.clone(),
628                });
629            }
630            if args1.len() != args2.len() {
631                return Err(ProofError::ArityMismatch {
632                    expected: args1.len(),
633                    found: args2.len(),
634                });
635            }
636            for (a1, a2) in args1.iter().zip(args2.iter()) {
637                unify_terms_with_subst(a1, a2, subst)?;
638            }
639            Ok(())
640        }
641
642        // Group unification: same length, unify elements pairwise
643        (ProofTerm::Group(g1), ProofTerm::Group(g2)) => {
644            if g1.len() != g2.len() {
645                return Err(ProofError::ArityMismatch {
646                    expected: g1.len(),
647                    found: g2.len(),
648                });
649            }
650            for (t1, t2) in g1.iter().zip(g2.iter()) {
651                unify_terms_with_subst(t1, t2, subst)?;
652            }
653            Ok(())
654        }
655
656        // Any other combination fails
657        _ => Err(ProofError::UnificationFailed {
658            left: t1,
659            right: t2,
660        }),
661    }
662}
663
664/// Check if a variable occurs in a term (for occurs check).
665/// Prevents infinite types like x = f(x).
666fn occurs(var: &str, term: &ProofTerm) -> bool {
667    match term {
668        ProofTerm::Variable(v) => v == var,
669        ProofTerm::BoundVarRef(v) => v == var, // BoundVarRef participates in occurs check
670        ProofTerm::Constant(_) => false,
671        ProofTerm::Function(_, args) => args.iter().any(|a| occurs(var, a)),
672        ProofTerm::Group(terms) => terms.iter().any(|t| occurs(var, t)),
673    }
674}
675
676/// Apply a substitution to a term.
677///
678/// Replaces all variables in the term according to the substitution mapping.
679/// Handles transitive chains: if `{x ↦ y, y ↦ z}`, then applying to `x` yields `z`.
680///
681/// # Arguments
682///
683/// * `term` - The term to transform
684/// * `subst` - The substitution mapping variables to replacement terms
685///
686/// # Returns
687///
688/// A new term with all substitutions applied.
689///
690/// # Example
691///
692/// ```
693/// use logicaffeine_proof::{ProofTerm, unify::{Substitution, apply_subst_to_term}};
694/// use std::collections::HashMap;
695///
696/// let mut subst = Substitution::new();
697/// subst.insert("x".into(), ProofTerm::Constant("Socrates".into()));
698///
699/// let term = ProofTerm::Function(
700///     "Mortal".into(),
701///     vec![ProofTerm::Variable("x".into())]
702/// );
703///
704/// let result = apply_subst_to_term(&term, &subst);
705/// // result = Mortal(Socrates)
706/// ```
707pub fn apply_subst_to_term(term: &ProofTerm, subst: &Substitution) -> ProofTerm {
708    match term {
709        ProofTerm::Variable(v) => {
710            if let Some(replacement) = subst.get(v) {
711                // Recursively apply to handle chains like {x ↦ y, y ↦ z}
712                apply_subst_to_term(replacement, subst)
713            } else {
714                term.clone()
715            }
716        }
717        // BoundVarRef participates in substitution (for instantiating quantified formulas)
718        ProofTerm::BoundVarRef(v) => {
719            if let Some(replacement) = subst.get(v) {
720                apply_subst_to_term(replacement, subst)
721            } else {
722                term.clone()
723            }
724        }
725        ProofTerm::Constant(_) => term.clone(),
726        ProofTerm::Function(name, args) => {
727            let new_args = args.iter().map(|a| apply_subst_to_term(a, subst)).collect();
728            ProofTerm::Function(name.clone(), new_args)
729        }
730        ProofTerm::Group(terms) => {
731            let new_terms = terms.iter().map(|t| apply_subst_to_term(t, subst)).collect();
732            ProofTerm::Group(new_terms)
733        }
734    }
735}
736
737// =============================================================================
738// EXPRESSION-LEVEL UNIFICATION
739// =============================================================================
740
741/// Unify two expressions, returning the Most General Unifier.
742///
743/// Expression unification extends term unification with support for logical
744/// connectives, quantifiers, and alpha-equivalence for bound variables.
745///
746/// # Alpha-Equivalence
747///
748/// Bound variable names are considered arbitrary:
749/// - `∀x P(x)` unifies with `∀y P(y)`
750/// - `∃e Run(e)` unifies with `∃x Run(x)`
751///
752/// This is achieved by substituting fresh constants for bound variables
753/// before comparing bodies.
754///
755/// # Beta-Reduction
756///
757/// Both expressions are beta-reduced before comparison, so
758/// `(λx. P(x))(a)` will unify with `P(a)`.
759///
760/// # Arguments
761///
762/// * `e1` - The first expression
763/// * `e2` - The second expression
764///
765/// # Returns
766///
767/// * `Ok(subst)` - A substitution unifying the expressions
768/// * `Err(ExprUnificationFailed)` - If expressions cannot be unified
769///
770/// # See Also
771///
772/// * [`unify_terms`] - Underlying term unification
773/// * [`beta_reduce`] - Normalization applied before unification
774/// * [`unify_pattern`] - Higher-order pattern unification for holes
775pub fn unify_exprs(e1: &ProofExpr, e2: &ProofExpr) -> ProofResult<Substitution> {
776    let mut subst = Substitution::new();
777    unify_exprs_with_subst(e1, e2, &mut subst)?;
778    Ok(subst)
779}
780
781/// Internal expression unification with accumulating substitution.
782fn unify_exprs_with_subst(
783    e1: &ProofExpr,
784    e2: &ProofExpr,
785    subst: &mut Substitution,
786) -> ProofResult<()> {
787    // Beta-reduce both expressions before unification.
788    // This normalizes lambda applications: (λx. P(x))(a) → P(a)
789    let e1 = beta_reduce(e1);
790    let e2 = beta_reduce(e2);
791
792    match (&e1, &e2) {
793        // Atom unification
794        (ProofExpr::Atom(a1), ProofExpr::Atom(a2)) if a1 == a2 => Ok(()),
795
796        // Predicate unification: same name, unify arguments
797        (
798            ProofExpr::Predicate { name: n1, args: a1, world: w1 },
799            ProofExpr::Predicate { name: n2, args: a2, world: w2 },
800        ) => {
801            if n1 != n2 {
802                return Err(ProofError::SymbolMismatch {
803                    left: n1.clone(),
804                    right: n2.clone(),
805                });
806            }
807            if a1.len() != a2.len() {
808                return Err(ProofError::ArityMismatch {
809                    expected: a1.len(),
810                    found: a2.len(),
811                });
812            }
813            // Unify worlds if both present
814            match (w1, w2) {
815                (Some(w1), Some(w2)) if w1 != w2 => {
816                    return Err(ProofError::SymbolMismatch {
817                        left: w1.clone(),
818                        right: w2.clone(),
819                    });
820                }
821                _ => {}
822            }
823            // Unify arguments
824            for (t1, t2) in a1.iter().zip(a2.iter()) {
825                unify_terms_with_subst(t1, t2, subst)?;
826            }
827            Ok(())
828        }
829
830        // Identity unification
831        (ProofExpr::Identity(l1, r1), ProofExpr::Identity(l2, r2)) => {
832            unify_terms_with_subst(l1, l2, subst)?;
833            unify_terms_with_subst(r1, r2, subst)?;
834            Ok(())
835        }
836
837        // Binary operators: same operator, unify both sides
838        (ProofExpr::And(l1, r1), ProofExpr::And(l2, r2))
839        | (ProofExpr::Or(l1, r1), ProofExpr::Or(l2, r2))
840        | (ProofExpr::Implies(l1, r1), ProofExpr::Implies(l2, r2))
841        | (ProofExpr::Iff(l1, r1), ProofExpr::Iff(l2, r2)) => {
842            unify_exprs_with_subst(l1, l2, subst)?;
843            unify_exprs_with_subst(r1, r2, subst)?;
844            Ok(())
845        }
846
847        // Negation
848        (ProofExpr::Not(inner1), ProofExpr::Not(inner2)) => {
849            unify_exprs_with_subst(inner1, inner2, subst)
850        }
851
852        // Quantifiers: Alpha-equivalence - bound variable names are arbitrary
853        // ∃e P(e) ≡ ∃x P(x) because they describe the same logical content
854        (
855            ProofExpr::ForAll { variable: v1, body: b1 },
856            ProofExpr::ForAll { variable: v2, body: b2 },
857        )
858        | (
859            ProofExpr::Exists { variable: v1, body: b1 },
860            ProofExpr::Exists { variable: v2, body: b2 },
861        ) => {
862            // Generate a fresh constant to substitute for both bound variables.
863            // Using a constant (not a variable) avoids capture issues.
864            let fresh = fresh_alpha_constant();
865
866            // Create substitutions for each bound variable
867            let subst1: Substitution = [(v1.clone(), fresh.clone())].into_iter().collect();
868            let subst2: Substitution = [(v2.clone(), fresh)].into_iter().collect();
869
870            // Apply substitutions to bodies
871            let body1_renamed = apply_subst_to_expr(b1, &subst1);
872            let body2_renamed = apply_subst_to_expr(b2, &subst2);
873
874            // Recursively unify the renamed bodies
875            unify_exprs_with_subst(&body1_renamed, &body2_renamed, subst)
876        }
877
878        // Lambda expressions: Alpha-equivalence - λx.P(x) ≡ λy.P(y)
879        (
880            ProofExpr::Lambda { variable: v1, body: b1 },
881            ProofExpr::Lambda { variable: v2, body: b2 },
882        ) => {
883            // Same alpha-renaming technique as quantifiers
884            let fresh = fresh_alpha_constant();
885            let subst1: Substitution = [(v1.clone(), fresh.clone())].into_iter().collect();
886            let subst2: Substitution = [(v2.clone(), fresh)].into_iter().collect();
887            let body1_renamed = apply_subst_to_expr(b1, &subst1);
888            let body2_renamed = apply_subst_to_expr(b2, &subst2);
889            unify_exprs_with_subst(&body1_renamed, &body2_renamed, subst)
890        }
891
892        // Application
893        (ProofExpr::App(f1, a1), ProofExpr::App(f2, a2)) => {
894            unify_exprs_with_subst(f1, f2, subst)?;
895            unify_exprs_with_subst(a1, a2, subst)?;
896            Ok(())
897        }
898
899        // NeoEvent: Alpha-equivalence for event variables
900        // ∃e(Run(e) ∧ Agent(e, John)) should unify with ∃x(Run(x) ∧ Agent(x, John))
901        (
902            ProofExpr::NeoEvent {
903                event_var: e1,
904                verb: v1,
905                roles: r1,
906            },
907            ProofExpr::NeoEvent {
908                event_var: e2,
909                verb: v2,
910                roles: r2,
911            },
912        ) => {
913            // Verb names must match (case-insensitive for robustness)
914            if v1.to_lowercase() != v2.to_lowercase() {
915                return Err(ProofError::SymbolMismatch {
916                    left: v1.clone(),
917                    right: v2.clone(),
918                });
919            }
920
921            // Roles must have same length
922            if r1.len() != r2.len() {
923                return Err(ProofError::ArityMismatch {
924                    expected: r1.len(),
925                    found: r2.len(),
926                });
927            }
928
929            // Alpha-equivalence: generate fresh constant for event variable
930            let fresh = fresh_alpha_constant();
931            let subst1: Substitution = [(e1.clone(), fresh.clone())].into_iter().collect();
932            let subst2: Substitution = [(e2.clone(), fresh)].into_iter().collect();
933
934            // Unify roles pairwise with alpha-renamed event variables
935            for ((role1, term1), (role2, term2)) in r1.iter().zip(r2.iter()) {
936                // Role names must match
937                if role1 != role2 {
938                    return Err(ProofError::SymbolMismatch {
939                        left: role1.clone(),
940                        right: role2.clone(),
941                    });
942                }
943                // Apply alpha-renaming to terms and unify
944                let t1_renamed = apply_subst_to_term(term1, &subst1);
945                let t2_renamed = apply_subst_to_term(term2, &subst2);
946                unify_terms_with_subst(&t1_renamed, &t2_renamed, subst)?;
947            }
948            Ok(())
949        }
950
951        // Temporal operators: Past(P), Future(P)
952        // Same operator required, then unify bodies
953        (
954            ProofExpr::Temporal { operator: op1, body: b1 },
955            ProofExpr::Temporal { operator: op2, body: b2 },
956        ) => {
957            if op1 != op2 {
958                return Err(ProofError::ExprUnificationFailed {
959                    left: e1.clone(),
960                    right: e2.clone(),
961                });
962            }
963            unify_exprs_with_subst(b1, b2, subst)
964        }
965
966        // Anything else fails
967        _ => Err(ProofError::ExprUnificationFailed {
968            left: e1.clone(),
969            right: e2.clone(),
970        }),
971    }
972}
973
974/// Apply a substitution to an expression.
975///
976/// Recursively replaces variables in terms within the expression according
977/// to the substitution mapping. Also handles binder renaming when the
978/// substitution maps a bound variable name to a new variable.
979///
980/// # Arguments
981///
982/// * `expr` - The expression to transform
983/// * `subst` - The substitution mapping variables to replacement terms
984///
985/// # Returns
986///
987/// A new expression with all substitutions applied to embedded terms.
988///
989/// # Binder Handling
990///
991/// When a quantifier or lambda binds a variable that maps to another variable
992/// in the substitution, the binder is renamed:
993///
994/// ```text
995/// apply_subst_to_expr(∀x P(x), {x ↦ y}) = ∀y P(y)
996/// ```
997pub fn apply_subst_to_expr(expr: &ProofExpr, subst: &Substitution) -> ProofExpr {
998    match expr {
999        ProofExpr::Predicate { name, args, world } => ProofExpr::Predicate {
1000            name: name.clone(),
1001            args: args.iter().map(|a| apply_subst_to_term(a, subst)).collect(),
1002            world: world.clone(),
1003        },
1004        ProofExpr::Identity(l, r) => ProofExpr::Identity(
1005            apply_subst_to_term(l, subst),
1006            apply_subst_to_term(r, subst),
1007        ),
1008        ProofExpr::Atom(a) => ProofExpr::Atom(a.clone()),
1009        ProofExpr::And(l, r) => ProofExpr::And(
1010            Box::new(apply_subst_to_expr(l, subst)),
1011            Box::new(apply_subst_to_expr(r, subst)),
1012        ),
1013        ProofExpr::Or(l, r) => ProofExpr::Or(
1014            Box::new(apply_subst_to_expr(l, subst)),
1015            Box::new(apply_subst_to_expr(r, subst)),
1016        ),
1017        ProofExpr::Implies(l, r) => ProofExpr::Implies(
1018            Box::new(apply_subst_to_expr(l, subst)),
1019            Box::new(apply_subst_to_expr(r, subst)),
1020        ),
1021        ProofExpr::Iff(l, r) => ProofExpr::Iff(
1022            Box::new(apply_subst_to_expr(l, subst)),
1023            Box::new(apply_subst_to_expr(r, subst)),
1024        ),
1025        ProofExpr::Not(inner) => ProofExpr::Not(Box::new(apply_subst_to_expr(inner, subst))),
1026        ProofExpr::ForAll { variable, body } => {
1027            // If the variable is being renamed (maps to a Variable), update the binder
1028            let new_variable = match subst.get(variable) {
1029                Some(ProofTerm::Variable(new_name)) => new_name.clone(),
1030                _ => variable.clone(),
1031            };
1032            ProofExpr::ForAll {
1033                variable: new_variable,
1034                body: Box::new(apply_subst_to_expr(body, subst)),
1035            }
1036        }
1037        ProofExpr::Exists { variable, body } => {
1038            // If the variable is being renamed (maps to a Variable), update the binder
1039            let new_variable = match subst.get(variable) {
1040                Some(ProofTerm::Variable(new_name)) => new_name.clone(),
1041                _ => variable.clone(),
1042            };
1043            ProofExpr::Exists {
1044                variable: new_variable,
1045                body: Box::new(apply_subst_to_expr(body, subst)),
1046            }
1047        }
1048        ProofExpr::Modal { domain, force, flavor, body } => ProofExpr::Modal {
1049            domain: domain.clone(),
1050            force: *force,
1051            flavor: flavor.clone(),
1052            body: Box::new(apply_subst_to_expr(body, subst)),
1053        },
1054        ProofExpr::Temporal { operator, body } => ProofExpr::Temporal {
1055            operator: operator.clone(),
1056            body: Box::new(apply_subst_to_expr(body, subst)),
1057        },
1058        ProofExpr::Lambda { variable, body } => {
1059            // If the variable is being renamed (maps to a Variable), update the binder
1060            let new_variable = match subst.get(variable) {
1061                Some(ProofTerm::Variable(new_name)) => new_name.clone(),
1062                _ => variable.clone(),
1063            };
1064            ProofExpr::Lambda {
1065                variable: new_variable,
1066                body: Box::new(apply_subst_to_expr(body, subst)),
1067            }
1068        }
1069        ProofExpr::App(f, a) => ProofExpr::App(
1070            Box::new(apply_subst_to_expr(f, subst)),
1071            Box::new(apply_subst_to_expr(a, subst)),
1072        ),
1073        ProofExpr::NeoEvent { event_var, verb, roles } => ProofExpr::NeoEvent {
1074            event_var: event_var.clone(),
1075            verb: verb.clone(),
1076            roles: roles
1077                .iter()
1078                .map(|(r, t)| (r.clone(), apply_subst_to_term(t, subst)))
1079                .collect(),
1080        },
1081        // Peano / Inductive Types
1082        ProofExpr::Ctor { name, args } => ProofExpr::Ctor {
1083            name: name.clone(),
1084            args: args.iter().map(|a| apply_subst_to_expr(a, subst)).collect(),
1085        },
1086        ProofExpr::Match { scrutinee, arms } => ProofExpr::Match {
1087            scrutinee: Box::new(apply_subst_to_expr(scrutinee, subst)),
1088            arms: arms
1089                .iter()
1090                .map(|arm| MatchArm {
1091                    ctor: arm.ctor.clone(),
1092                    bindings: arm.bindings.clone(),
1093                    body: apply_subst_to_expr(&arm.body, subst),
1094                })
1095                .collect(),
1096        },
1097        ProofExpr::Fixpoint { name, body } => ProofExpr::Fixpoint {
1098            name: name.clone(),
1099            body: Box::new(apply_subst_to_expr(body, subst)),
1100        },
1101        ProofExpr::TypedVar { name, typename } => ProofExpr::TypedVar {
1102            name: name.clone(),
1103            typename: typename.clone(),
1104        },
1105        ProofExpr::Unsupported(s) => ProofExpr::Unsupported(s.clone()),
1106        // Holes are meta-variables - term substitution doesn't apply
1107        ProofExpr::Hole(name) => ProofExpr::Hole(name.clone()),
1108        // Terms: apply substitution to the inner term
1109        ProofExpr::Term(t) => ProofExpr::Term(apply_subst_to_term(t, subst)),
1110    }
1111}
1112
1113/// Compose two substitutions: apply s2 after s1.
1114///
1115/// The resulting substitution applies s1 first, then s2. This is the standard
1116/// composition operation for Most General Unifiers.
1117///
1118/// # Semantics
1119///
1120/// For any term t: `apply(compose(s1, s2), t) = apply(s2, apply(s1, t))`
1121///
1122/// # Arguments
1123///
1124/// * `s1` - The first substitution (applied first)
1125/// * `s2` - The second substitution (applied second)
1126///
1127/// # Returns
1128///
1129/// A combined substitution equivalent to applying s1 then s2.
1130///
1131/// # Example
1132///
1133/// ```
1134/// use logicaffeine_proof::{ProofTerm, unify::{Substitution, compose_substitutions}};
1135///
1136/// let mut s1 = Substitution::new();
1137/// s1.insert("x".into(), ProofTerm::Variable("y".into()));
1138///
1139/// let mut s2 = Substitution::new();
1140/// s2.insert("y".into(), ProofTerm::Constant("a".into()));
1141///
1142/// let composed = compose_substitutions(s1, s2);
1143/// // x ↦ a (via y), y ↦ a
1144/// ```
1145pub fn compose_substitutions(s1: Substitution, s2: Substitution) -> Substitution {
1146    let mut result: Substitution = s1
1147        .into_iter()
1148        .map(|(k, v)| (k, apply_subst_to_term(&v, &s2)))
1149        .collect();
1150
1151    // Add bindings from s2 that aren't in s1
1152    for (k, v) in s2 {
1153        result.entry(k).or_insert(v);
1154    }
1155
1156    result
1157}
1158
1159// =============================================================================
1160// HIGHER-ORDER PATTERN UNIFICATION
1161// =============================================================================
1162
1163/// Attempt higher-order pattern unification (Miller patterns).
1164///
1165/// Given `lhs` and `rhs`, if `lhs` is of the form `Hole(h)(args...)` where
1166/// args are distinct `BoundVarRef`s, solves: `h = λargs. rhs`.
1167///
1168/// # Miller Patterns
1169///
1170/// A Miller pattern is a meta-variable applied to distinct bound variables:
1171/// - `?P(x)` is a valid pattern
1172/// - `?F(x, y)` is a valid pattern (x ≠ y)
1173/// - `?G(x, x)` is NOT valid (duplicate variable)
1174/// - `?H(f(x))` is NOT valid (non-variable argument)
1175///
1176/// # Why This Matters
1177///
1178/// Higher-order pattern unification is decidable and has unique most general
1179/// solutions, unlike full higher-order unification. This restriction enables
1180/// automatic motive inference for structural induction.
1181///
1182/// # Returns
1183///
1184/// * `Ok(subst)` - An [`ExprSubstitution`] mapping hole names to lambdas
1185/// * `Err(PatternNotDistinct)` - If pattern has duplicate variables
1186/// * `Err(NotAPattern)` - If arguments aren't bound variable references
1187/// * `Err(ScopeViolation)` - If RHS uses variables not in pattern scope
1188///
1189/// # Example
1190///
1191/// ```text
1192/// unify_pattern(?P(x), Even(x))
1193/// → { "P" ↦ λx. Even(x) }
1194/// ```
1195///
1196/// # See Also
1197///
1198/// * [`ExprSubstitution`] - The result type for hole solutions
1199/// * [`ProofError::PatternNotDistinct`] - Duplicate variable error
1200pub fn unify_pattern(lhs: &ProofExpr, rhs: &ProofExpr) -> ProofResult<ExprSubstitution> {
1201    let mut solution = ExprSubstitution::new();
1202    unify_pattern_internal(lhs, rhs, &mut solution)?;
1203    Ok(solution)
1204}
1205
1206/// Internal pattern unification with accumulating solution.
1207fn unify_pattern_internal(
1208    lhs: &ProofExpr,
1209    rhs: &ProofExpr,
1210    solution: &mut ExprSubstitution,
1211) -> ProofResult<()> {
1212    // Beta-reduce both sides first
1213    let lhs = beta_reduce(lhs);
1214    let rhs = beta_reduce(rhs);
1215
1216    match &lhs {
1217        // Case: Bare hole ?P = rhs
1218        ProofExpr::Hole(h) => {
1219            solution.insert(h.clone(), rhs.clone());
1220            Ok(())
1221        }
1222
1223        // Case: Application - might be Hole(h)(args...)
1224        ProofExpr::App(_, _) => {
1225            // Collect all arguments and find the head
1226            let (head, args) = collect_app_args(&lhs);
1227
1228            if let ProofExpr::Hole(h) = head {
1229                // Check Miller pattern: all args must be distinct BoundVarRefs
1230                let var_args = extract_distinct_vars(&args)?;
1231
1232                // Check that rhs only uses variables from var_args (scope check)
1233                check_scope(&rhs, &var_args)?;
1234
1235                // Construct solution: h = λargs. rhs
1236                // Need to rename variables in rhs to match the BoundVarRef names
1237                let renamed_rhs = rename_vars_to_bound(&rhs, &var_args);
1238                let lambda = build_lambda(var_args, renamed_rhs);
1239                solution.insert(h.clone(), lambda);
1240                Ok(())
1241            } else {
1242                // Not a pattern, try structural equality
1243                if lhs == rhs {
1244                    Ok(())
1245                } else {
1246                    Err(ProofError::ExprUnificationFailed {
1247                        left: lhs.clone(),
1248                        right: rhs.clone(),
1249                    })
1250                }
1251            }
1252        }
1253
1254        // Other cases: structural equality
1255        _ => {
1256            if lhs == rhs {
1257                Ok(())
1258            } else {
1259                Err(ProofError::ExprUnificationFailed {
1260                    left: lhs.clone(),
1261                    right: rhs.clone(),
1262                })
1263            }
1264        }
1265    }
1266}
1267
1268/// Collect f(a)(b)(c) into (f, [a, b, c])
1269fn collect_app_args(expr: &ProofExpr) -> (ProofExpr, Vec<ProofExpr>) {
1270    let mut args = Vec::new();
1271    let mut current = expr.clone();
1272
1273    while let ProofExpr::App(func, arg) = current {
1274        args.push(*arg);
1275        current = *func;
1276    }
1277
1278    args.reverse();
1279    (current, args)
1280}
1281
1282/// Extract distinct variable names from pattern arguments.
1283/// Fails if any arg is not a Term(BoundVarRef) or if duplicates exist.
1284fn extract_distinct_vars(args: &[ProofExpr]) -> ProofResult<Vec<String>> {
1285    let mut vars = Vec::new();
1286    for arg in args {
1287        match arg {
1288            ProofExpr::Term(ProofTerm::BoundVarRef(v)) => {
1289                if vars.contains(v) {
1290                    return Err(ProofError::PatternNotDistinct(v.clone()));
1291                }
1292                vars.push(v.clone());
1293            }
1294            _ => return Err(ProofError::NotAPattern(arg.clone())),
1295        }
1296    }
1297    Ok(vars)
1298}
1299
1300/// Check that all free variables in expr are in the allowed set.
1301fn check_scope(expr: &ProofExpr, allowed: &[String]) -> ProofResult<()> {
1302    let free_vars = collect_free_vars(expr);
1303    for var in free_vars {
1304        if !allowed.contains(&var) {
1305            return Err(ProofError::ScopeViolation {
1306                var,
1307                allowed: allowed.to_vec(),
1308            });
1309        }
1310    }
1311    Ok(())
1312}
1313
1314/// Collect free variables from an expression.
1315fn collect_free_vars(expr: &ProofExpr) -> Vec<String> {
1316    let mut vars = Vec::new();
1317    collect_free_vars_impl(expr, &mut vars, &mut Vec::new());
1318    vars
1319}
1320
1321fn collect_free_vars_impl(expr: &ProofExpr, vars: &mut Vec<String>, bound: &mut Vec<String>) {
1322    match expr {
1323        ProofExpr::Predicate { args, .. } => {
1324            for arg in args {
1325                collect_free_vars_term(arg, vars, bound);
1326            }
1327        }
1328        ProofExpr::Identity(l, r) => {
1329            collect_free_vars_term(l, vars, bound);
1330            collect_free_vars_term(r, vars, bound);
1331        }
1332        ProofExpr::Atom(s) => {
1333            if !bound.contains(s) && !vars.contains(s) {
1334                vars.push(s.clone());
1335            }
1336        }
1337        ProofExpr::And(l, r)
1338        | ProofExpr::Or(l, r)
1339        | ProofExpr::Implies(l, r)
1340        | ProofExpr::Iff(l, r) => {
1341            collect_free_vars_impl(l, vars, bound);
1342            collect_free_vars_impl(r, vars, bound);
1343        }
1344        ProofExpr::Not(inner) => collect_free_vars_impl(inner, vars, bound),
1345        ProofExpr::ForAll { variable, body }
1346        | ProofExpr::Exists { variable, body }
1347        | ProofExpr::Lambda { variable, body } => {
1348            bound.push(variable.clone());
1349            collect_free_vars_impl(body, vars, bound);
1350            bound.pop();
1351        }
1352        ProofExpr::App(f, a) => {
1353            collect_free_vars_impl(f, vars, bound);
1354            collect_free_vars_impl(a, vars, bound);
1355        }
1356        ProofExpr::Term(t) => collect_free_vars_term(t, vars, bound),
1357        ProofExpr::Hole(_) => {} // Holes don't contribute free vars
1358        _ => {} // Other cases don't add free vars
1359    }
1360}
1361
1362fn collect_free_vars_term(term: &ProofTerm, vars: &mut Vec<String>, bound: &[String]) {
1363    match term {
1364        ProofTerm::Variable(v) => {
1365            if !bound.contains(v) && !vars.contains(v) {
1366                vars.push(v.clone());
1367            }
1368        }
1369        ProofTerm::Function(_, args) => {
1370            for arg in args {
1371                collect_free_vars_term(arg, vars, bound);
1372            }
1373        }
1374        ProofTerm::Group(terms) => {
1375            for t in terms {
1376                collect_free_vars_term(t, vars, bound);
1377            }
1378        }
1379        ProofTerm::Constant(_) | ProofTerm::BoundVarRef(_) => {}
1380    }
1381}
1382
1383/// Rename Variable(x) to Variable(x) if x is in the bound vars list.
1384/// This ensures the solution lambda binds the right names.
1385fn rename_vars_to_bound(expr: &ProofExpr, bound_vars: &[String]) -> ProofExpr {
1386    // For the basic case, we don't need to rename since the RHS already uses
1387    // Variable("x") and we want the lambda to bind "x".
1388    // The key is just to use the same names.
1389    expr.clone()
1390}
1391
1392/// Build λx₁.λx₂...λxₙ. body
1393fn build_lambda(vars: Vec<String>, body: ProofExpr) -> ProofExpr {
1394    vars.into_iter().rev().fold(body, |acc, var| {
1395        ProofExpr::Lambda {
1396            variable: var,
1397            body: Box::new(acc),
1398        }
1399    })
1400}
1401
1402#[cfg(test)]
1403mod tests {
1404    use super::*;
1405
1406    #[test]
1407    fn test_unify_same_constant() {
1408        let t1 = ProofTerm::Constant("a".into());
1409        let t2 = ProofTerm::Constant("a".into());
1410        let result = unify_terms(&t1, &t2);
1411        assert!(result.is_ok());
1412        assert!(result.unwrap().is_empty());
1413    }
1414
1415    #[test]
1416    fn test_unify_different_constants() {
1417        let t1 = ProofTerm::Constant("a".into());
1418        let t2 = ProofTerm::Constant("b".into());
1419        let result = unify_terms(&t1, &t2);
1420        assert!(result.is_err());
1421    }
1422
1423    #[test]
1424    fn test_unify_var_constant() {
1425        let t1 = ProofTerm::Variable("x".into());
1426        let t2 = ProofTerm::Constant("a".into());
1427        let result = unify_terms(&t1, &t2);
1428        assert!(result.is_ok());
1429        let subst = result.unwrap();
1430        assert_eq!(subst.get("x"), Some(&ProofTerm::Constant("a".into())));
1431    }
1432
1433    #[test]
1434    fn test_occurs_check() {
1435        let t1 = ProofTerm::Variable("x".into());
1436        let t2 = ProofTerm::Function("f".into(), vec![ProofTerm::Variable("x".into())]);
1437        let result = unify_terms(&t1, &t2);
1438        assert!(matches!(result, Err(ProofError::OccursCheck { .. })));
1439    }
1440
1441    #[test]
1442    fn test_compose_substitutions() {
1443        let mut s1 = Substitution::new();
1444        s1.insert("x".into(), ProofTerm::Variable("y".into()));
1445
1446        let mut s2 = Substitution::new();
1447        s2.insert("y".into(), ProofTerm::Constant("a".into()));
1448
1449        let composed = compose_substitutions(s1, s2);
1450
1451        // x should map to a (via y)
1452        assert_eq!(composed.get("x"), Some(&ProofTerm::Constant("a".into())));
1453        // y should also map to a
1454        assert_eq!(composed.get("y"), Some(&ProofTerm::Constant("a".into())));
1455    }
1456
1457    // =========================================================================
1458    // ALPHA-EQUIVALENCE TESTS
1459    // =========================================================================
1460
1461    #[test]
1462    fn test_alpha_equivalence_exists() {
1463        // ∃e P(e) should unify with ∃x P(x)
1464        let e1 = ProofExpr::Exists {
1465            variable: "e".to_string(),
1466            body: Box::new(ProofExpr::Predicate {
1467                name: "run".to_string(),
1468                args: vec![ProofTerm::Variable("e".to_string())],
1469                world: None,
1470            }),
1471        };
1472
1473        let e2 = ProofExpr::Exists {
1474            variable: "x".to_string(),
1475            body: Box::new(ProofExpr::Predicate {
1476                name: "run".to_string(),
1477                args: vec![ProofTerm::Variable("x".to_string())],
1478                world: None,
1479            }),
1480        };
1481
1482        let result = unify_exprs(&e1, &e2);
1483        assert!(
1484            result.is_ok(),
1485            "Alpha-equivalent expressions should unify: {:?}",
1486            result
1487        );
1488    }
1489
1490    #[test]
1491    fn test_alpha_equivalence_forall() {
1492        // ∀x P(x) should unify with ∀y P(y)
1493        let e1 = ProofExpr::ForAll {
1494            variable: "x".to_string(),
1495            body: Box::new(ProofExpr::Predicate {
1496                name: "mortal".to_string(),
1497                args: vec![ProofTerm::Variable("x".to_string())],
1498                world: None,
1499            }),
1500        };
1501
1502        let e2 = ProofExpr::ForAll {
1503            variable: "y".to_string(),
1504            body: Box::new(ProofExpr::Predicate {
1505                name: "mortal".to_string(),
1506                args: vec![ProofTerm::Variable("y".to_string())],
1507                world: None,
1508            }),
1509        };
1510
1511        let result = unify_exprs(&e1, &e2);
1512        assert!(
1513            result.is_ok(),
1514            "Alpha-equivalent universals should unify: {:?}",
1515            result
1516        );
1517    }
1518
1519    #[test]
1520    fn test_alpha_equivalence_nested() {
1521        // ∃e (Run(e) ∧ Agent(e, John)) should unify with ∃x (Run(x) ∧ Agent(x, John))
1522        let e1 = ProofExpr::Exists {
1523            variable: "e".to_string(),
1524            body: Box::new(ProofExpr::And(
1525                Box::new(ProofExpr::Predicate {
1526                    name: "run".to_string(),
1527                    args: vec![ProofTerm::Variable("e".to_string())],
1528                    world: None,
1529                }),
1530                Box::new(ProofExpr::Predicate {
1531                    name: "agent".to_string(),
1532                    args: vec![
1533                        ProofTerm::Variable("e".to_string()),
1534                        ProofTerm::Constant("John".to_string()),
1535                    ],
1536                    world: None,
1537                }),
1538            )),
1539        };
1540
1541        let e2 = ProofExpr::Exists {
1542            variable: "x".to_string(),
1543            body: Box::new(ProofExpr::And(
1544                Box::new(ProofExpr::Predicate {
1545                    name: "run".to_string(),
1546                    args: vec![ProofTerm::Variable("x".to_string())],
1547                    world: None,
1548                }),
1549                Box::new(ProofExpr::Predicate {
1550                    name: "agent".to_string(),
1551                    args: vec![
1552                        ProofTerm::Variable("x".to_string()),
1553                        ProofTerm::Constant("John".to_string()),
1554                    ],
1555                    world: None,
1556                }),
1557            )),
1558        };
1559
1560        let result = unify_exprs(&e1, &e2);
1561        assert!(
1562            result.is_ok(),
1563            "Nested alpha-equivalent expressions should unify: {:?}",
1564            result
1565        );
1566    }
1567}