Elaborate a FreeUnionConfig.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Given a DecidableEq Var instance, this elaborator automatically constructs
the union of any variables, finite sets of variables, and optionally the
results of provided functions mapping to variables. This is configurable with
optional boolean boolean arguments singleton and finset.
As an example, consider the following:
variable (x : ℕ) (xs : Finset ℕ) (var : String)
def f (_ : String) : Finset ℕ := {1, 2, 3}
def g (_ : String) : Finset ℕ := {4, 5, 6}
-- info: ∅ ∪ {x} ∪ id xs : Finset ℕ
#check free_union ℕ
-- info: ∅ ∪ {x} ∪ id xs ∪ f var ∪ g var : Finset ℕ
#check free_union [f, g] ℕ
info: ∅ ∪ id xs : Finset ℕ
#check free_union (singleton := false) ℕ
-- info: ∅ ∪ {x} : Finset ℕ
#check free_union (finset := false) ℕ
-- info: ∅ : Finset ℕ
#check free_union (singleton := false) (finset := false) ℕ
Equations
- One or more equations did not get rendered due to their size.
Instances For
Elaborator for free_union.
Equations
- One or more equations did not get rendered due to their size.
Instances For
All infinite types have an associated (at least noncomputable) fresh function.
This, in conjunction with HasFresh.not_of_finite, characterizes HasFresh.
Equations
- Cslib.HasFresh.of_infinite α = { fresh := fun (s : Finset α) => Exists.choose ⋯, fresh_notMem := ⋯ }
Instances For
Construct a fresh element given a function f with x < f x.
Equations
Instances For
ℕ has a computable fresh function.
Equations
- Cslib.instHasFreshNat = Cslib.HasFresh.ofSucc (fun (x : ℕ) => x + 1) Nat.lt_add_one
ℤ has a computable fresh function.
Equations
- Cslib.instHasFreshInt = Cslib.HasFresh.ofSucc (fun (x : ℤ) => x + 1) Int.lt_succ
ℚ has a computable fresh function.
Equations
- Cslib.instHasFreshRat = Cslib.HasFresh.ofSucc (fun (x : ℚ) => x + 1) Cslib.instHasFreshRat._proof_1
If α has a computable fresh function, then so does Finset α.
Equations
- Cslib.instHasFreshFinsetOfDecidableEq = Cslib.HasFresh.ofSucc (fun (s : Finset α) => insert (Cslib.fresh s) s) ⋯
If α is inhabited, then Multiset α has a computable fresh function.
Equations
- Cslib.instHasFreshMultisetOfDecidableEqOfInhabited = Cslib.HasFresh.ofSucc (fun (s : Multiset α) => default ::ₘ s) ⋯
ℕ → ℕ has a computable fresh function.
Equations
- Cslib.instHasFreshForallNat = Cslib.HasFresh.ofSucc (fun (f : ℕ → ℕ) (x : ℕ) => f x + 1) Cslib.instHasFreshForallNat._proof_1