Skip to content

Commit e6dd412

Browse files
authored
feat: upstream ReduceEval instances from quote4 (#10563)
This PR moves some `ReduceEval` instances about basic types up from the `quote4` library.
1 parent cfc46ac commit e6dd412

File tree

1 file changed

+84
-0
lines changed

1 file changed

+84
-0
lines changed

src/Lean/Meta/ReduceEval.lean

Lines changed: 84 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -63,4 +63,88 @@ private partial def evalName (e : Expr) : MetaM Name := do
6363
instance : ReduceEval Name where
6464
reduceEval := private evalName
6565

66+
private partial def evalList [ReduceEval α] (e : Expr) : MetaM (List α) := do
67+
let e ← whnf e
68+
let .const c _ := e.getAppFn | throwFailedToEval e
69+
let nargs := e.getAppNumArgs
70+
match c, nargs with
71+
| ``List.nil, 1 => pure []
72+
| ``List.cons, 3 => return (← reduceEval (e.getArg! 1)) :: (← evalList (e.getArg! 2))
73+
| _, _ => throwFailedToEval e
74+
75+
instance [ReduceEval α] : ReduceEval (List α) where
76+
reduceEval := private evalList
77+
78+
instance [NeZero n] : ReduceEval (Fin n) where
79+
reduceEval := private fun e => do
80+
let e ← whnf e
81+
if e.isAppOfArity ``Fin.mk 3 then
82+
return Fin.ofNat _ (← reduceEval (e.getArg! 1))
83+
else
84+
throwFailedToEval e
85+
86+
instance {n : Nat} : ReduceEval (BitVec n) where
87+
reduceEval := private fun e => do
88+
let e ← whnf e
89+
if e.isAppOfArity ``BitVec.ofFin 2 then
90+
have : 2^n - 1 + 1 = 2^n := Nat.sub_one_add_one_eq_of_pos (Nat.two_pow_pos n)
91+
let _ : ReduceEval (Fin (2^n)) := this ▸ (inferInstanceAs <| ReduceEval (Fin (2^n - 1 + 1)))
92+
pure ⟨(← reduceEval (e.getArg! 1))⟩
93+
else
94+
throwFailedToEval e
95+
96+
instance : ReduceEval Bool where
97+
reduceEval := private fun e => do
98+
let e ← whnf e
99+
if e.isAppOf ``true then
100+
pure true
101+
else if e.isAppOf ``false then
102+
pure false
103+
else
104+
throwFailedToEval e
105+
106+
instance : ReduceEval BinderInfo where
107+
reduceEval := private fun e => do
108+
match (← whnf e).constName? with
109+
| some ``BinderInfo.default => pure .default
110+
| some ``BinderInfo.implicit => pure .implicit
111+
| some ``BinderInfo.strictImplicit => pure .strictImplicit
112+
| some ``BinderInfo.instImplicit => pure .instImplicit
113+
| _ => throwFailedToEval e
114+
115+
instance : ReduceEval Literal where
116+
reduceEval := private fun e => do
117+
let e ← whnf e
118+
if e.isAppOfArity ``Literal.natVal 1 then
119+
return .natVal (← reduceEval (e.getArg! 0))
120+
else if e.isAppOfArity ``Literal.strVal 1 then
121+
return .strVal (← reduceEval (e.getArg! 0))
122+
else
123+
throwFailedToEval e
124+
125+
instance : ReduceEval MVarId where
126+
reduceEval e := private do
127+
let e ← whnf e
128+
if e.isAppOfArity ``MVarId.mk 1 then
129+
return ⟨← reduceEval (e.getArg! 0)⟩
130+
else
131+
throwFailedToEval e
132+
133+
instance : ReduceEval LevelMVarId where
134+
reduceEval e := private do
135+
let e ← whnf e
136+
if e.isAppOfArity ``LevelMVarId.mk 1 then
137+
return ⟨← reduceEval (e.getArg! 0)⟩
138+
else
139+
throwFailedToEval e
140+
141+
instance : ReduceEval FVarId where
142+
reduceEval e := private do
143+
let e ← whnf e
144+
if e.isAppOfArity ``FVarId.mk 1 then
145+
return ⟨← reduceEval (e.getArg! 0)⟩
146+
else
147+
throwFailedToEval e
148+
149+
66150
end Lean.Meta

0 commit comments

Comments
 (0)