Skip to content

Commit

Permalink
chore: upgrade to Lean v4.7.0
Browse files Browse the repository at this point in the history
  • Loading branch information
alexkeizer committed May 17, 2024
1 parent 219d397 commit 789630c
Show file tree
Hide file tree
Showing 12 changed files with 71 additions and 98 deletions.
20 changes: 9 additions & 11 deletions Qpf/Builtins/List.lean
Original file line number Diff line number Diff line change
Expand Up @@ -6,7 +6,7 @@ import Qpf.Macro.Data
import Qpf.Qpf.Multivariate.ofPolynomial
import Qpf.Util

namespace MvQPF
namespace MvQPF
namespace List

def ListPFunctor : MvPFunctor.{u} 1
Expand All @@ -16,15 +16,15 @@ namespace List
abbrev QpfList' := ListPFunctor.Obj
abbrev List' := @TypeFun.ofCurried 1 List

abbrev box {Γ} (x : List' Γ) : QpfList' Γ
abbrev box {Γ} (x : List' Γ) : QpfList' Γ
:= ⟨
ULift.up x.length,
ULift.up x.length,
fun .fz j => Vec.ofList x (PFin2.toFin2 j)

abbrev unbox {Γ} (x : QpfList' Γ) : List' Γ
abbrev unbox {Γ} (x : QpfList' Γ) : List' Γ
:= Vec.toList fun i => x.snd 0 (PFin2.ofFin2 i)

private theorem typeext {α} {f g : α → Sort _} (f_eq_g: f = g) :
((a : α) → f a) = ((a : α) → g a) :=
by
Expand All @@ -34,7 +34,7 @@ namespace List
instance : MvFunctor List' :=
MvFunctor.ofIsomorphism _ box unbox

instance : MvQPF List' :=
instance : MvQPF List' :=
.ofIsomorphism _ box unbox (
by
intros Γ x;
Expand All @@ -59,19 +59,17 @@ namespace List
apply HEq.trans cast_fun_arg
rfl

case H₃ =>
case H₃ =>
apply typeext;
fin_destr;

simp only [ListPFunctor, TypeVec.Arrow, DVec];
have : List.length (unbox { fst := { down := n }, snd := v }) = n; {
simp
}
have : List.length (unbox { fst := { down := n }, snd := v }) = n := by simp
simp[this]

case H₄ => intros; rfl
) (
by
by
intros _ x;
induction x
. rfl
Expand Down
2 changes: 1 addition & 1 deletion Qpf/Macro/Comp.lean
Original file line number Diff line number Diff line change
Expand Up @@ -178,7 +178,7 @@ partial def elabQpf {arity : Nat} (vars : Vector Q(Type u) arity) (target : Q(Ty
| none => throwError "Free variable {target} is not one of the qpf arguments"
| some ind => pure ind

let ind : Fin2 arity := cast (by simp) ind.inv
let ind : Fin2 arity := cast (by simp [vars']) ind.inv
let prj := q(@Prj.{u} $arity $ind)
trace[QPF] "represented by: {prj}"
pure ⟨prj, q(Prj.mvfunctor _), q(Prj.mvqpf _)⟩
Expand Down
29 changes: 5 additions & 24 deletions Qpf/Macro/Data.lean
Original file line number Diff line number Diff line change
Expand Up @@ -150,23 +150,6 @@ def mkHeadT (view : InductiveView) : CommandElabM Name := do
elabInductiveViews #[view]
pure declName


open Parser in
private def matchAltsOfArray (matchAlts : Array Syntax) : Syntax :=
mkNode ``Term.matchAlts #[mkNullNode matchAlts]


open Parser in
/--
Wraps an array of `matchAltExpr` syntax objects into a single `Command.declValEqns` node, for
use in inductive definitions
-/
private def declValEqnsOfMatchAltArray (matchAlts : Array Syntax) : TSyntax ``Command.declValEqns :=
let body := matchAltsOfArray matchAlts
let body := mkNode ``Term.matchAltsWhereDecls #[body, mkNullNode]
mkNode ``Command.declValEqns #[body]


open Parser Parser.Term Parser.Command in
/--
Defines the "child" family of type vectors for an `n`-ary polynomial functor
Expand All @@ -190,11 +173,11 @@ def mkChildT (view : InductiveView) (r : Replace) (headTName : Name) : CommandEl

`(matchAltExpr| | $head => (!![ $counts,* ]))

let body := declValEqnsOfMatchAltArray matchAlts
let body ← `(declValEqns|
$[$matchAlts:matchAlt]*
)
let headT := mkIdent headTName



let cmd ← `(
def $declId : $headT → $target_type
$body:declValEqns
Expand Down Expand Up @@ -403,11 +386,10 @@ def mkShape (view: DataView) : TermElabM MkShapeResult := do
let childTId := mkIdent childTName

elabCommand <|<- `(
def $PDeclId :=
def $PDeclId:declId :=
MvPFunctor.mk $headTId $childTId
)


mkQpf view ctorArgs headTId PId r.expr.size

Expand Down Expand Up @@ -540,8 +522,7 @@ def mkConstructors (view : DataView) (shape : Name) : CommandElabM Unit := do

let cmd ← `(
$(quote modifiers):declModifiers
def $declId:ident : $type
:= $body:term
def $declId:ident : $type := $body:term
)

trace[QPF] "mkConstructor.cmd = {cmd}"
Expand Down
18 changes: 9 additions & 9 deletions Qpf/Macro/Tactic/FinDestr.lean
Original file line number Diff line number Diff line change
Expand Up @@ -12,7 +12,7 @@ open Lean Syntax Elab Elab.Tactic Meta

def elabFinDestrAux (i_stx : TSyntax `ident) : TacticM Unit := do
let n ← mkFreshExprMVar (mkConst ``Nat) (kind:=MetavarKind.synthetic);

try {
let u ← mkFreshLevelMVar;
let us := [u]
Expand All @@ -38,7 +38,7 @@ def elabFinDestrAux (i_stx : TSyntax `ident) : TacticM Unit := do
if let some n := n.natLit? then
let rec genTactic : Nat → TacticM (TSyntax `tactic)
| 0 => `(tactic| cases $i_stx:ident)
| n+1 => do
| n+1 => do
let tct ← genTactic n
`(tactic| (cases $i_stx:ident; swap; rename_i $i_stx:ident; $tct:tactic))

Expand All @@ -47,7 +47,7 @@ def elabFinDestrAux (i_stx : TSyntax `ident) : TacticM Unit := do

else
let rec genTacticExpr : Expr → TacticM (Option <| TSyntax `tactic)
| Expr.const ``Nat.zero .. =>
| Expr.const ``Nat.zero .. =>
`(tactic| cases $i_stx:ident)

| Expr.app (Expr.const ``Nat.succ ..) n .. => do
Expand All @@ -66,26 +66,26 @@ elab "fin_destr_one " i:ident : tactic => do
withMainContext <|
elabFinDestrAux i

syntax "fin_destr' " ident* : tactic
syntax "fin_destr' " ident* : tactic
macro_rules
| `(tactic| fin_destr' $i:ident $is:ident*) => `(tactic| fin_destr_one $i <;> dsimp (config := {failIfUnchanged := false}) <;> fin_destr' $is:ident*)
| `(tactic| fin_destr') => `(tactic| skip)

syntax "fin_destr " ident* : tactic
syntax "fin_destr " ident* : tactic
macro_rules
| `(tactic| fin_destr $i:ident $is:ident*) => `(tactic|
fin_destr' $i:ident $is:ident*
| `(tactic| fin_destr $i:ident $is:ident*) => `(tactic|
fin_destr' $i:ident $is:ident*
<;> try fin_destr
)

| `(tactic| fin_destr) => `(tactic|
| `(tactic| fin_destr) => `(tactic|
first
| intro i; fin_destr i
| funext i; fin_destr i
)


syntax "vec_eq " (tactic)? : tactic
syntax "vec_eq " (tactic)? : tactic
macro_rules
| `(tactic| vec_eq) => `(tactic| vec_eq trivial)
| `(tactic| vec_eq $tct:tactic ) => `(tactic| funext i; fin_destr i <;> $tct:tactic)
1 change: 0 additions & 1 deletion Qpf/MathlibPort/Fin2.lean
Original file line number Diff line number Diff line change
Expand Up @@ -5,7 +5,6 @@ Authors: Mario Carneiro
-/

import Mathlib.Data.Fin.Fin2
import Mathlib.Data.Nat.Order.Basic
import Mathlib.Init.Order.Defs

/-!
Expand Down
10 changes: 5 additions & 5 deletions Qpf/PFunctor/Multivariate/Constructions/Prod.lean
Original file line number Diff line number Diff line change
Expand Up @@ -12,9 +12,9 @@ namespace Prod

open PFin2 (fz fs)

def P : MvPFunctor 2
def P : MvPFunctor 2
:= .mk Unit fun _ _ => PFin2 1


abbrev QpfProd' := P.Obj
abbrev QpfProd := TypeFun.curried QpfProd'
Expand All @@ -31,8 +31,8 @@ abbrev Prod' : TypeFun 2
-/
def mk (a : Γ 1) (b : Γ 0) : QpfProd' Γ
:= ⟨
(),
fun
(),
fun
| 1, _ => a
| 0, _ => b
Expand All @@ -55,7 +55,7 @@ instance : MvFunctor Prod' where
map f x := equiv.invFun <| P.map f <| equiv.toFun <| x

instance : MvQPF.IsPolynomial Prod' := .ofEquiv _ equiv


end Prod

Expand Down
35 changes: 17 additions & 18 deletions Qpf/Qpf/Multivariate/Constructions/Comp.lean
Original file line number Diff line number Diff line change
Expand Up @@ -14,12 +14,12 @@ import Qpf.Util

universe u


namespace MvQPF.Comp
open MvPFunctor MvFunctor

variable {n m : ℕ}
(F : TypeFun n)
open MvPFunctor MvFunctor

variable {n m : ℕ}
(F : TypeFun n)
(G : Vec (TypeFun m) n)
[p : IsPolynomial F]
[p' : ∀ i, IsPolynomial <| G i]
Expand All @@ -46,25 +46,25 @@ variable {n m : ℕ}
snd := fun x a =>
(repr (abs { fst := f' x a, snd := fun j b => f j { fst := x, snd := { fst := a, snd := b } } })).fst }
i =
B (P (Comp F G)) { fst := a', snd := f' } i :=
B (P (Comp F G)) { fst := a', snd := f' } i :=
by simp only [IsPolynomial.repr_abs]; rfl

apply HEq.funext'
case type_eq_β =>
case type_eq_β =>
intros; rfl
case type_eq_α =>
case type_eq_α =>
exact type_eq_α

rintro ⟨b, ⟨b', g'⟩⟩
simp only [heq_eq_eq]
have : repr (abs
⟨ f' b b',
fun j b_1 =>

have : repr (abs
⟨ f' b b',
fun j b_1 =>
f j ⟨b, ⟨b', b_1⟩⟩
⟩)
= ⟨ f' b b',
fun j b_1 =>
= ⟨ f' b b',
fun j b_1 =>
f j ⟨b, ⟨b', b_1⟩⟩
⟩;
{
Expand All @@ -77,10 +77,9 @@ variable {n m : ℕ}
sorry
}






end MvQPF.Comp



end MvQPF.Comp
23 changes: 11 additions & 12 deletions Qpf/Util/TypeFun.lean
Original file line number Diff line number Diff line change
Expand Up @@ -14,15 +14,15 @@ abbrev CurriedFun (α : Type u) (β : Type v) : Nat → Type (max u v)
abbrev CurriedTypeFun : Nat → Type ((max u v) + 1)
:= CurriedFun (Type u) (Type v)

/-- An uncurried type function, all `n` arguments are collected into a single `TypeVec n`
/-- An uncurried type function, all `n` arguments are collected into a single `TypeVec n`
Note that all arguments live in the same universe, but the result may be in a different universe
-/
abbrev TypeFun (n : Nat) : Type ((max u v) + 1) :=
abbrev TypeFun (n : Nat) : Type ((max u v) + 1) :=
TypeVec.{u} n → Type v

namespace TypeFun
def reverseArgs : TypeFun n → TypeFun n :=
fun v α => v <| Vec.reverse α
fun v α => v <| Vec.reverse α

@[simp]
theorem reverseArgs_involution (F : TypeFun n) :
Expand All @@ -33,7 +33,7 @@ namespace TypeFun

def curriedAux : {n : Nat} → TypeFun n → CurriedTypeFun n
| 0, F => fun _ => F !![]
| 1, F => fun a => F !![a]
| 1, F => fun a => F !![a]
| _+2, F => fun a => curriedAux fun αs => F (αs ::: a)

def curried (F : TypeFun n) : CurriedTypeFun n
Expand All @@ -53,7 +53,7 @@ namespace TypeFun
@[simp]
theorem curriedAux_ofCurriedAux_involution {F : CurriedTypeFun n} :
curriedAux (ofCurriedAux F) = F :=
by
by
cases n
case zero => simp [curriedAux, ofCurriedAux]
case succ n => {
Expand All @@ -64,12 +64,12 @@ namespace TypeFun
funext a;
apply @ih (F a);
}
}
}

@[simp]
theorem curried_ofCurried_involution {F : CurriedTypeFun n} :
curried (ofCurried F) = F :=
by
by
simp only [curried, ofCurried, reverseArgs_involution]
apply curriedAux_ofCurriedAux_involution

Expand All @@ -78,9 +78,9 @@ namespace TypeFun
@[simp]
theorem ofCurriedAux_curriedAux_involution {F : TypeFun n} :
ofCurriedAux (curriedAux F) = F :=
by
by
cases n
case zero =>
case zero =>
funext x;
simp [curriedAux, ofCurriedAux, Matrix.vecEmpty]
congr
Expand All @@ -102,8 +102,7 @@ namespace TypeFun
funext x;
simp [ofCurriedAux, curriedAux];
let F' := fun α => F (α ::: x.last);
have : F x = F' x.drop;
. simp
have : F x = F' x.drop := by simp [F']
rw [this]
rw [@ih F']
}
Expand All @@ -112,7 +111,7 @@ namespace TypeFun
@[simp]
theorem ofCurried_curried_involution {F : TypeFun n} :
ofCurried (curried F) = F :=
by
by
simp only [ofCurried, curried, ofCurriedAux_curriedAux_involution]
apply reverseArgs_involution
end TypeFun
Loading

0 comments on commit 789630c

Please sign in to comment.