From a8c6d63d3e5ffbdc7c54a8029ca46d640bfdfd07 Mon Sep 17 00:00:00 2001 From: Josh Cohen Date: Tue, 16 Sep 2025 16:58:27 -0400 Subject: [PATCH 01/19] Allow LTy to take in TypeRestrictions parameter Instantiate with unit as default for compatibiity --- Strata/DL/Lambda/LExpr.lean | 176 ++++++++++-------- Strata/DL/Lambda/LTy.lean | 69 +++---- Strata/DL/Lambda/LTyUnify.lean | 8 +- .../Boogie/Examples/DDMAxiomsExtraction.lean | 33 ++-- Strata/Languages/Boogie/Identifiers.lean | 2 +- 5 files changed, 159 insertions(+), 129 deletions(-) diff --git a/Strata/DL/Lambda/LExpr.lean b/Strata/DL/Lambda/LExpr.lean index 1d45f06a..907e7340 100644 --- a/Strata/DL/Lambda/LExpr.lean +++ b/Strata/DL/Lambda/LExpr.lean @@ -322,51 +322,51 @@ scoped syntax lconstmono : lexprmono def elabLConstMono (Identifier : Type) [MkIdent Identifier] : Lean.Syntax → MetaM Expr | `(lconstmono| #$n:num) => do - let none ← mkNone (mkConst ``LMonoTy) - let typeTypeExpr := mkConst ``LMonoTy + let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit (toString n.getNat), none] | `(lconstmono| (#$n:num : $ty:lmonoty)) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy ty - let lmonoty ← mkSome (mkConst ``LMonoTy) lmonoty - let typeTypeExpr := mkConst ``LMonoTy + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Unit)) lmonoty + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit (toString n.getNat), lmonoty] | `(lconstmono| #-$n:num) => do - let none ← mkNone (mkConst ``LMonoTy) - let typeTypeExpr := mkConst ``LMonoTy + let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit ("-" ++ (toString n.getNat)), none] | `(lconstmono| (#-$n:num : $ty:lmonoty)) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy ty - let lmonoty ← mkSome (mkConst ``LMonoTy) lmonoty - let typeTypeExpr := mkConst ``LMonoTy + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Unit)) lmonoty + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit ("-" ++ (toString n.getNat)), lmonoty] | `(lconstmono| #true) => do - let none ← mkNone (mkConst ``LMonoTy) - let typeTypeExpr := mkConst ``LMonoTy + let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit "true", none] | `(lconstmono| (#true : $ty:lmonoty)) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy ty - let lmonoty ← mkSome (mkConst ``LMonoTy) lmonoty - let typeTypeExpr := mkConst ``LMonoTy + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Unit)) lmonoty + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit "true", lmonoty] | `(lconstmono| #false) => do - let none ← mkNone (mkConst ``LMonoTy) - let typeTypeExpr := mkConst ``LMonoTy + let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit "false", none] | `(lconstmono| (#false : $ty:lmonoty)) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy ty - let lmonoty ← mkSome (mkConst ``LMonoTy) lmonoty - let typeTypeExpr := mkConst ``LMonoTy + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Unit)) lmonoty + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit "false", lmonoty] | `(lconstmono| #$s:ident) => do - let none ← mkNone (mkConst ``LMonoTy) + let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Unit)) let s := toString s.getId - let typeTypeExpr := mkConst ``LMonoTy + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit s, none] | `(lconstmono| (#$s:ident : $ty:lmonoty)) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy ty - let lmonoty ← mkSome (mkConst ``LMonoTy) lmonoty + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Unit)) lmonoty let s := toString s.getId - let typeTypeExpr := mkConst ``LMonoTy + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit s, lmonoty] | _ => throwUnsupportedSyntax @@ -377,13 +377,13 @@ scoped syntax lopmono : lexprmono def elabLOpMono (Identifier : Type) [MkIdent Identifier] : Lean.Syntax → MetaM Expr | `(lopmono| ~$s:lidentmono) => do - let none ← mkNone (mkConst ``LMonoTy) + let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Unit)) let ident ← MkIdent.elabIdent Identifier s - let typeTypeExpr := mkConst ``LMonoTy + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) return mkAppN (.const ``LExpr.op []) #[typeTypeExpr, MkIdent.toExpr Identifier, ident, none] | `(lopmono| (~$s:lidentmono : $ty:lmonoty)) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy ty - let lmonoty ← mkSome (mkConst ``LMonoTy) lmonoty + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Unit)) lmonoty mkAppM ``LExpr.op #[← MkIdent.elabIdent Identifier s, lmonoty] | _ => throwUnsupportedSyntax @@ -391,7 +391,7 @@ declare_syntax_cat lbvarmono scoped syntax "%" noWs num : lbvarmono def elabLBVarMono (Identifier : Type) [MkIdent Identifier] : Lean.Syntax → MetaM Expr | `(lbvarmono| %$n:num) => - let typeTypeExpr := mkConst ``LMonoTy + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) return mkAppN (.const ``LExpr.bvar []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkNatLit n.getNat] | _ => throwUnsupportedSyntax scoped syntax lbvarmono : lexprmono @@ -402,11 +402,11 @@ scoped syntax "(" lidentmono ":" lmonoty ")" : lfvarmono def elabLFVarMono (Identifier : Type) [MkIdent Identifier] : Lean.Syntax → MetaM Expr | `(lfvarmono| $i:lidentmono) => do - let none ← mkNone (mkConst ``LMonoTy) + let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Unit)) mkAppM ``LExpr.fvar #[← MkIdent.elabIdent Identifier i, none] | `(lfvarmono| ($i:lidentmono : $ty:lmonoty)) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy ty - let lmonoty ← mkSome (mkConst ``LMonoTy) lmonoty + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Unit)) lmonoty mkAppM ``LExpr.fvar #[← MkIdent.elabIdent Identifier i, lmonoty] | _ => throwUnsupportedSyntax scoped syntax lfvarmono : lexprmono @@ -455,48 +455,48 @@ partial def elabLExprMono (Identifier : Type) [MkIdent Identifier] : Lean.Syntax | `(lexprmono| $f:lfvarmono) => elabLFVarMono Identifier f | `(lexprmono| λ $e:lexprmono) => do let e' ← elabLExprMono Identifier e - let typeTypeExpr := mkConst ``LMonoTy + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) return mkAppN (.const ``LExpr.absUntyped []) #[typeTypeExpr, MkIdent.toExpr Identifier, e'] | `(lexprmono| λ ($mty:lmonoty): $e:lexprmono) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy mty - let lmonoty ← mkSome (mkConst ``LMonoTy) lmonoty + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Unit)) lmonoty let e' ← elabLExprMono Identifier e - let typeTypeExpr := mkConst ``LMonoTy + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) return mkAppN (.const ``LExpr.abs []) #[typeTypeExpr, MkIdent.toExpr Identifier, lmonoty, e'] | `(lexprmono| ∀ $e:lexprmono) => do let e' ← elabLExprMono Identifier e - let typeTypeExpr := mkConst ``LMonoTy + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) return mkAppN (.const ``LExpr.allUntyped []) #[typeTypeExpr, MkIdent.toExpr Identifier, e'] | `(lexprmono| ∀ {$tr}$e:lexprmono) => do let e' ← elabLExprMono Identifier e let tr' ← elabLExprMono Identifier tr - let typeTypeExpr := mkConst ``LMonoTy + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) return mkAppN (.const ``LExpr.allUntypedTr []) #[typeTypeExpr, MkIdent.toExpr Identifier, tr', e'] | `(lexprmono| ∀ ($mty:lmonoty): $e:lexprmono) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy mty - let lmonoty ← mkSome (mkConst ``LMonoTy) lmonoty + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Unit)) lmonoty let e' ← elabLExprMono Identifier e - let typeTypeExpr := mkConst ``LMonoTy + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) return mkAppN (.const ``LExpr.all []) #[typeTypeExpr, MkIdent.toExpr Identifier, lmonoty, e'] | `(lexprmono| ∀ ($mty:lmonoty):{$tr} $e:lexprmono) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy mty - let lmonoty ← mkSome (mkConst ``LMonoTy) lmonoty + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Unit)) lmonoty let e' ← elabLExprMono Identifier e let tr' ← elabLExprMono Identifier tr - let typeTypeExpr := mkConst ``LMonoTy + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) return mkAppN (.const ``LExpr.allTr []) #[typeTypeExpr, MkIdent.toExpr Identifier, lmonoty, tr', e'] | `(lexprmono| ∃ ($mty:lmonoty): $e:lexprmono) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy mty - let lmonoty ← mkSome (mkConst ``LMonoTy) lmonoty + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Unit)) lmonoty let e' ← elabLExprMono Identifier e - let typeTypeExpr := mkConst ``LMonoTy + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) return mkAppN (.const ``LExpr.exist []) #[typeTypeExpr, MkIdent.toExpr Identifier, lmonoty, e'] | `(lexprmono| ∃ ($mty:lmonoty):{$tr} $e:lexprmono) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy mty - let lmonoty ← mkSome (mkConst ``LMonoTy) lmonoty + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Unit)) lmonoty let e' ← elabLExprMono Identifier e let tr' ← elabLExprMono Identifier tr - let typeTypeExpr := mkConst ``LMonoTy + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) return mkAppN (.const ``LExpr.existTr []) #[typeTypeExpr, MkIdent.toExpr Identifier, lmonoty, tr', e'] | `(lexprmono| ∃ $e:lexprmono) => do let e' ← elabLExprMono Identifier e @@ -508,18 +508,18 @@ partial def elabLExprMono (Identifier : Type) [MkIdent Identifier] : Lean.Syntax | `(lexprmono| ($e1:lexprmono $e2:lexprmono)) => do let e1' ← elabLExprMono Identifier e1 let e2' ← elabLExprMono Identifier e2 - let typeTypeExpr := mkConst ``LMonoTy + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) return mkAppN (.const ``LExpr.app []) #[typeTypeExpr, MkIdent.toExpr Identifier, e1', e2'] | `(lexprmono| $e1:lexprmono == $e2:lexprmono) => do let e1' ← elabLExprMono Identifier e1 let e2' ← elabLExprMono Identifier e2 - let typeTypeExpr := mkConst ``LMonoTy + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) return mkAppN (.const ``LExpr.eq []) #[typeTypeExpr, MkIdent.toExpr Identifier, e1', e2'] | `(lexprmono| if $e1:lexprmono then $e2:lexprmono else $e3:lexprmono) => do let e1' ← elabLExprMono Identifier e1 let e2' ← elabLExprMono Identifier e2 let e3' ← elabLExprMono Identifier e3 - let typeTypeExpr := mkConst ``LMonoTy + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) return mkAppN (.const ``LExpr.ite []) #[typeTypeExpr, MkIdent.toExpr Identifier, e1', e2', e3'] | `(lexprmono| ($e:lexprmono)) => elabLExprMono Identifier e | _ => throwUnsupportedSyntax @@ -544,7 +544,7 @@ open LTy.Syntax #guard_msgs in #check esM[((λ %0) #5)] -/-- info: (abs (some (LMonoTy.tcons "bool" [])) (bvar 0)).app (const "true" none) : LExpr LMonoTy String -/ +/-- info: (abs (some (LMonoTy.tcons Unit "bool" [] ())) (bvar 0)).app (const "true" none) : LExpr LMonoTy String -/ #guard_msgs in #check esM[((λ (bool): %0) #true)] @@ -556,36 +556,44 @@ open LTy.Syntax #guard_msgs in #check esM[(∃ %0 == #5)] -/-- info: exist (some (LMonoTy.tcons "int" [])) ((bvar 0).eq (const "5" none)) : LExpr LMonoTy String -/ +/-- info: exist (some (LMonoTy.tcons Unit "int" [] ())) ((bvar 0).eq (const "5" none)) : LExpr LMonoTy String -/ #guard_msgs in #check esM[(∃ (int): %0 == #5)] -/-- info: fvar "x" (some (LMonoTy.tcons "bool" [])) : LExpr LMonoTy String -/ +/-- info: fvar "x" (some (LMonoTy.tcons Unit "bool" [] ())) : LExpr LMonoTy String -/ #guard_msgs in #check esM[(x : bool)] -- axiom [updateSelect]: forall m: Map k v, kk: k, vv: v :: m[kk := vv][kk] == vv; /-- -info: all (some (LMonoTy.tcons "Map" [LMonoTy.ftvar "k", LMonoTy.ftvar "v"])) - (all (some (LMonoTy.ftvar "k")) - (all (some (LMonoTy.ftvar "v")) +info: all (some (LMonoTy.tcons Unit "Map" [LMonoTy.ftvar Unit "k", LMonoTy.ftvar Unit "v"] ())) + (all (some (LMonoTy.ftvar Unit "k")) + (all (some (LMonoTy.ftvar Unit "v")) ((((op "select" (some - (LMonoTy.tcons "Map" - [LMonoTy.ftvar "k", - LMonoTy.tcons "arrow" - [LMonoTy.ftvar "v", LMonoTy.tcons "arrow" [LMonoTy.ftvar "k", LMonoTy.ftvar "v"]]]))).app + (LMonoTy.tcons Unit "Map" + [LMonoTy.ftvar Unit "k", + LMonoTy.tcons Unit "arrow" + [LMonoTy.ftvar Unit "v", + LMonoTy.tcons Unit "arrow" [LMonoTy.ftvar Unit "k", LMonoTy.ftvar Unit "v"] ()] + ()] + ()))).app ((((op "update" (some - (LMonoTy.tcons "Map" - [LMonoTy.ftvar "k", - LMonoTy.tcons "arrow" - [LMonoTy.ftvar "v", - LMonoTy.tcons "arrow" - [LMonoTy.ftvar "k", - LMonoTy.tcons "arrow" - [LMonoTy.ftvar "v", - LMonoTy.tcons "Map" [LMonoTy.ftvar "k", LMonoTy.ftvar "v"]]]]]))).app + (LMonoTy.tcons Unit "Map" + [LMonoTy.ftvar Unit "k", + LMonoTy.tcons Unit "arrow" + [LMonoTy.ftvar Unit "v", + LMonoTy.tcons Unit "arrow" + [LMonoTy.ftvar Unit "k", + LMonoTy.tcons Unit "arrow" + [LMonoTy.ftvar Unit "v", + LMonoTy.tcons Unit "Map" + [LMonoTy.ftvar Unit "k", LMonoTy.ftvar Unit "v"] ()] + ()] + ()] + ()] + ()))).app (bvar 2)).app (bvar 1)).app (bvar 0))).app @@ -858,7 +866,7 @@ open LTy.Syntax #guard_msgs in #check es[((λ %0) #5)] -/-- info: (abs (some (LTy.forAll [] (LMonoTy.tcons "bool" []))) (bvar 0)).app (const "true" none) : LExpr LTy String -/ +/-- info: (abs (some (LTy.forAll [] (LMonoTy.tcons Unit "bool" [] ()))) (bvar 0)).app (const "true" none) : LExpr LTy String -/ #guard_msgs in #check es[((λ (bool): %0) #true)] @@ -870,38 +878,46 @@ open LTy.Syntax #guard_msgs in #check es[(∃ %0 == #5)] -/-- info: exist (some (LTy.forAll [] (LMonoTy.tcons "int" []))) ((bvar 0).eq (const "5" none)) : LExpr LTy String -/ +/-- info: exist (some (LTy.forAll [] (LMonoTy.tcons Unit "int" [] ()))) ((bvar 0).eq (const "5" none)) : LExpr LTy String -/ #guard_msgs in #check es[(∃ (int): %0 == #5)] -/-- info: fvar "x" (some (LTy.forAll [] (LMonoTy.tcons "bool" []))) : LExpr LTy String -/ +/-- info: fvar "x" (some (LTy.forAll [] (LMonoTy.tcons Unit "bool" [] ()))) : LExpr LTy String -/ #guard_msgs in #check es[(x : bool)] -- axiom [updateSelect]: forall m: Map k v, kk: k, vv: v :: m[kk := vv][kk] == vv; /-- -info: all (some (LTy.forAll [] (LMonoTy.tcons "Map" [LMonoTy.ftvar "k", LMonoTy.ftvar "v"]))) - (all (some (LTy.forAll [] (LMonoTy.ftvar "k"))) - (all (some (LTy.forAll [] (LMonoTy.ftvar "v"))) +info: all (some (LTy.forAll [] (LMonoTy.tcons Unit "Map" [LMonoTy.ftvar Unit "k", LMonoTy.ftvar Unit "v"] ()))) + (all (some (LTy.forAll [] (LMonoTy.ftvar Unit "k"))) + (all (some (LTy.forAll [] (LMonoTy.ftvar Unit "v"))) ((((op "select" (some (LTy.forAll [] - (LMonoTy.tcons "Map" - [LMonoTy.ftvar "k", - LMonoTy.tcons "arrow" - [LMonoTy.ftvar "v", LMonoTy.tcons "arrow" [LMonoTy.ftvar "k", LMonoTy.ftvar "v"]]])))).app + (LMonoTy.tcons Unit "Map" + [LMonoTy.ftvar Unit "k", + LMonoTy.tcons Unit "arrow" + [LMonoTy.ftvar Unit "v", + LMonoTy.tcons Unit "arrow" [LMonoTy.ftvar Unit "k", LMonoTy.ftvar Unit "v"] ()] + ()] + ())))).app ((((op "update" (some (LTy.forAll [] - (LMonoTy.tcons "Map" - [LMonoTy.ftvar "k", - LMonoTy.tcons "arrow" - [LMonoTy.ftvar "v", - LMonoTy.tcons "arrow" - [LMonoTy.ftvar "k", - LMonoTy.tcons "arrow" - [LMonoTy.ftvar "v", - LMonoTy.tcons "Map" [LMonoTy.ftvar "k", LMonoTy.ftvar "v"]]]]])))).app + (LMonoTy.tcons Unit "Map" + [LMonoTy.ftvar Unit "k", + LMonoTy.tcons Unit "arrow" + [LMonoTy.ftvar Unit "v", + LMonoTy.tcons Unit "arrow" + [LMonoTy.ftvar Unit "k", + LMonoTy.tcons Unit "arrow" + [LMonoTy.ftvar Unit "v", + LMonoTy.tcons Unit "Map" + [LMonoTy.ftvar Unit "k", LMonoTy.ftvar Unit "v"] ()] + ()] + ()] + ()] + ())))).app (bvar 2)).app (bvar 1)).app (bvar 0))).app diff --git a/Strata/DL/Lambda/LTy.lean b/Strata/DL/Lambda/LTy.lean index f033f8c2..6fe8d3f2 100644 --- a/Strata/DL/Lambda/LTy.lean +++ b/Strata/DL/Lambda/LTy.lean @@ -29,11 +29,11 @@ instance : Coe String TyIdentifier where Types in Lambda: these are mono-types. Note that all free type variables (`.ftvar`) are implicitly universally quantified. -/ -inductive LMonoTy : Type where +inductive LMonoTy (TypeRestrictions: Type := Unit) : Type where /-- Type variable. -/ | ftvar (name : TyIdentifier) /-- Type constructor. -/ - | tcons (name : String) (args : List LMonoTy) + | tcons (name : String) (args : List (LMonoTy TypeRestrictions)) (r: TypeRestrictions := by exact ()) /-- Special support for bitvector types of every size. -/ | bitvec (size : Nat) deriving Inhabited, Repr @@ -124,10 +124,10 @@ Lean's default `induction` tactic does not support nested or mutual inductive types. So we define our own induction principle below. -/ @[induction_eliminator] -theorem LMonoTy.induct {P : LMonoTy → Prop} +theorem LMonoTy.induct {Meta: Type} {P : LMonoTy Meta → Prop} (ftvar : ∀f, P (.ftvar f)) (bitvec : ∀n, P (.bitvec n)) - (tcons : ∀name args, (∀ ty ∈ args, P ty) → P (.tcons name args)) : + (tcons : ∀name args r, (∀ ty ∈ args, P ty) → P (.tcons name args r)) : ∀ ty, P ty := by intro n apply LMonoTy.rec <;> try assumption @@ -179,7 +179,7 @@ def LMonoTy.BEq (x y : LMonoTy) : Bool := @[simp] theorem LMonoTy.BEq_refl : LMonoTy.BEq ty ty := by induction ty <;> simp_all [LMonoTy.BEq] - rename_i name args ih + rename_i name args r ih induction args case tcons.nil => simp [LMonoTy.BEq.go] case tcons.cons => @@ -197,9 +197,9 @@ instance : DecidableEq LMonoTy := case bitvec => unfold LMonoTy.BEq at h <;> split at h <;> try simp_all case tcons => - rename_i name args ih + rename_i name args r ih cases y <;> try simp_all [LMonoTy.BEq] - rename_i name' args' + rename_i name' args' r' obtain ⟨⟨h1, h2⟩, h3⟩ := h induction args generalizing args' case nil => unfold List.length at h2; split at h2 <;> simp_all @@ -215,9 +215,9 @@ instance : DecidableEq LMonoTy := cases y <;> try simp_all [LMonoTy.BEq] case bitvec n => cases y <;> try simp_all [LMonoTy.BEq] - case tcons name args ih => + case tcons name args r ih => cases y <;> try simp_all [LMonoTy.BEq] - rename_i name' args' + rename_i name' args' r' intro hname; simp [hname] at h induction args generalizing args' case tcons.nil => @@ -279,10 +279,10 @@ def LMonoTy.getTyConstructors (mty : LMonoTy) : List LMonoTy := | [] => [] | m :: mrest => LMonoTy.getTyConstructors m ++ go mrest /-- -info: [Lambda.LMonoTy.tcons "arrow" [Lambda.LMonoTy.ftvar "_dummy0", Lambda.LMonoTy.ftvar "_dummy1"], - Lambda.LMonoTy.tcons "bool" [], - Lambda.LMonoTy.tcons "foo" [Lambda.LMonoTy.ftvar "_dummy0"], - Lambda.LMonoTy.tcons "a" [Lambda.LMonoTy.ftvar "_dummy0", Lambda.LMonoTy.ftvar "_dummy1"]] +info: [Lambda.LMonoTy.tcons "arrow" [Lambda.LMonoTy.ftvar "_dummy0", Lambda.LMonoTy.ftvar "_dummy1"] (), + Lambda.LMonoTy.tcons "bool" [] (), + Lambda.LMonoTy.tcons "foo" [Lambda.LMonoTy.ftvar "_dummy0"] (), + Lambda.LMonoTy.tcons "a" [Lambda.LMonoTy.ftvar "_dummy0", Lambda.LMonoTy.ftvar "_dummy1"] ()] -/ #guard_msgs in #eval LMonoTy.getTyConstructors @@ -408,18 +408,18 @@ scoped syntax "(" lmonoty ")" : lmonoty open Lean Elab Meta in partial def elabLMonoTy : Lean.Syntax → MetaM Expr | `(lmonoty| %$f:ident) => do - mkAppM ``LMonoTy.ftvar #[mkStrLit (toString f.getId)] + mkAppOptM ``LMonoTy.ftvar #[mkConst `Unit, mkStrLit (toString f.getId)] | `(lmonoty| $ty1:lmonoty → $ty2:lmonoty) => do let ty1' ← elabLMonoTy ty1 let ty2' ← elabLMonoTy ty2 - let tys ← mkListLit (mkConst ``LMonoTy) [ty1', ty2'] - mkAppM ``LMonoTy.tcons #[(mkStrLit "arrow"), tys] + let tys ← mkListLit (.app (mkConst ``LMonoTy) (mkConst `Unit)) [ty1', ty2'] + mkAppM ``LMonoTy.tcons #[(mkStrLit "arrow"), tys, (mkConst `Unit.unit)] | `(lmonoty| int) => do - let argslist ← mkListLit (mkConst ``LMonoTy) [] - mkAppM ``LMonoTy.tcons #[(mkStrLit "int"), argslist] + let argslist ← mkListLit (.app (mkConst ``LMonoTy) (mkConst `Unit)) [] + mkAppM ``LMonoTy.tcons #[(mkStrLit "int"), argslist, (mkConst `Unit.unit)] | `(lmonoty| bool) => do - let argslist ← mkListLit (mkConst ``LMonoTy) [] - mkAppM ``LMonoTy.tcons #[(mkStrLit "bool"), argslist] + let argslist ← mkListLit (.app (mkConst ``LMonoTy) (mkConst `Unit)) [] + mkAppM ``LMonoTy.tcons #[(mkStrLit "bool"), argslist, (mkConst `Unit.unit)] | `(lmonoty| bv1) => mkAppM ``LMonoTy.bv1 #[] | `(lmonoty| bv8) => mkAppM ``LMonoTy.bv8 #[] | `(lmonoty| bv16) => mkAppM ``LMonoTy.bv16 #[] @@ -427,8 +427,8 @@ partial def elabLMonoTy : Lean.Syntax → MetaM Expr | `(lmonoty| bv64) => mkAppM ``LMonoTy.bv64 #[] | `(lmonoty| $i:ident $args:lmonoty*) => do let args' ← go args - let argslist ← mkListLit (mkConst ``LMonoTy) args'.toList - mkAppM ``LMonoTy.tcons #[(mkStrLit (toString i.getId)), argslist] + let argslist ← mkListLit (.app (mkConst ``LMonoTy) (mkConst `Unit)) args'.toList + mkAppM ``LMonoTy.tcons #[(mkStrLit (toString i.getId)), argslist, (mkConst `Unit.unit)] | `(lmonoty| ($ty:lmonoty)) => elabLMonoTy ty | _ => throwUnsupportedSyntax where go (args : TSyntaxArray `lmonoty) : MetaM (Array Expr) := do @@ -440,26 +440,29 @@ partial def elabLMonoTy : Lean.Syntax → MetaM Expr elab "mty[" ty:lmonoty "]" : term => elabLMonoTy ty -/-- info: LMonoTy.tcons "list" [LMonoTy.tcons "int" []] : LMonoTy -/ +/-- info: LMonoTy.tcons Unit "list" [LMonoTy.tcons Unit "int" [] ()] () : LMonoTy -/ #guard_msgs in #check mty[list int] -/-- info: LMonoTy.tcons "pair" [LMonoTy.tcons "int" [], LMonoTy.tcons "bool" []] : LMonoTy -/ +/-- info: LMonoTy.tcons Unit "pair" [LMonoTy.tcons Unit "int" [] (), LMonoTy.tcons Unit "bool" [] ()] () : LMonoTy -/ #guard_msgs in #check mty[pair int bool] /-- -info: LMonoTy.tcons "arrow" - [LMonoTy.tcons "Map" [LMonoTy.ftvar "k", LMonoTy.ftvar "v"], - LMonoTy.tcons "arrow" [LMonoTy.ftvar "k", LMonoTy.ftvar "v"]] : LMonoTy +info: LMonoTy.tcons Unit "arrow" + [LMonoTy.tcons Unit "Map" [LMonoTy.ftvar Unit "k", LMonoTy.ftvar Unit "v"] (), + LMonoTy.tcons Unit "arrow" [LMonoTy.ftvar Unit "k", LMonoTy.ftvar Unit "v"] ()] + () : LMonoTy -/ #guard_msgs in #check mty[(Map %k %v) → %k → %v] /-- -info: LMonoTy.tcons "arrow" - [LMonoTy.ftvar "a", - LMonoTy.tcons "arrow" [LMonoTy.ftvar "b", LMonoTy.tcons "arrow" [LMonoTy.ftvar "c", LMonoTy.ftvar "d"]]] : LMonoTy +info: LMonoTy.tcons Unit "arrow" + [LMonoTy.ftvar Unit "a", + LMonoTy.tcons Unit "arrow" + [LMonoTy.ftvar Unit "b", LMonoTy.tcons Unit "arrow" [LMonoTy.ftvar Unit "c", LMonoTy.ftvar Unit "d"] ()] ()] + () : LMonoTy -/ #guard_msgs in #check mty[%a → %b → %c → %d] @@ -485,13 +488,15 @@ partial def elabLTy : Lean.Syntax → MetaM Expr elab "t[" lsch:lty "]" : term => elabLTy lsch -/-- info: forAll ["α"] (LMonoTy.tcons "myType" [LMonoTy.ftvar "α"]) : LTy -/ +/-- info: forAll ["α"] (LMonoTy.tcons Unit "myType" [LMonoTy.ftvar Unit "α"] ()) : LTy -/ #guard_msgs in #check t[∀α. myType %α] /-- info: forAll ["α"] - (LMonoTy.tcons "arrow" [LMonoTy.ftvar "α", LMonoTy.tcons "arrow" [LMonoTy.ftvar "α", LMonoTy.tcons "int" []]]) : LTy + (LMonoTy.tcons Unit "arrow" + [LMonoTy.ftvar Unit "α", LMonoTy.tcons Unit "arrow" [LMonoTy.ftvar Unit "α", LMonoTy.tcons Unit "int" [] ()] ()] + ()) : LTy -/ #guard_msgs in #check t[∀α. %α → %α → int] diff --git a/Strata/DL/Lambda/LTyUnify.lean b/Strata/DL/Lambda/LTyUnify.lean index 8db8be3c..11a6d5dd 100644 --- a/Strata/DL/Lambda/LTyUnify.lean +++ b/Strata/DL/Lambda/LTyUnify.lean @@ -122,8 +122,8 @@ def LMonoTy.subst (S : Subst) (mty : LMonoTy) : LMonoTy := | .ftvar x => match S.find? x with | some sty => sty | none => mty | .bitvec _ => mty - | .tcons name ltys => - .tcons name (LMonoTys.subst S ltys) + | .tcons name ltys r => + .tcons name (LMonoTys.subst S ltys) r /-- Apply substitution `S` to monotypes `mtys`. -/ @@ -200,7 +200,7 @@ theorem LMonoTy.subst_keys_not_in_substituted_type (h : SubstWF S) : simp_all [LMonoTy.subst] unfold LMonoTy.freeVars simp - case tcons name args h1 => + case tcons name args r h1 => simp_all simp [subst] induction args @@ -243,7 +243,7 @@ theorem LMonoTy.freeVars_of_subst_subset (S : Subst) (mty : LMonoTy) : simp [freeVars] case bitvec n => simp [subst] - case tcons name args ih => + case tcons name args r ih => simp [LMonoTy.subst, LMonoTy.freeVars] induction args case nil => diff --git a/Strata/Languages/Boogie/Examples/DDMAxiomsExtraction.lean b/Strata/Languages/Boogie/Examples/DDMAxiomsExtraction.lean index bd17a8ca..91e3399b 100644 --- a/Strata/Languages/Boogie/Examples/DDMAxiomsExtraction.lean +++ b/Strata/Languages/Boogie/Examples/DDMAxiomsExtraction.lean @@ -239,7 +239,7 @@ info: #[{ name := { dialect := "Boogie", name := "command_typedecl" }, /-- info: [Lambda.LExpr.quant (Lambda.QuantifierKind.all) - (some (Lambda.LMonoTy.tcons "Map" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"])) + (some (Lambda.LMonoTy.tcons "Map" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] ())) (Lambda.LExpr.bvar 0) (Lambda.LExpr.quant (Lambda.QuantifierKind.all) @@ -256,8 +256,9 @@ info: [Lambda.LExpr.quant u:select (some (Lambda.LMonoTy.tcons "arrow" - [Lambda.LMonoTy.tcons "Map" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"], - Lambda.LMonoTy.tcons "arrow" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"]]))) + [Lambda.LMonoTy.tcons "Map" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] (), + Lambda.LMonoTy.tcons "arrow" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] ()] + ()))) (Lambda.LExpr.app (Lambda.LExpr.app (Lambda.LExpr.app @@ -265,14 +266,17 @@ info: [Lambda.LExpr.quant u:update (some (Lambda.LMonoTy.tcons "arrow" - [Lambda.LMonoTy.tcons "Map" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"], + [Lambda.LMonoTy.tcons "Map" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] (), Lambda.LMonoTy.tcons "arrow" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.tcons "arrow" [Lambda.LMonoTy.ftvar "v", - Lambda.LMonoTy.tcons "Map" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"]]]]))) + Lambda.LMonoTy.tcons "Map" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] ()] + ()] + ()] + ()))) (Lambda.LExpr.bvar 2)) (Lambda.LExpr.bvar 1)) (Lambda.LExpr.bvar 0))) @@ -280,7 +284,7 @@ info: [Lambda.LExpr.quant (Lambda.LExpr.bvar 0)))), Lambda.LExpr.quant (Lambda.QuantifierKind.all) - (some (Lambda.LMonoTy.tcons "Map" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"])) + (some (Lambda.LMonoTy.tcons "Map" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] ())) (Lambda.LExpr.bvar 0) (Lambda.LExpr.quant (Lambda.QuantifierKind.all) @@ -301,8 +305,9 @@ info: [Lambda.LExpr.quant u:select (some (Lambda.LMonoTy.tcons "arrow" - [Lambda.LMonoTy.tcons "Map" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"], - Lambda.LMonoTy.tcons "arrow" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"]]))) + [Lambda.LMonoTy.tcons "Map" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] (), + Lambda.LMonoTy.tcons "arrow" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] ()] + ()))) (Lambda.LExpr.app (Lambda.LExpr.app (Lambda.LExpr.app @@ -310,14 +315,17 @@ info: [Lambda.LExpr.quant u:update (some (Lambda.LMonoTy.tcons "arrow" - [Lambda.LMonoTy.tcons "Map" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"], + [Lambda.LMonoTy.tcons "Map" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] (), Lambda.LMonoTy.tcons "arrow" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.tcons "arrow" [Lambda.LMonoTy.ftvar "v", - Lambda.LMonoTy.tcons "Map" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"]]]]))) + Lambda.LMonoTy.tcons "Map" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] ()] + ()] + ()] + ()))) (Lambda.LExpr.bvar 3)) (Lambda.LExpr.bvar 1)) (Lambda.LExpr.bvar 0))) @@ -328,8 +336,9 @@ info: [Lambda.LExpr.quant u:select (some (Lambda.LMonoTy.tcons "arrow" - [Lambda.LMonoTy.tcons "Map" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"], - Lambda.LMonoTy.tcons "arrow" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"]]))) + [Lambda.LMonoTy.tcons "Map" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] (), + Lambda.LMonoTy.tcons "arrow" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] ()] + ()))) (Lambda.LExpr.bvar 3)) (Lambda.LExpr.bvar 2))))))] -/ diff --git a/Strata/Languages/Boogie/Identifiers.lean b/Strata/Languages/Boogie/Identifiers.lean index 298edd68..535a30ee 100644 --- a/Strata/Languages/Boogie/Identifiers.lean +++ b/Strata/Languages/Boogie/Identifiers.lean @@ -141,7 +141,7 @@ info: (Lambda.LExpr.op (BoogieIdent.unres "old") none).app open Lambda.LTy.Syntax in /-- info: Lambda.LExpr.fvar (BoogieIdent.unres "x") - (some (Lambda.LMonoTy.tcons "bool" [])) : Lambda.LExpr Lambda.LMonoTy (Visibility × String) -/ + (some (Lambda.LMonoTy.tcons Unit "bool" [] ())) : Lambda.LExpr Lambda.LMonoTy (Visibility × String) -/ #guard_msgs in #check eb[(x : bool)] From 6da14b0c1e3170a6196f52a16e1938203d11520c Mon Sep 17 00:00:00 2001 From: Josh Cohen Date: Wed, 17 Sep 2025 13:42:10 -0400 Subject: [PATCH 02/19] Refactor bitvectors as type metadata Bitvectors are no longer a special case in LTy but rather an instance of a type restriction/metadata --- Strata/DL/Lambda/LExpr.lean | 208 ++++++++++-------- Strata/DL/Lambda/LExprT.lean | 4 +- Strata/DL/Lambda/LExprTypeEnv.lean | 16 +- Strata/DL/Lambda/LTy.lean | 121 +++++----- Strata/DL/Lambda/LTyUnify.lean | 21 +- .../Boogie/DDMTransform/Translate.lean | 12 +- .../Boogie/Examples/DDMAxiomsExtraction.lean | 72 ++++-- Strata/Languages/Boogie/Identifiers.lean | 4 +- Strata/Languages/Boogie/SMTEncoder.lean | 14 +- 9 files changed, 262 insertions(+), 210 deletions(-) diff --git a/Strata/DL/Lambda/LExpr.lean b/Strata/DL/Lambda/LExpr.lean index 907e7340..a4e22671 100644 --- a/Strata/DL/Lambda/LExpr.lean +++ b/Strata/DL/Lambda/LExpr.lean @@ -168,7 +168,7 @@ Note that we are type-agnostic here. -/ def denoteBitVec (n : Nat) (e : (LExpr LMonoTy Identifier)) : Option (BitVec n) := match e with - | .const x (.some (.bitvec n')) => + | .const x (.some (bitvec n')) => if n == n' then .map (.ofNat n) x.toNat? else none | .const x none => .map (.ofNat n) x.toNat? | _ => none @@ -322,51 +322,51 @@ scoped syntax lconstmono : lexprmono def elabLConstMono (Identifier : Type) [MkIdent Identifier] : Lean.Syntax → MetaM Expr | `(lconstmono| #$n:num) => do - let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Unit)) - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit (toString n.getNat), none] | `(lconstmono| (#$n:num : $ty:lmonoty)) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy ty - let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Unit)) lmonoty - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) lmonoty + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit (toString n.getNat), lmonoty] | `(lconstmono| #-$n:num) => do - let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Unit)) - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit ("-" ++ (toString n.getNat)), none] | `(lconstmono| (#-$n:num : $ty:lmonoty)) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy ty - let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Unit)) lmonoty - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) lmonoty + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit ("-" ++ (toString n.getNat)), lmonoty] | `(lconstmono| #true) => do - let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Unit)) - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit "true", none] | `(lconstmono| (#true : $ty:lmonoty)) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy ty - let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Unit)) lmonoty - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) lmonoty + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit "true", lmonoty] | `(lconstmono| #false) => do - let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Unit)) - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit "false", none] | `(lconstmono| (#false : $ty:lmonoty)) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy ty - let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Unit)) lmonoty - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) lmonoty + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit "false", lmonoty] | `(lconstmono| #$s:ident) => do - let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) let s := toString s.getId - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit s, none] | `(lconstmono| (#$s:ident : $ty:lmonoty)) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy ty - let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Unit)) lmonoty + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) lmonoty let s := toString s.getId - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit s, lmonoty] | _ => throwUnsupportedSyntax @@ -377,13 +377,13 @@ scoped syntax lopmono : lexprmono def elabLOpMono (Identifier : Type) [MkIdent Identifier] : Lean.Syntax → MetaM Expr | `(lopmono| ~$s:lidentmono) => do - let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) let ident ← MkIdent.elabIdent Identifier s - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) return mkAppN (.const ``LExpr.op []) #[typeTypeExpr, MkIdent.toExpr Identifier, ident, none] | `(lopmono| (~$s:lidentmono : $ty:lmonoty)) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy ty - let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Unit)) lmonoty + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) lmonoty mkAppM ``LExpr.op #[← MkIdent.elabIdent Identifier s, lmonoty] | _ => throwUnsupportedSyntax @@ -391,7 +391,7 @@ declare_syntax_cat lbvarmono scoped syntax "%" noWs num : lbvarmono def elabLBVarMono (Identifier : Type) [MkIdent Identifier] : Lean.Syntax → MetaM Expr | `(lbvarmono| %$n:num) => - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) return mkAppN (.const ``LExpr.bvar []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkNatLit n.getNat] | _ => throwUnsupportedSyntax scoped syntax lbvarmono : lexprmono @@ -402,11 +402,11 @@ scoped syntax "(" lidentmono ":" lmonoty ")" : lfvarmono def elabLFVarMono (Identifier : Type) [MkIdent Identifier] : Lean.Syntax → MetaM Expr | `(lfvarmono| $i:lidentmono) => do - let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) mkAppM ``LExpr.fvar #[← MkIdent.elabIdent Identifier i, none] | `(lfvarmono| ($i:lidentmono : $ty:lmonoty)) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy ty - let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Unit)) lmonoty + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) lmonoty mkAppM ``LExpr.fvar #[← MkIdent.elabIdent Identifier i, lmonoty] | _ => throwUnsupportedSyntax scoped syntax lfvarmono : lexprmono @@ -455,48 +455,48 @@ partial def elabLExprMono (Identifier : Type) [MkIdent Identifier] : Lean.Syntax | `(lexprmono| $f:lfvarmono) => elabLFVarMono Identifier f | `(lexprmono| λ $e:lexprmono) => do let e' ← elabLExprMono Identifier e - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) return mkAppN (.const ``LExpr.absUntyped []) #[typeTypeExpr, MkIdent.toExpr Identifier, e'] | `(lexprmono| λ ($mty:lmonoty): $e:lexprmono) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy mty - let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Unit)) lmonoty + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) lmonoty let e' ← elabLExprMono Identifier e - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) return mkAppN (.const ``LExpr.abs []) #[typeTypeExpr, MkIdent.toExpr Identifier, lmonoty, e'] | `(lexprmono| ∀ $e:lexprmono) => do let e' ← elabLExprMono Identifier e - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) return mkAppN (.const ``LExpr.allUntyped []) #[typeTypeExpr, MkIdent.toExpr Identifier, e'] | `(lexprmono| ∀ {$tr}$e:lexprmono) => do let e' ← elabLExprMono Identifier e let tr' ← elabLExprMono Identifier tr - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) return mkAppN (.const ``LExpr.allUntypedTr []) #[typeTypeExpr, MkIdent.toExpr Identifier, tr', e'] | `(lexprmono| ∀ ($mty:lmonoty): $e:lexprmono) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy mty - let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Unit)) lmonoty + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) lmonoty let e' ← elabLExprMono Identifier e - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) return mkAppN (.const ``LExpr.all []) #[typeTypeExpr, MkIdent.toExpr Identifier, lmonoty, e'] | `(lexprmono| ∀ ($mty:lmonoty):{$tr} $e:lexprmono) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy mty - let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Unit)) lmonoty + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) lmonoty let e' ← elabLExprMono Identifier e let tr' ← elabLExprMono Identifier tr - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) return mkAppN (.const ``LExpr.allTr []) #[typeTypeExpr, MkIdent.toExpr Identifier, lmonoty, tr', e'] | `(lexprmono| ∃ ($mty:lmonoty): $e:lexprmono) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy mty - let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Unit)) lmonoty + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) lmonoty let e' ← elabLExprMono Identifier e - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) return mkAppN (.const ``LExpr.exist []) #[typeTypeExpr, MkIdent.toExpr Identifier, lmonoty, e'] | `(lexprmono| ∃ ($mty:lmonoty):{$tr} $e:lexprmono) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy mty - let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Unit)) lmonoty + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) lmonoty let e' ← elabLExprMono Identifier e let tr' ← elabLExprMono Identifier tr - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) return mkAppN (.const ``LExpr.existTr []) #[typeTypeExpr, MkIdent.toExpr Identifier, lmonoty, tr', e'] | `(lexprmono| ∃ $e:lexprmono) => do let e' ← elabLExprMono Identifier e @@ -508,18 +508,18 @@ partial def elabLExprMono (Identifier : Type) [MkIdent Identifier] : Lean.Syntax | `(lexprmono| ($e1:lexprmono $e2:lexprmono)) => do let e1' ← elabLExprMono Identifier e1 let e2' ← elabLExprMono Identifier e2 - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) return mkAppN (.const ``LExpr.app []) #[typeTypeExpr, MkIdent.toExpr Identifier, e1', e2'] | `(lexprmono| $e1:lexprmono == $e2:lexprmono) => do let e1' ← elabLExprMono Identifier e1 let e2' ← elabLExprMono Identifier e2 - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) return mkAppN (.const ``LExpr.eq []) #[typeTypeExpr, MkIdent.toExpr Identifier, e1', e2'] | `(lexprmono| if $e1:lexprmono then $e2:lexprmono else $e3:lexprmono) => do let e1' ← elabLExprMono Identifier e1 let e2' ← elabLExprMono Identifier e2 let e3' ← elabLExprMono Identifier e3 - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Unit)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) return mkAppN (.const ``LExpr.ite []) #[typeTypeExpr, MkIdent.toExpr Identifier, e1', e2', e3'] | `(lexprmono| ($e:lexprmono)) => elabLExprMono Identifier e | _ => throwUnsupportedSyntax @@ -544,7 +544,8 @@ open LTy.Syntax #guard_msgs in #check esM[((λ %0) #5)] -/-- info: (abs (some (LMonoTy.tcons Unit "bool" [] ())) (bvar 0)).app (const "true" none) : LExpr LMonoTy String -/ +/-- info: (abs (some (LMonoTy.tcons LTyRestrict "bool" [] LTyRestrict.nodata)) (bvar 0)).app + (const "true" none) : LExpr LMonoTy String -/ #guard_msgs in #check esM[((λ (bool): %0) #true)] @@ -556,44 +557,49 @@ open LTy.Syntax #guard_msgs in #check esM[(∃ %0 == #5)] -/-- info: exist (some (LMonoTy.tcons Unit "int" [] ())) ((bvar 0).eq (const "5" none)) : LExpr LMonoTy String -/ +/-- info: exist (some (LMonoTy.tcons LTyRestrict "int" [] LTyRestrict.nodata)) + ((bvar 0).eq (const "5" none)) : LExpr LMonoTy String -/ #guard_msgs in #check esM[(∃ (int): %0 == #5)] -/-- info: fvar "x" (some (LMonoTy.tcons Unit "bool" [] ())) : LExpr LMonoTy String -/ +/-- info: fvar "x" (some (LMonoTy.tcons LTyRestrict "bool" [] LTyRestrict.nodata)) : LExpr LMonoTy String -/ #guard_msgs in #check esM[(x : bool)] -- axiom [updateSelect]: forall m: Map k v, kk: k, vv: v :: m[kk := vv][kk] == vv; /-- -info: all (some (LMonoTy.tcons Unit "Map" [LMonoTy.ftvar Unit "k", LMonoTy.ftvar Unit "v"] ())) - (all (some (LMonoTy.ftvar Unit "k")) - (all (some (LMonoTy.ftvar Unit "v")) +info: all + (some + (LMonoTy.tcons LTyRestrict "Map" [LMonoTy.ftvar LTyRestrict "k", LMonoTy.ftvar LTyRestrict "v"] LTyRestrict.nodata)) + (all (some (LMonoTy.ftvar LTyRestrict "k")) + (all (some (LMonoTy.ftvar LTyRestrict "v")) ((((op "select" (some - (LMonoTy.tcons Unit "Map" - [LMonoTy.ftvar Unit "k", - LMonoTy.tcons Unit "arrow" - [LMonoTy.ftvar Unit "v", - LMonoTy.tcons Unit "arrow" [LMonoTy.ftvar Unit "k", LMonoTy.ftvar Unit "v"] ()] - ()] - ()))).app + (LMonoTy.tcons LTyRestrict "Map" + [LMonoTy.ftvar LTyRestrict "k", + LMonoTy.tcons LTyRestrict "arrow" + [LMonoTy.ftvar LTyRestrict "v", + LMonoTy.tcons LTyRestrict "arrow" + [LMonoTy.ftvar LTyRestrict "k", LMonoTy.ftvar LTyRestrict "v"] LTyRestrict.nodata] + LTyRestrict.nodata] + LTyRestrict.nodata))).app ((((op "update" (some - (LMonoTy.tcons Unit "Map" - [LMonoTy.ftvar Unit "k", - LMonoTy.tcons Unit "arrow" - [LMonoTy.ftvar Unit "v", - LMonoTy.tcons Unit "arrow" - [LMonoTy.ftvar Unit "k", - LMonoTy.tcons Unit "arrow" - [LMonoTy.ftvar Unit "v", - LMonoTy.tcons Unit "Map" - [LMonoTy.ftvar Unit "k", LMonoTy.ftvar Unit "v"] ()] - ()] - ()] - ()] - ()))).app + (LMonoTy.tcons LTyRestrict "Map" + [LMonoTy.ftvar LTyRestrict "k", + LMonoTy.tcons LTyRestrict "arrow" + [LMonoTy.ftvar LTyRestrict "v", + LMonoTy.tcons LTyRestrict "arrow" + [LMonoTy.ftvar LTyRestrict "k", + LMonoTy.tcons LTyRestrict "arrow" + [LMonoTy.ftvar LTyRestrict "v", + LMonoTy.tcons LTyRestrict "Map" + [LMonoTy.ftvar LTyRestrict "k", LMonoTy.ftvar LTyRestrict "v"] + LTyRestrict.nodata] + LTyRestrict.nodata] + LTyRestrict.nodata] + LTyRestrict.nodata] + LTyRestrict.nodata))).app (bvar 2)).app (bvar 1)).app (bvar 0))).app @@ -866,7 +872,8 @@ open LTy.Syntax #guard_msgs in #check es[((λ %0) #5)] -/-- info: (abs (some (LTy.forAll [] (LMonoTy.tcons Unit "bool" [] ()))) (bvar 0)).app (const "true" none) : LExpr LTy String -/ +/-- info: (abs (some (LTy.forAll [] (LMonoTy.tcons LTyRestrict "bool" [] LTyRestrict.nodata))) (bvar 0)).app + (const "true" none) : LExpr LTy String -/ #guard_msgs in #check es[((λ (bool): %0) #true)] @@ -878,46 +885,53 @@ open LTy.Syntax #guard_msgs in #check es[(∃ %0 == #5)] -/-- info: exist (some (LTy.forAll [] (LMonoTy.tcons Unit "int" [] ()))) ((bvar 0).eq (const "5" none)) : LExpr LTy String -/ +/-- info: exist (some (LTy.forAll [] (LMonoTy.tcons LTyRestrict "int" [] LTyRestrict.nodata))) + ((bvar 0).eq (const "5" none)) : LExpr LTy String -/ #guard_msgs in #check es[(∃ (int): %0 == #5)] -/-- info: fvar "x" (some (LTy.forAll [] (LMonoTy.tcons Unit "bool" [] ()))) : LExpr LTy String -/ +/-- info: fvar "x" (some (LTy.forAll [] (LMonoTy.tcons LTyRestrict "bool" [] LTyRestrict.nodata))) : LExpr LTy String -/ #guard_msgs in #check es[(x : bool)] -- axiom [updateSelect]: forall m: Map k v, kk: k, vv: v :: m[kk := vv][kk] == vv; /-- -info: all (some (LTy.forAll [] (LMonoTy.tcons Unit "Map" [LMonoTy.ftvar Unit "k", LMonoTy.ftvar Unit "v"] ()))) - (all (some (LTy.forAll [] (LMonoTy.ftvar Unit "k"))) - (all (some (LTy.forAll [] (LMonoTy.ftvar Unit "v"))) +info: all + (some + (LTy.forAll [] + (LMonoTy.tcons LTyRestrict "Map" [LMonoTy.ftvar LTyRestrict "k", LMonoTy.ftvar LTyRestrict "v"] + LTyRestrict.nodata))) + (all (some (LTy.forAll [] (LMonoTy.ftvar LTyRestrict "k"))) + (all (some (LTy.forAll [] (LMonoTy.ftvar LTyRestrict "v"))) ((((op "select" (some (LTy.forAll [] - (LMonoTy.tcons Unit "Map" - [LMonoTy.ftvar Unit "k", - LMonoTy.tcons Unit "arrow" - [LMonoTy.ftvar Unit "v", - LMonoTy.tcons Unit "arrow" [LMonoTy.ftvar Unit "k", LMonoTy.ftvar Unit "v"] ()] - ()] - ())))).app + (LMonoTy.tcons LTyRestrict "Map" + [LMonoTy.ftvar LTyRestrict "k", + LMonoTy.tcons LTyRestrict "arrow" + [LMonoTy.ftvar LTyRestrict "v", + LMonoTy.tcons LTyRestrict "arrow" + [LMonoTy.ftvar LTyRestrict "k", LMonoTy.ftvar LTyRestrict "v"] LTyRestrict.nodata] + LTyRestrict.nodata] + LTyRestrict.nodata)))).app ((((op "update" (some (LTy.forAll [] - (LMonoTy.tcons Unit "Map" - [LMonoTy.ftvar Unit "k", - LMonoTy.tcons Unit "arrow" - [LMonoTy.ftvar Unit "v", - LMonoTy.tcons Unit "arrow" - [LMonoTy.ftvar Unit "k", - LMonoTy.tcons Unit "arrow" - [LMonoTy.ftvar Unit "v", - LMonoTy.tcons Unit "Map" - [LMonoTy.ftvar Unit "k", LMonoTy.ftvar Unit "v"] ()] - ()] - ()] - ()] - ())))).app + (LMonoTy.tcons LTyRestrict "Map" + [LMonoTy.ftvar LTyRestrict "k", + LMonoTy.tcons LTyRestrict "arrow" + [LMonoTy.ftvar LTyRestrict "v", + LMonoTy.tcons LTyRestrict "arrow" + [LMonoTy.ftvar LTyRestrict "k", + LMonoTy.tcons LTyRestrict "arrow" + [LMonoTy.ftvar LTyRestrict "v", + LMonoTy.tcons LTyRestrict "Map" + [LMonoTy.ftvar LTyRestrict "k", LMonoTy.ftvar LTyRestrict "v"] + LTyRestrict.nodata] + LTyRestrict.nodata] + LTyRestrict.nodata] + LTyRestrict.nodata] + LTyRestrict.nodata)))).app (bvar 2)).app (bvar 1)).app (bvar 0))).app diff --git a/Strata/DL/Lambda/LExprT.lean b/Strata/DL/Lambda/LExprT.lean index dd7a30c2..e626b35d 100644 --- a/Strata/DL/Lambda/LExprT.lean +++ b/Strata/DL/Lambda/LExprT.lean @@ -226,8 +226,8 @@ def inferConst (T : (TEnv Identifier)) (c : String) (cty : Option LMonoTy) : {@LExpr.const LMonoTy Identifier c cty}\n\ Known Types: {T.knownTypes}" -- Annotated BitVecs - | c, some (LMonoTy.bitvec n) => - let ty := LMonoTy.bitvec n + | c, some (bitvec n) => + let ty := bitvec n if { name := "bitvec", arity := 1 } ∈ T.knownTypes then (.ok (ty, T)) else diff --git a/Strata/DL/Lambda/LExprTypeEnv.lean b/Strata/DL/Lambda/LExprTypeEnv.lean index 184549c0..52c24711 100644 --- a/Strata/DL/Lambda/LExprTypeEnv.lean +++ b/Strata/DL/Lambda/LExprTypeEnv.lean @@ -193,7 +193,7 @@ def KnownType.toLTy (k : KnownType) : LTy := def LTy.toKnownType! (lty : LTy) : KnownType := match lty with | .forAll _ (.tcons name args) => { name, arity := args.length } - | .forAll [] (.bitvec _) => { name := "bitvec", arity := 1 } + | .forAll [] (bitvec _) => { name := "bitvec", arity := 1 } | _ => panic! s!"Unsupported known type: {lty}" instance : ToFormat KnownType where @@ -421,10 +421,10 @@ def LMonoTy.aliasDef? (mty : LMonoTy) (T : (TEnv Identifier)) : (Option LMonoTy -- We can't have a free variable be the LHS of an alias definition because -- then it will unify with every type. (none, T) - | .bitvec _ => + | bitvec _ => -- A bitvector cannot be a type alias. (none, T) - | .tcons name args => + | .tcons name args r => match T.context.aliases.find? (fun a => a.name == name && a.typeArgs.length == args.length) with | none => (none, T) | some alias => @@ -538,12 +538,12 @@ partial def LMonoTy.resolveAliases (mty : LMonoTy) (T : TEnv Identifier) : (Opti | none => match mty with | .ftvar _ => (some mty, T) - | .bitvec _ => (some mty, T) - | .tcons name mtys => + | bitvec _ => (some mty, T) + | .tcons name mtys r => let (maybe_mtys, T) := LMonoTys.resolveAliases mtys T.context.aliases T match maybe_mtys with | none => (none, T) - | some mtys' => (some (.tcons name mtys'), T) + | some mtys' => (some (.tcons name mtys' r), T) /-- De-alias `mtys`, including at the subtrees. @@ -610,8 +610,8 @@ de-aliased. -/ def LMonoTy.knownInstance (ty : LMonoTy) (ks : KnownTypes) : Bool := match ty with - | .ftvar _ | .bitvec _ => true - | .tcons name args => + | .ftvar _ | bitvec _ => true + | .tcons name args _ => (ks.contains { name := name, arity := args.length }) && LMonoTys.knownInstances args ks diff --git a/Strata/DL/Lambda/LTy.lean b/Strata/DL/Lambda/LTy.lean index 6fe8d3f2..87413d35 100644 --- a/Strata/DL/Lambda/LTy.lean +++ b/Strata/DL/Lambda/LTy.lean @@ -25,21 +25,33 @@ abbrev TyIdentifier := String instance : Coe String TyIdentifier where coe := id +/-- +An additional parameter that restricts the values of a type +or attaches additional metadata to a type. +For now, only bitvectors are supported. +-/ +inductive LTyRestrict : Type where + /-- Special support for bitvector types of every size. -/ + | bitvecdata (size: Nat) + | nodata +deriving Inhabited, Repr, DecidableEq + /-- Types in Lambda: these are mono-types. Note that all free type variables (`.ftvar`) are implicitly universally quantified. -/ -inductive LMonoTy (TypeRestrictions: Type := Unit) : Type where +inductive LMonoTy (TypeRestrictions: Type := LTyRestrict) : Type where /-- Type variable. -/ | ftvar (name : TyIdentifier) /-- Type constructor. -/ - | tcons (name : String) (args : List (LMonoTy TypeRestrictions)) (r: TypeRestrictions := by exact ()) - /-- Special support for bitvector types of every size. -/ - | bitvec (size : Nat) + | tcons (name : String) (args : List (LMonoTy TypeRestrictions)) (r: TypeRestrictions := by exact .nodata) deriving Inhabited, Repr abbrev LMonoTys := List LMonoTy +@[match_pattern] +def bitvec (n: Nat) := LMonoTy.tcons "bitvec" [] (LTyRestrict.bitvecdata n) + @[match_pattern] def LMonoTy.bool : LMonoTy := .tcons "bool" [] @@ -54,23 +66,23 @@ def LMonoTy.real : LMonoTy := @[match_pattern] def LMonoTy.bv1 : LMonoTy := - .bitvec 1 + bitvec 1 @[match_pattern] def LMonoTy.bv8 : LMonoTy := - .bitvec 8 + bitvec 8 @[match_pattern] def LMonoTy.bv16 : LMonoTy := - .bitvec 16 + bitvec 16 @[match_pattern] def LMonoTy.bv32 : LMonoTy := - .bitvec 32 + bitvec 32 @[match_pattern] def LMonoTy.bv64 : LMonoTy := - .bitvec 64 + bitvec 64 @[match_pattern] def LMonoTy.string : LMonoTy := @@ -126,7 +138,6 @@ types. So we define our own induction principle below. @[induction_eliminator] theorem LMonoTy.induct {Meta: Type} {P : LMonoTy Meta → Prop} (ftvar : ∀f, P (.ftvar f)) - (bitvec : ∀n, P (.bitvec n)) (tcons : ∀name args r, (∀ ty ∈ args, P ty) → P (.tcons name args r)) : ∀ ty, P ty := by intro n @@ -144,8 +155,7 @@ Compute the size of `ty` as a tree. def LMonoTy.size (ty : LMonoTy) : Nat := match ty with | .ftvar _ => 1 - | .tcons _ args => 1 + LMonoTys.size args - | .bitvec _ => 1 + | .tcons _ args _ => 1 + LMonoTys.size args def LMonoTys.size (args : LMonoTys) : Nat := match args with @@ -165,9 +175,8 @@ Boolean equality for `LMonoTy`. def LMonoTy.BEq (x y : LMonoTy) : Bool := match x, y with | .ftvar i, .ftvar j => i == j - | .bitvec i, .bitvec j => i == j - | .tcons i1 j1, .tcons i2 j2 => - i1 == i2 && j1.length == j2.length && go j1 j2 + | .tcons i1 j1 r1, .tcons i2 j2 r2 => + i1 == i2 && j1.length == j2.length && r1 == r2 && go j1 j2 | _, _ => false where go j1 j2 := match j1, j2 with @@ -194,13 +203,11 @@ instance : DecidableEq LMonoTy := induction x generalizing y case ftvar => unfold LMonoTy.BEq at h <;> split at h <;> try simp_all - case bitvec => - unfold LMonoTy.BEq at h <;> split at h <;> try simp_all case tcons => rename_i name args r ih cases y <;> try simp_all [LMonoTy.BEq] rename_i name' args' r' - obtain ⟨⟨h1, h2⟩, h3⟩ := h + obtain ⟨⟨⟨h1, h2⟩, _⟩, h3⟩ := h induction args generalizing args' case nil => unfold List.length at h2; split at h2 <;> simp_all case cons head' tail' ih' => @@ -213,8 +220,6 @@ instance : DecidableEq LMonoTy := isFalse (by induction x generalizing y case ftvar => cases y <;> try simp_all [LMonoTy.BEq] - case bitvec n => - cases y <;> try simp_all [LMonoTy.BEq] case tcons name args r ih => cases y <;> try simp_all [LMonoTy.BEq] rename_i name' args' r' @@ -244,8 +249,7 @@ it. def LMonoTy.freeVars (mty : LMonoTy) : List TyIdentifier := match mty with | .ftvar x => [x] - | .bitvec _ => [] - | .tcons _ ltys => LMonoTys.freeVars ltys + | .tcons _ ltys _ => LMonoTys.freeVars ltys /-- Get the free type variables in monotypes `mtys`, which are simply the `.ftvar`s @@ -267,8 +271,7 @@ Get all type constructors in monotype `mty`. def LMonoTy.getTyConstructors (mty : LMonoTy) : List LMonoTy := match mty with | .ftvar _ => [] - | .bitvec _ => [] - | .tcons name mtys => + | .tcons name mtys _ => let typeargs := List.replicate mtys.length "_dummy" let args := typeargs.mapIdx (fun i elem => LMonoTy.ftvar (elem ++ toString i)) let mty := .tcons name args @@ -279,10 +282,13 @@ def LMonoTy.getTyConstructors (mty : LMonoTy) : List LMonoTy := | [] => [] | m :: mrest => LMonoTy.getTyConstructors m ++ go mrest /-- -info: [Lambda.LMonoTy.tcons "arrow" [Lambda.LMonoTy.ftvar "_dummy0", Lambda.LMonoTy.ftvar "_dummy1"] (), - Lambda.LMonoTy.tcons "bool" [] (), - Lambda.LMonoTy.tcons "foo" [Lambda.LMonoTy.ftvar "_dummy0"] (), - Lambda.LMonoTy.tcons "a" [Lambda.LMonoTy.ftvar "_dummy0", Lambda.LMonoTy.ftvar "_dummy1"] ()] +info: [Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.ftvar "_dummy0", Lambda.LMonoTy.ftvar "_dummy1"] + (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "foo" [Lambda.LMonoTy.ftvar "_dummy0"] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "a" [Lambda.LMonoTy.ftvar "_dummy0", Lambda.LMonoTy.ftvar "_dummy1"] (Lambda.LTyRestrict.nodata)] -/ #guard_msgs in #eval LMonoTy.getTyConstructors @@ -355,8 +361,8 @@ instance : ToString LMonoTy where private def formatLMonoTy (lmonoty : LMonoTy) : Format := match lmonoty with | .ftvar x => toString x - | .bitvec n => f!"bv{n}" - | .tcons name tys => + | bitvec n => f!"bv{n}" + | .tcons name tys _ => if tys.isEmpty then f!"{name}" else @@ -408,18 +414,18 @@ scoped syntax "(" lmonoty ")" : lmonoty open Lean Elab Meta in partial def elabLMonoTy : Lean.Syntax → MetaM Expr | `(lmonoty| %$f:ident) => do - mkAppOptM ``LMonoTy.ftvar #[mkConst `Unit, mkStrLit (toString f.getId)] + mkAppOptM ``LMonoTy.ftvar #[mkConst `Lambda.LTyRestrict, mkStrLit (toString f.getId)] | `(lmonoty| $ty1:lmonoty → $ty2:lmonoty) => do let ty1' ← elabLMonoTy ty1 let ty2' ← elabLMonoTy ty2 - let tys ← mkListLit (.app (mkConst ``LMonoTy) (mkConst `Unit)) [ty1', ty2'] - mkAppM ``LMonoTy.tcons #[(mkStrLit "arrow"), tys, (mkConst `Unit.unit)] + let tys ← mkListLit (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) [ty1', ty2'] + mkAppM ``LMonoTy.tcons #[(mkStrLit "arrow"), tys, (mkConst `Lambda.LTyRestrict.nodata)] | `(lmonoty| int) => do - let argslist ← mkListLit (.app (mkConst ``LMonoTy) (mkConst `Unit)) [] - mkAppM ``LMonoTy.tcons #[(mkStrLit "int"), argslist, (mkConst `Unit.unit)] + let argslist ← mkListLit (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) [] + mkAppM ``LMonoTy.tcons #[(mkStrLit "int"), argslist, (mkConst `Lambda.LTyRestrict.nodata)] | `(lmonoty| bool) => do - let argslist ← mkListLit (.app (mkConst ``LMonoTy) (mkConst `Unit)) [] - mkAppM ``LMonoTy.tcons #[(mkStrLit "bool"), argslist, (mkConst `Unit.unit)] + let argslist ← mkListLit (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) [] + mkAppM ``LMonoTy.tcons #[(mkStrLit "bool"), argslist, (mkConst `Lambda.LTyRestrict.nodata)] | `(lmonoty| bv1) => mkAppM ``LMonoTy.bv1 #[] | `(lmonoty| bv8) => mkAppM ``LMonoTy.bv8 #[] | `(lmonoty| bv16) => mkAppM ``LMonoTy.bv16 #[] @@ -427,8 +433,8 @@ partial def elabLMonoTy : Lean.Syntax → MetaM Expr | `(lmonoty| bv64) => mkAppM ``LMonoTy.bv64 #[] | `(lmonoty| $i:ident $args:lmonoty*) => do let args' ← go args - let argslist ← mkListLit (.app (mkConst ``LMonoTy) (mkConst `Unit)) args'.toList - mkAppM ``LMonoTy.tcons #[(mkStrLit (toString i.getId)), argslist, (mkConst `Unit.unit)] + let argslist ← mkListLit (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) args'.toList + mkAppM ``LMonoTy.tcons #[(mkStrLit (toString i.getId)), argslist, (mkConst `Lambda.LTyRestrict.nodata)] | `(lmonoty| ($ty:lmonoty)) => elabLMonoTy ty | _ => throwUnsupportedSyntax where go (args : TSyntaxArray `lmonoty) : MetaM (Array Expr) := do @@ -440,29 +446,34 @@ partial def elabLMonoTy : Lean.Syntax → MetaM Expr elab "mty[" ty:lmonoty "]" : term => elabLMonoTy ty -/-- info: LMonoTy.tcons Unit "list" [LMonoTy.tcons Unit "int" [] ()] () : LMonoTy -/ +/-- info: LMonoTy.tcons LTyRestrict "list" [LMonoTy.tcons LTyRestrict "int" [] LTyRestrict.nodata] LTyRestrict.nodata : LMonoTy -/ #guard_msgs in #check mty[list int] -/-- info: LMonoTy.tcons Unit "pair" [LMonoTy.tcons Unit "int" [] (), LMonoTy.tcons Unit "bool" [] ()] () : LMonoTy -/ +/-- info: LMonoTy.tcons LTyRestrict "pair" + [LMonoTy.tcons LTyRestrict "int" [] LTyRestrict.nodata, LMonoTy.tcons LTyRestrict "bool" [] LTyRestrict.nodata] + LTyRestrict.nodata : LMonoTy -/ #guard_msgs in #check mty[pair int bool] /-- -info: LMonoTy.tcons Unit "arrow" - [LMonoTy.tcons Unit "Map" [LMonoTy.ftvar Unit "k", LMonoTy.ftvar Unit "v"] (), - LMonoTy.tcons Unit "arrow" [LMonoTy.ftvar Unit "k", LMonoTy.ftvar Unit "v"] ()] - () : LMonoTy +info: LMonoTy.tcons LTyRestrict "arrow" + [LMonoTy.tcons LTyRestrict "Map" [LMonoTy.ftvar LTyRestrict "k", LMonoTy.ftvar LTyRestrict "v"] LTyRestrict.nodata, + LMonoTy.tcons LTyRestrict "arrow" [LMonoTy.ftvar LTyRestrict "k", LMonoTy.ftvar LTyRestrict "v"] LTyRestrict.nodata] + LTyRestrict.nodata : LMonoTy -/ #guard_msgs in #check mty[(Map %k %v) → %k → %v] /-- -info: LMonoTy.tcons Unit "arrow" - [LMonoTy.ftvar Unit "a", - LMonoTy.tcons Unit "arrow" - [LMonoTy.ftvar Unit "b", LMonoTy.tcons Unit "arrow" [LMonoTy.ftvar Unit "c", LMonoTy.ftvar Unit "d"] ()] ()] - () : LMonoTy +info: LMonoTy.tcons LTyRestrict "arrow" + [LMonoTy.ftvar LTyRestrict "a", + LMonoTy.tcons LTyRestrict "arrow" + [LMonoTy.ftvar LTyRestrict "b", + LMonoTy.tcons LTyRestrict "arrow" [LMonoTy.ftvar LTyRestrict "c", LMonoTy.ftvar LTyRestrict "d"] + LTyRestrict.nodata] + LTyRestrict.nodata] + LTyRestrict.nodata : LMonoTy -/ #guard_msgs in #check mty[%a → %b → %c → %d] @@ -488,15 +499,17 @@ partial def elabLTy : Lean.Syntax → MetaM Expr elab "t[" lsch:lty "]" : term => elabLTy lsch -/-- info: forAll ["α"] (LMonoTy.tcons Unit "myType" [LMonoTy.ftvar Unit "α"] ()) : LTy -/ +/-- info: forAll ["α"] (LMonoTy.tcons LTyRestrict "myType" [LMonoTy.ftvar LTyRestrict "α"] LTyRestrict.nodata) : LTy -/ #guard_msgs in #check t[∀α. myType %α] /-- info: forAll ["α"] - (LMonoTy.tcons Unit "arrow" - [LMonoTy.ftvar Unit "α", LMonoTy.tcons Unit "arrow" [LMonoTy.ftvar Unit "α", LMonoTy.tcons Unit "int" [] ()] ()] - ()) : LTy + (LMonoTy.tcons LTyRestrict "arrow" + [LMonoTy.ftvar LTyRestrict "α", + LMonoTy.tcons LTyRestrict "arrow" + [LMonoTy.ftvar LTyRestrict "α", LMonoTy.tcons LTyRestrict "int" [] LTyRestrict.nodata] LTyRestrict.nodata] + LTyRestrict.nodata) : LTy -/ #guard_msgs in #check t[∀α. %α → %α → int] diff --git a/Strata/DL/Lambda/LTyUnify.lean b/Strata/DL/Lambda/LTyUnify.lean index 11a6d5dd..f0ebbc24 100644 --- a/Strata/DL/Lambda/LTyUnify.lean +++ b/Strata/DL/Lambda/LTyUnify.lean @@ -121,7 +121,6 @@ def LMonoTy.subst (S : Subst) (mty : LMonoTy) : LMonoTy := match mty with | .ftvar x => match S.find? x with | some sty => sty | none => mty - | .bitvec _ => mty | .tcons name ltys r => .tcons name (LMonoTys.subst S ltys) r /-- @@ -196,10 +195,6 @@ theorem LMonoTy.subst_keys_not_in_substituted_type (h : SubstWF S) : have := @Map.find?_of_not_mem_values _ _ i _ S simp_all exact ne_of_mem_of_not_mem hid this - case bitvec n => - simp_all [LMonoTy.subst] - unfold LMonoTy.freeVars - simp case tcons name args r h1 => simp_all simp [subst] @@ -241,8 +236,6 @@ theorem LMonoTy.freeVars_of_subst_subset (S : Subst) (mty : LMonoTy) : apply @Map.find?_mem_values _ _ x sty _ S h_find · -- Case: S.find? x = none simp [freeVars] - case bitvec n => - simp [subst] case tcons name args r ih => simp [LMonoTy.subst, LMonoTy.freeVars] induction args @@ -963,25 +956,25 @@ def Constraint.unifyOne (c : Constraint) (S : SubstInfo) : have h_sub2 : Subst.freeVars_subset_prop [(orig_lty, LMonoTy.ftvar id)] newS S := by simp_all [Subst.freeVars_subset_prop_single_constraint_comm] .ok { newS := newS, goodSubset := by all_goals simp [h_sub1, h_sub2] } - | .bitvec n1, .bitvec n2 => + | bitvec n1, bitvec n2 => if _h7 : n1 == n2 then .ok { newS := SubstInfo.mk [] (by simp [SubstWF]), goodSubset := by grind } else .error f!"Cannot unify differently sized bitvector types {t1} and {t2}!" - | .tcons name1 args1, .tcons name2 args2 => do + | bitvec _, .tcons _ _ _ => + .error f!"Cannot unify bv type {t1} and type constructor {t2}!" + | .tcons _ _ _, bitvec _ => + .error f!"Cannot unify type constructor {t1} and bv type {t2}!" + | .tcons name1 args1 r1, .tcons name2 args2 r2 => do if _h6 : name1 == name2 && args1.length == args2.length then let new_constraints := List.zip args1 args2 let relS ← Constraints.unifyCore new_constraints S have h_sub : Subst.freeVars_subset_prop - [(LMonoTy.tcons name1 args1, LMonoTy.tcons name2 args2)] relS.newS S := by + [(LMonoTy.tcons name1 args1 r1, LMonoTy.tcons name2 args2 r2)] relS.newS S := by exact Subst.freeVars_subset_prop_of_tcons S name1 name2 args1 args2 rfl relS .ok { newS := relS.newS, goodSubset := by simp [h_sub] } else .error f!"Cannot unify differently named type constructors {t1} and {t2}!" - | .bitvec _, .tcons _ _ => - .error f!"Cannot unify bv type {t1} and type constructor {t2}!" - | .tcons _ _, .bitvec _ => - .error f!"Cannot unify type constructor {t1} and bv type {t2}!" termination_by ((((Constraints.freeVars [c]) ++ S.subst.freeVars).dedup.length), Constraints.size [c], 0) diff --git a/Strata/Languages/Boogie/DDMTransform/Translate.lean b/Strata/Languages/Boogie/DDMTransform/Translate.lean index 0b9b8296..0ab64511 100644 --- a/Strata/Languages/Boogie/DDMTransform/Translate.lean +++ b/Strata/Languages/Boogie/DDMTransform/Translate.lean @@ -202,11 +202,11 @@ partial def translateLMonoTy (bindings : TransBindings) (arg : Arg) : let .type tp := arg | TransM.error s!"translateLMonoTy expected type {repr arg}" match tp with - | .ident q`Boogie.bv1 #[] => pure <| .bitvec 1 - | .ident q`Boogie.bv8 #[] => pure <| .bitvec 8 - | .ident q`Boogie.bv16 #[] => pure <| .bitvec 16 - | .ident q`Boogie.bv32 #[] => pure <| .bitvec 32 - | .ident q`Boogie.bv64 #[] => pure <| .bitvec 64 + | .ident q`Boogie.bv1 #[] => pure <| bitvec 1 + | .ident q`Boogie.bv8 #[] => pure <| bitvec 8 + | .ident q`Boogie.bv16 #[] => pure <| bitvec 16 + | .ident q`Boogie.bv32 #[] => pure <| bitvec 32 + | .ident q`Boogie.bv64 #[] => pure <| bitvec 64 | .ident i argst => let argst' ← translateLMonoTys bindings (argst.map Arg.type) pure <| (.tcons i.name argst'.toList.reverse) @@ -399,7 +399,7 @@ def dealiasTypeArg (p : Program) (a : Arg) : Arg := def isArithTy : LMonoTy → Bool | .int => true | .real => true -| .bitvec _ => true +| bitvec _ => true | _ => false def translateFn (ty? : Option LMonoTy) (q : QualifiedIdent) : TransM Boogie.Expression.Expr := diff --git a/Strata/Languages/Boogie/Examples/DDMAxiomsExtraction.lean b/Strata/Languages/Boogie/Examples/DDMAxiomsExtraction.lean index 91e3399b..9ac2f236 100644 --- a/Strata/Languages/Boogie/Examples/DDMAxiomsExtraction.lean +++ b/Strata/Languages/Boogie/Examples/DDMAxiomsExtraction.lean @@ -239,7 +239,7 @@ info: #[{ name := { dialect := "Boogie", name := "command_typedecl" }, /-- info: [Lambda.LExpr.quant (Lambda.QuantifierKind.all) - (some (Lambda.LMonoTy.tcons "Map" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] ())) + (some (Lambda.LMonoTy.tcons "Map" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] (Lambda.LTyRestrict.nodata))) (Lambda.LExpr.bvar 0) (Lambda.LExpr.quant (Lambda.QuantifierKind.all) @@ -256,9 +256,15 @@ info: [Lambda.LExpr.quant u:select (some (Lambda.LMonoTy.tcons "arrow" - [Lambda.LMonoTy.tcons "Map" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] (), - Lambda.LMonoTy.tcons "arrow" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] ()] - ()))) + [Lambda.LMonoTy.tcons + "Map" + [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] + (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) (Lambda.LExpr.app (Lambda.LExpr.app (Lambda.LExpr.app @@ -266,17 +272,23 @@ info: [Lambda.LExpr.quant u:update (some (Lambda.LMonoTy.tcons "arrow" - [Lambda.LMonoTy.tcons "Map" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] (), + [Lambda.LMonoTy.tcons + "Map" + [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] + (Lambda.LTyRestrict.nodata), Lambda.LMonoTy.tcons "arrow" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.tcons "arrow" [Lambda.LMonoTy.ftvar "v", - Lambda.LMonoTy.tcons "Map" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] ()] - ()] - ()] - ()))) + Lambda.LMonoTy.tcons + "Map" + [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) (Lambda.LExpr.bvar 2)) (Lambda.LExpr.bvar 1)) (Lambda.LExpr.bvar 0))) @@ -284,7 +296,7 @@ info: [Lambda.LExpr.quant (Lambda.LExpr.bvar 0)))), Lambda.LExpr.quant (Lambda.QuantifierKind.all) - (some (Lambda.LMonoTy.tcons "Map" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] ())) + (some (Lambda.LMonoTy.tcons "Map" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] (Lambda.LTyRestrict.nodata))) (Lambda.LExpr.bvar 0) (Lambda.LExpr.quant (Lambda.QuantifierKind.all) @@ -305,9 +317,15 @@ info: [Lambda.LExpr.quant u:select (some (Lambda.LMonoTy.tcons "arrow" - [Lambda.LMonoTy.tcons "Map" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] (), - Lambda.LMonoTy.tcons "arrow" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] ()] - ()))) + [Lambda.LMonoTy.tcons + "Map" + [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] + (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) (Lambda.LExpr.app (Lambda.LExpr.app (Lambda.LExpr.app @@ -315,17 +333,23 @@ info: [Lambda.LExpr.quant u:update (some (Lambda.LMonoTy.tcons "arrow" - [Lambda.LMonoTy.tcons "Map" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] (), + [Lambda.LMonoTy.tcons + "Map" + [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] + (Lambda.LTyRestrict.nodata), Lambda.LMonoTy.tcons "arrow" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.tcons "arrow" [Lambda.LMonoTy.ftvar "v", - Lambda.LMonoTy.tcons "Map" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] ()] - ()] - ()] - ()))) + Lambda.LMonoTy.tcons + "Map" + [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) (Lambda.LExpr.bvar 3)) (Lambda.LExpr.bvar 1)) (Lambda.LExpr.bvar 0))) @@ -336,9 +360,15 @@ info: [Lambda.LExpr.quant u:select (some (Lambda.LMonoTy.tcons "arrow" - [Lambda.LMonoTy.tcons "Map" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] (), - Lambda.LMonoTy.tcons "arrow" [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] ()] - ()))) + [Lambda.LMonoTy.tcons + "Map" + [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] + (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.ftvar "k", Lambda.LMonoTy.ftvar "v"] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) (Lambda.LExpr.bvar 3)) (Lambda.LExpr.bvar 2))))))] -/ diff --git a/Strata/Languages/Boogie/Identifiers.lean b/Strata/Languages/Boogie/Identifiers.lean index 535a30ee..f2955c88 100644 --- a/Strata/Languages/Boogie/Identifiers.lean +++ b/Strata/Languages/Boogie/Identifiers.lean @@ -141,7 +141,9 @@ info: (Lambda.LExpr.op (BoogieIdent.unres "old") none).app open Lambda.LTy.Syntax in /-- info: Lambda.LExpr.fvar (BoogieIdent.unres "x") - (some (Lambda.LMonoTy.tcons Unit "bool" [] ())) : Lambda.LExpr Lambda.LMonoTy (Visibility × String) -/ + (some + (Lambda.LMonoTy.tcons Lambda.LTyRestrict "bool" [] + Lambda.LTyRestrict.nodata)) : Lambda.LExpr Lambda.LMonoTy (Visibility × String) -/ #guard_msgs in #check eb[(x : bool)] diff --git a/Strata/Languages/Boogie/SMTEncoder.lean b/Strata/Languages/Boogie/SMTEncoder.lean index c7ae1cc5..fa383a12 100644 --- a/Strata/Languages/Boogie/SMTEncoder.lean +++ b/Strata/Languages/Boogie/SMTEncoder.lean @@ -87,12 +87,12 @@ mutual def LMonoTy.toSMTType (ty : LMonoTy) (ctx : SMT.Context) : Except Format (TermType × SMT.Context) := do match ty with - | .bitvec n => .ok (.bitvec n, ctx) - | .tcons "bool" [] => .ok (.bool, ctx) - | .tcons "int" [] => .ok (.int, ctx) - | .tcons "real" [] => .ok (.real, ctx) - | .tcons "string" [] => .ok (.string, ctx) - | .tcons id args => + | bitvec n => .ok (TermType.bitvec n, ctx) + | .tcons "bool" [] _ => .ok (.bool, ctx) + | .tcons "int" [] _ => .ok (.int, ctx) + | .tcons "real" [] _ => .ok (.real, ctx) + | .tcons "string" [] _ => .ok (.string, ctx) + | .tcons id args _ => let ctx := ctx.addSort { name := id, arity := args.length } let (args', ctx) ← LMonoTys.toSMTType args ctx .ok ((.constr id args'), ctx) @@ -141,7 +141,7 @@ partial def toSMTTerm (E : Env) (bvs : BoundVars) (e : LExpr LMonoTy BoogieIdent | none => .error f!"Unexpected real constant {e}" | some r => .ok ((Term.real r), ctx) - | .bitvec n => + | bitvec n => match e.denoteBitVec n with | none => .error f!"Unexpected bv constant {e}" From c6443418b350d84bdda8779719e27ebcb86f429b Mon Sep 17 00:00:00 2001 From: Josh Cohen Date: Mon, 22 Sep 2025 17:14:48 -0400 Subject: [PATCH 03/19] Minor changes to fix bounded types after merge --- Strata/DL/Lambda/LExpr.lean | 2 +- Strata/DL/Lambda/LExprT.lean | 4 ++-- Strata/DL/Lambda/LExprTypeEnv.lean | 6 +++--- Strata/DL/Lambda/LTy.lean | 4 ++-- Strata/DL/Lambda/LTyUnify.lean | 6 +++--- Strata/Languages/Boogie/DDMTransform/Translate.lean | 12 ++++++------ Strata/Languages/Boogie/SMTEncoder.lean | 4 ++-- 7 files changed, 19 insertions(+), 19 deletions(-) diff --git a/Strata/DL/Lambda/LExpr.lean b/Strata/DL/Lambda/LExpr.lean index a4e22671..5e4e9ff9 100644 --- a/Strata/DL/Lambda/LExpr.lean +++ b/Strata/DL/Lambda/LExpr.lean @@ -168,7 +168,7 @@ Note that we are type-agnostic here. -/ def denoteBitVec (n : Nat) (e : (LExpr LMonoTy Identifier)) : Option (BitVec n) := match e with - | .const x (.some (bitvec n')) => + | .const x (.some (.bitvec n')) => if n == n' then .map (.ofNat n) x.toNat? else none | .const x none => .map (.ofNat n) x.toNat? | _ => none diff --git a/Strata/DL/Lambda/LExprT.lean b/Strata/DL/Lambda/LExprT.lean index 5d577578..dd0b5103 100644 --- a/Strata/DL/Lambda/LExprT.lean +++ b/Strata/DL/Lambda/LExprT.lean @@ -278,8 +278,8 @@ def inferConst (T : (TEnv Identifier)) (c : String) (cty : Option LMonoTy) : {@LExpr.const LMonoTy Identifier c cty}\n\ Known Types: {T.knownTypes}" -- Annotated BitVecs - | c, some (bitvec n) => - let ty := bitvec n + | c, some (.bitvec n) => + let ty := .bitvec n if { name := "bitvec", arity := 1 } ∈ T.knownTypes then (.ok (ty, T)) else diff --git a/Strata/DL/Lambda/LExprTypeEnv.lean b/Strata/DL/Lambda/LExprTypeEnv.lean index 52c24711..4b6bb041 100644 --- a/Strata/DL/Lambda/LExprTypeEnv.lean +++ b/Strata/DL/Lambda/LExprTypeEnv.lean @@ -192,8 +192,8 @@ def KnownType.toLTy (k : KnownType) : LTy := def LTy.toKnownType! (lty : LTy) : KnownType := match lty with + | .forAll [] (.bitvec _) => { name := "bitvec", arity := 1 } | .forAll _ (.tcons name args) => { name, arity := args.length } - | .forAll [] (bitvec _) => { name := "bitvec", arity := 1 } | _ => panic! s!"Unsupported known type: {lty}" instance : ToFormat KnownType where @@ -538,7 +538,7 @@ partial def LMonoTy.resolveAliases (mty : LMonoTy) (T : TEnv Identifier) : (Opti | none => match mty with | .ftvar _ => (some mty, T) - | bitvec _ => (some mty, T) + | .bitvec _ => (some mty, T) | .tcons name mtys r => let (maybe_mtys, T) := LMonoTys.resolveAliases mtys T.context.aliases T match maybe_mtys with @@ -610,7 +610,7 @@ de-aliased. -/ def LMonoTy.knownInstance (ty : LMonoTy) (ks : KnownTypes) : Bool := match ty with - | .ftvar _ | bitvec _ => true + | .ftvar _ | .bitvec _ => true | .tcons name args _ => (ks.contains { name := name, arity := args.length }) && LMonoTys.knownInstances args ks diff --git a/Strata/DL/Lambda/LTy.lean b/Strata/DL/Lambda/LTy.lean index 87413d35..20ce9d22 100644 --- a/Strata/DL/Lambda/LTy.lean +++ b/Strata/DL/Lambda/LTy.lean @@ -50,7 +50,7 @@ inductive LMonoTy (TypeRestrictions: Type := LTyRestrict) : Type where abbrev LMonoTys := List LMonoTy @[match_pattern] -def bitvec (n: Nat) := LMonoTy.tcons "bitvec" [] (LTyRestrict.bitvecdata n) +def LMonoTy.bitvec (n: Nat) := LMonoTy.tcons "bitvec" [] (LTyRestrict.bitvecdata n) @[match_pattern] def LMonoTy.bool : LMonoTy := @@ -361,7 +361,7 @@ instance : ToString LMonoTy where private def formatLMonoTy (lmonoty : LMonoTy) : Format := match lmonoty with | .ftvar x => toString x - | bitvec n => f!"bv{n}" + | .bitvec n => f!"bv{n}" | .tcons name tys _ => if tys.isEmpty then f!"{name}" diff --git a/Strata/DL/Lambda/LTyUnify.lean b/Strata/DL/Lambda/LTyUnify.lean index 576f90a8..79987de8 100644 --- a/Strata/DL/Lambda/LTyUnify.lean +++ b/Strata/DL/Lambda/LTyUnify.lean @@ -979,14 +979,14 @@ def Constraint.unifyOne (c : Constraint) (S : SubstInfo) : have h_sub2 : Subst.freeVars_subset_prop [(orig_lty, LMonoTy.ftvar id)] newS S := by simp_all [Subst.freeVars_subset_prop_single_constraint_comm] .ok { newS := newS, goodSubset := by all_goals simp [h_sub1, h_sub2] } - | bitvec n1, bitvec n2 => + | .bitvec n1, .bitvec n2 => if _h7 : n1 == n2 then .ok { newS := SubstInfo.mk [] (by simp [SubstWF]), goodSubset := by grind } else .error f!"Cannot unify differently sized bitvector types {t1} and {t2}!" - | bitvec _, .tcons _ _ _ => + | .bitvec _, .tcons _ _ _ => .error f!"Cannot unify bv type {t1} and type constructor {t2}!" - | .tcons _ _ _, bitvec _ => + | .tcons _ _ _, .bitvec _ => .error f!"Cannot unify type constructor {t1} and bv type {t2}!" | .tcons name1 args1 r1, .tcons name2 args2 r2 => do if _h6 : name1 == name2 && args1.length == args2.length then diff --git a/Strata/Languages/Boogie/DDMTransform/Translate.lean b/Strata/Languages/Boogie/DDMTransform/Translate.lean index a32c3a78..c92a03e4 100644 --- a/Strata/Languages/Boogie/DDMTransform/Translate.lean +++ b/Strata/Languages/Boogie/DDMTransform/Translate.lean @@ -207,11 +207,11 @@ partial def translateLMonoTy (bindings : TransBindings) (arg : Arg) : let .type tp := arg | TransM.error s!"translateLMonoTy expected type {repr arg}" match tp with - | .ident q`Boogie.bv1 #[] => pure <| bitvec 1 - | .ident q`Boogie.bv8 #[] => pure <| bitvec 8 - | .ident q`Boogie.bv16 #[] => pure <| bitvec 16 - | .ident q`Boogie.bv32 #[] => pure <| bitvec 32 - | .ident q`Boogie.bv64 #[] => pure <| bitvec 64 + | .ident q`Boogie.bv1 #[] => pure <| .bitvec 1 + | .ident q`Boogie.bv8 #[] => pure <| .bitvec 8 + | .ident q`Boogie.bv16 #[] => pure <| .bitvec 16 + | .ident q`Boogie.bv32 #[] => pure <| .bitvec 32 + | .ident q`Boogie.bv64 #[] => pure <| .bitvec 64 | .ident i argst => let argst' ← translateLMonoTys bindings (argst.map Arg.type) pure <| (.tcons i.name argst'.toList.reverse) @@ -402,7 +402,7 @@ def dealiasTypeArg (p : Program) (a : Arg) : Arg := def isArithTy : LMonoTy → Bool | .int => true | .real => true -| bitvec _ => true +| .bitvec _ => true | _ => false def translateFn (ty? : Option LMonoTy) (q : QualifiedIdent) : TransM Boogie.Expression.Expr := diff --git a/Strata/Languages/Boogie/SMTEncoder.lean b/Strata/Languages/Boogie/SMTEncoder.lean index 8a4742b3..887d9a4b 100644 --- a/Strata/Languages/Boogie/SMTEncoder.lean +++ b/Strata/Languages/Boogie/SMTEncoder.lean @@ -87,7 +87,7 @@ mutual def LMonoTy.toSMTType (ty : LMonoTy) (ctx : SMT.Context) : Except Format (TermType × SMT.Context) := do match ty with - | bitvec n => .ok (TermType.bitvec n, ctx) + | .bitvec n => .ok (TermType.bitvec n, ctx) | .tcons "bool" [] _ => .ok (.bool, ctx) | .tcons "int" [] _ => .ok (.int, ctx) | .tcons "real" [] _ => .ok (.real, ctx) @@ -141,7 +141,7 @@ partial def toSMTTerm (E : Env) (bvs : BoundVars) (e : LExpr LMonoTy BoogieIdent | none => .error f!"Unexpected real constant {e}" | some r => .ok ((Term.real r), ctx) - | bitvec n => + | .bitvec n => match e.denoteBitVec n with | none => .error f!"Unexpected bv constant {e}" From 61787d44138aaa01383f973bebc4281168b11575 Mon Sep 17 00:00:00 2001 From: Josh Cohen Date: Mon, 22 Sep 2025 17:15:09 -0400 Subject: [PATCH 04/19] Initial version of bounded int expression translation --- Strata/DL/Bounded/BExpr.lean | 246 +++++++++++++++++++++++++++++++++++ Strata/DL/Bounded/BTy.lean | 61 +++++++++ 2 files changed, 307 insertions(+) create mode 100644 Strata/DL/Bounded/BExpr.lean create mode 100644 Strata/DL/Bounded/BTy.lean diff --git a/Strata/DL/Bounded/BExpr.lean b/Strata/DL/Bounded/BExpr.lean new file mode 100644 index 00000000..f7cc5776 --- /dev/null +++ b/Strata/DL/Bounded/BExpr.lean @@ -0,0 +1,246 @@ +/- + Copyright Strata Contributors + + SPDX-License-Identifier: Apache-2.0 OR MIT +-/ + +import Strata.DL.Lambda.LTy +-- import Strata.DL.Bounded.BTy +import Strata.DL.Lambda.LExprT +import Strata.DL.Lambda.IntBoolFactory + +/-! ## Lambda Expressions with Bounded Ints + +Here, we parameterize LExprs with bounded int types. These augmented expressions +can be desugared into regular Lambda expressions by erasing the int bounds. To make +the bounds meaningful, we also generate constraints. +All bounds must be explicitly stated; they will not be inferred. + +-/ + +/-! +In general, each input bound gives rise to an assumption in the expression body +as well as an obligation for any arguments, whereas an output bound is the opposite. +-/ + +namespace Bounded +open Std (ToFormat Format format) +open Lambda + +/- + +There are two parts to desugaring BLambda to Lambda: translating the terms +by (almost) simply removing the integer bounds, and generating the +corresponding well-formedness conditions. Well-formedness is more general than +just bounded ints, though that is all we have for now. +-/ + +def isBounded (t: LMonoTy) : Option BoundExpr := + match t with + | LMonoTy.bounded b => .some b + | _ => .none + +def removeBound (t: LMonoTy) : LMonoTy := + match t with + | LMonoTy.bounded _ => LMonoTy.int + | .tcons name args m => .tcons name (List.map removeBound args) m + | _ => t + + +-- generate a single constraint with a body +-- we rely on Lambda's existing substitution for this, simply constructing +-- an expression with a single free variable (called "x", but it will always +-- be substiuted so the name is irrelevant) +-- ugh, free var is Identifier, do substitution directly +--not great, unsafe maybe? +def BoundValToLExprT (b: BoundVal) (e: LExprT Identifier) : LExprT Identifier := + match b with + | .bvar => e + | .bconst val => .const (val.repr) LMonoTy.int + +-- a hack for now (from Factory.lean) +def LFunc.opExprT (f: LFunc Identifier) : LExprT Identifier := + let input_tys := f.inputs.values + let output_tys := Lambda.LMonoTy.destructArrow f.output + let ty := match input_tys with + | [] => f.output + | ity :: irest => Lambda.LMonoTy.mkArrow ity (irest ++ output_tys) + .op f.name ty + +--This is VERY messy,need to think about where this should take place +--MUCH nicer to do on untyped terms but need type for function application +-- Invariant, e must have type int +def BoundExprToLExprT [Coe String Identifier] (b: BoundExpr) (e: LExprT Identifier) : +LExprT Identifier := + match b with + | .beq b1 b2 => + .eq (BoundValToLExprT b1 e) (BoundValToLExprT b2 e) LMonoTy.bool + | .blt b1 b2 => + .app (.app (LFunc.opExprT intLtFunc) (BoundValToLExprT b1 e) (LMonoTy.arrow LMonoTy.int LMonoTy.bool)) (BoundValToLExprT b2 e) LMonoTy.bool + | .ble b1 b2 => + .app (.app (LFunc.opExprT intLeFunc) (BoundValToLExprT b1 e) (LMonoTy.arrow LMonoTy.int LMonoTy.bool)) (BoundValToLExprT b2 e) LMonoTy.bool + | .band e1 e2 => + .app (.app (LFunc.opExprT boolAndFunc) (BoundExprToLExprT e1 e) + (LMonoTy.arrow LMonoTy.bool LMonoTy.bool)) (BoundExprToLExprT e2 e) LMonoTy.bool + | .bnot e1 => + .app (LFunc.opExprT boolNotFunc) (BoundExprToLExprT e1 e) LMonoTy.bool + +-- e must have type bool +def addAssumptions [Coe String Identifier] (l: List (LExprT Identifier)) (e: LExprT Identifier) : LExprT Identifier := + List.foldr (fun x acc => .app (.app (LFunc.opExprT boolImpliesFunc) x LMonoTy.bool) acc LMonoTy.bool) e l + +def isBool (t: LMonoTy) : Bool := + match t with + | .bool => True + | _ => False + +-- Only deal with expressions of form bvar, eq, app (output of BoundExprToLExprT on bvar) +def increaseBVar (e: LExprT Identifier) : LExprT Identifier := + match e with + | .bvar b ty => .bvar (b+1) ty + | .eq e1 e2 ty => .eq (increaseBVar e1) (increaseBVar e2) ty + | .app e1 e2 ty => .app (increaseBVar e1) (increaseBVar e2) ty + | _ => e + + +def increaseBVars (l: List (LExprT Identifier)) : List (LExprT Identifier) := + List.map increaseBVar l + +/-- +Translate expression to one without bounded ints. +This is a non-trivial translation, as we want to preserve the semantics of the resulting term. +To see why this is tricky, consider the following examples: +1. ∀ (x: Nat), 0 <= x (should result in something semantically equivalent to true) +2. λ(x: Nat), if 0 <= x then 1 else 0 (should result in function that always evaluates to 1) +3. λ(x: int), if foo x >= 0 then 1 else 0 (supposing foo : int -> Nat, should also result in function that always evaluates to 1) +4. (x: Nat) + (y: Nat) >= 0 must be true + +To handle the first case, any bound variables induce constraints that must be +inserted in the resulting formula. The second case shows that adding the constraints is complicated, because the resulting term is not necessarily bool-valued. The third example shows that variables are not enough: any term of bounded type must result in an assumption somehow. The fourth shows that terms may not be closed; therefore, we must collect free variable assumptions as well. + +To deal with all of this, our translation function produces two outputs: the translated term and a set of constraints. Broadly, new constraints are added whenever we see a bounded bound variable or an expression of bounded type. We must propagate them back up to any boolean-valued expressions. There is a subtlety resulting from bound variables, as we must make sure the constraints are scoped appropriately (e.g. ∀ (y: Nat). λ(x: Nat). x + y >= 0, which becomes +∀ Nat. λ Nat. #0 + #1 >= 0 must result in ∀ #0 >= 0 -> λ Nat. #0 >= 0 -> #0 + #1 >= 0). + +Also note that assumptions percolate in two directions: A bound variable (possibly) induces a new assumption somewhere inside a subterm, while a (e.g.) function application induces a new assumption in a parent call. Thus, we need both input (the first kind) and output (the second kind) assumption lists. + +Note that we do NOT have to worry about the fact that the new functions have a larger domain when types are erased. We generate separate well-formedness checks (below) to ensure that any call of the function occurs on an argument satisfying the constraints. + +Invariant: assumptions must not have bounded types (TODO enforce this), same for inputs +Invariant (I think): All assumptions are of form: b(bvar #x) +-/ +def translateBounded [Coe String Identifier] (e: LExprT Identifier) (assumptions: List (LExprT Identifier)) : LExprT Identifier × List (LExprT Identifier) := + match e with + | .const c ty => (.const c (removeBound ty), []) + -- an op adds an assumption if it has bounded type (otherwise, its application may add an assumption further up the chain) + | .op o ty => let e1 := .op o (removeBound ty); + (if isBool ty then addAssumptions assumptions e1 else e1, ((isBounded ty).bind (fun b => BoundExprToLExprT b e1)).toList) + -- bvars should already have been handled if needed + | .bvar b ty => (.bvar b (removeBound ty), []) + -- fvars should be given back to the caller as a potential assumption + | .fvar f ty => let e1 := .fvar f (removeBound ty); (e1, ((isBounded ty).bind (fun b => BoundExprToLExprT b e1)).toList) + -- application: if the entire application has bool type, add assumptions (and remove from recursive call) + -- need to collect assumptions from both recursive calls and add if needed + | .app e1 e2 ty => + if isBool ty then + let e1' := translateBounded e1 []; + let e2' := translateBounded e2 []; + -- safe to add all assumptions, including recursive ones, nothing to propagate up + (addAssumptions (assumptions ++ e1'.2 ++ e2'.2) (.app e1'.1 e2'.1 (removeBound ty)), []) + else + --just recursive, may need assumptions in both but do not need assumptions from 1 in another + let e1' := translateBounded e1 assumptions; + let e2' := translateBounded e2 assumptions; + (.app e1'.1 e2'.1 (removeBound ty), e1'.2 ++ e2'.2) + -- abstraction: probably the trickiest one. + -- 1. if the body has bool type, easier, add assumptions (including new one) and translate + -- 2. Otherwise, need to add assumption and increase bvars of all in "assumptions" list (as they are passing through binder) + | .abs e ty => + --get new assumption (if variable is bounded) + let newAssumption := + match ty with + | .arrow (.bounded b) _ => + [BoundExprToLExprT b (.bvar 0 .int)] + | _ => []; + -- case 1: body has bool type + match ty with + | .arrow _ .bool => + let e' := translateBounded e []; + (.abs (addAssumptions (newAssumption ++ (increaseBVars assumptions) ++ e'.2) e'.1) (removeBound ty), []) + -- otherwise, propagate the assumptions, adding a new one + | _ => let e' := translateBounded e (newAssumption ++ (increaseBVars assumptions)); + (.abs e'.1 (removeBound ty), e'.2) + -- quantifiers are complex but a bit simpler because we know they are boolean-valued + | .quant .all ty tr e => + let newAssumption := + match ty with + | .bounded b => + [BoundExprToLExprT b (.bvar 0 .int)] + | _ => []; + let e' := translateBounded e []; + let tr' := translateBounded tr []; --TODO: need "clean" one here + (.quant .all (removeBound ty) tr'.1 (addAssumptions (newAssumption ++ (increaseBVars assumptions) ++ e'.2) e'.1), []) + -- only difference: need exists x, b(x) /\ (assumptions -> e) + | .quant .exist ty tr e => + let newAssumption := + match ty with + | .bounded b => + some (BoundExprToLExprT b (.bvar 0 .int)) + | _ => none; + let e' := translateBounded e []; + let tr' := translateBounded tr []; --TODO: need "clean" one here + let add_and x : LExprT Identifier := match newAssumption with + | .some f => (.app (.app (LFunc.opExprT boolAndFunc) f (LMonoTy.arrow LMonoTy.bool LMonoTy.bool)) x LMonoTy.bool) + | .none => x; + (.quant .exist (removeBound ty) tr'.1 (add_and (addAssumptions ((increaseBVars assumptions) ++ e'.2) e'.1)), []) + -- if-then-else is recursive, but we can add the assumptions to the condition. Likewise, we can test if the result is boolean-valued to see if we add to result + | .ite c t f .bool => + let c' := translateBounded c []; + let t' := translateBounded t []; + let f' := translateBounded f []; + (.ite (addAssumptions (assumptions ++ c'.2) c'.1) (addAssumptions (assumptions ++ t'.2) t'.1) (addAssumptions (assumptions ++ f'.2) f'.1) .bool, []) + | .ite c t f ty => + let c' := translateBounded c []; + let t' := translateBounded t assumptions; + let f' := translateBounded f assumptions; + (.ite (addAssumptions (assumptions ++ c'.2) c'.1) t'.1 f'.1 ty, t'.2 ++ f'.2) + -- eq is bool so easy + | .eq e1 e2 ty => --shouldnt' need to check bool + let e1' := translateBounded e1 []; + let e2' := translateBounded e2 []; + (addAssumptions (assumptions ++ e1'.2 ++ e2'.2) (.eq e1'.1 e2'.1 ty), []) + -- metadata just recursive + | .mdata m e => + let e' := translateBounded e assumptions; + (.mdata m e'.1, e'.2) + +-- TODO: test this with all examples from above + + + + + +/-- +Generate the constraints for a bounded lambda expression. They are: +1. For any int constant c : {x | b (x)}, generate constraint b(c) +2. For .app e1 e2, if e1 has type {x : b(x)} -> t, generate constraint b(e2) +3. For fvar +We also get information (NOTE: BOTH for translation and constraints): +3. For forall {x | b(x)} tr e have forall x, b(x) -> e +4. For exists {x | b(x)} tr e, have exists x, b(x) /\ e +For any free variables in the term, we add the assumption from their bound +(for example, suppose we have forall (x: Nat), x * (y: Nat) >= 0, + this results in y >= 0 -> forall (x: int), x * y >= 0) +-/ +def bounded_constraints [Coe String Identifier] +(e: Lambda.LExpr BMonoTy Identifier) : List (Lambda.LExpr Lambda.LMonoTy Identifier) := + match e with + | .const x (some (.boundint b)) => + -- Constant must be int + match x.toInt? with + | some i => [bound_to_expr b (.const x LMonoTy.int )] + | none => [] --this will be caught during typechecking + | .app e1 e2 => + --problem: need to be at given type but dont have types yet + + +end Bounded diff --git a/Strata/DL/Bounded/BTy.lean b/Strata/DL/Bounded/BTy.lean new file mode 100644 index 00000000..8a5575c1 --- /dev/null +++ b/Strata/DL/Bounded/BTy.lean @@ -0,0 +1,61 @@ +/- + Copyright Strata Contributors + + SPDX-License-Identifier: Apache-2.0 OR MIT +-/ + +import Strata.DL.Lambda.LTy + +/-! ## Lambda Types with Bounded Ints +-/ + +-- The grammar allows expressions like a <= b /\ b < c, where a, b, and c +-- are integer constants or a (single) variable. This lets us express e.g. +-- {x | 0 <= x < 2^32} + +namespace Bounded +open Std (ToFormat Format format) + + +abbrev TyIdentifier := String + +inductive BVal : Type where + /-- The bounded variable -/ + | bvar + /-- Integer constants -/ + | bconst (val: Int) +deriving Inhabited, Repr, DecidableEq + +inductive BExpr : Type where + /-- b1 = b2 -/ + | beq (b1 b2: BVal) + /-- b1 < b2 -/ + | blt (e1 e2: BVal) + /-- b1 <= b2 -/ + | ble (e1 e2: BVal) + /-- e1 /\ e2 -/ + | band (e1 e2: BExpr) + /-- not e -/ + | bnot (e: BExpr) +deriving Inhabited, Repr, DecidableEq + +inductive BMonoTy : Type where + /-- Type variable. -/ + | ftvar (name : TyIdentifier) + /-- Type constructor. -/ + | tcons (name : String) (args : List BMonoTy) + /-- Special support for bitvector types of every size. -/ + | bitvec (size : Nat) + /-- Added to LMonoTy: bounded ints -/ + | boundint (b: BExpr) +deriving Inhabited, Repr + +-- When desugaring to Lambda types, all bounded ints are erased +def BMonoTy_to_LMonoTy (b: BMonoTy) : Lambda.LMonoTy := + match b with + | .ftvar name => Lambda.LMonoTy.ftvar name + | .tcons name args => Lambda.LMonoTy.tcons name (List.map BMonoTy_to_LMonoTy args) + | .bitvec size => Lambda.LMonoTy.bitvec size + | .boundint _ => Lambda.LMonoTy.int + +end Bounded From 989ec244d5eccde772fc3a8b9d60ca0e95e8bfe9 Mon Sep 17 00:00:00 2001 From: Josh Cohen Date: Wed, 24 Sep 2025 13:17:26 -0400 Subject: [PATCH 05/19] Add bounded types to LTy, test cases for translateBounded --- Strata/DL/Bounded/BExpr.lean | 614 ++++++++++++++++++++++++++++++++++- Strata/DL/Lambda/LTy.lean | 31 ++ 2 files changed, 643 insertions(+), 2 deletions(-) diff --git a/Strata/DL/Bounded/BExpr.lean b/Strata/DL/Bounded/BExpr.lean index f7cc5776..53a57006 100644 --- a/Strata/DL/Bounded/BExpr.lean +++ b/Strata/DL/Bounded/BExpr.lean @@ -147,10 +147,12 @@ def translateBounded [Coe String Identifier] (e: LExprT Identifier) (assumptions -- safe to add all assumptions, including recursive ones, nothing to propagate up (addAssumptions (assumptions ++ e1'.2 ++ e2'.2) (.app e1'.1 e2'.1 (removeBound ty)), []) else + -- if the application has bounded type, need to add assumption further up the chain --just recursive, may need assumptions in both but do not need assumptions from 1 in another let e1' := translateBounded e1 assumptions; let e2' := translateBounded e2 assumptions; - (.app e1'.1 e2'.1 (removeBound ty), e1'.2 ++ e2'.2) + let e' := .app e1'.1 e2'.1 (removeBound ty); + (e', e1'.2 ++ e2'.2 ++ ((isBounded ty).bind (fun b => BoundExprToLExprT b e')).toList) -- abstraction: probably the trickiest one. -- 1. if the body has bool type, easier, add assumptions (including new one) and translate -- 2. Otherwise, need to add assumption and increase bvars of all in "assumptions" list (as they are passing through binder) @@ -213,10 +215,618 @@ def translateBounded [Coe String Identifier] (e: LExprT Identifier) (assumptions let e' := translateBounded e assumptions; (.mdata m e'.1, e'.2) --- TODO: test this with all examples from above +def translateBounded' [Coe String Identifier] (e: LExprT Identifier) : LExprT Identifier := + (translateBounded e []).1 + +-- Tests + +-- NOTE: with a semantics for LExpr/LExprT, we could prove the equivalences mentioned above + +def natTy : LMonoTy := LMonoTy.bounded (.ble (.bconst 0) .bvar) +def leOp (e1 e2: LExprT String) : LExprT String := .app (.app (LFunc.opExprT intLeFunc) e1 (.arrow .int .bool)) e2 .bool + +def geOp (e1 e2: LExprT String) : LExprT String := .app (.app (LFunc.opExprT intGeFunc) e1 (.arrow .int .bool)) e2 .bool + +def addOp (e1 e2: LExprT String) : LExprT String := .app (.app (LFunc.opExprT intAddFunc) e1 (.arrow .int .int)) e2 .int + +def notOp (e: LExprT String) : LExprT String := .app (LFunc.opExprT boolNotFunc) e .bool + +-- 1. ∀ (x: Nat), 0 <= x (quantified assumption) + +def test1 := (@LExprT.quant String .all natTy (.bvar 0 natTy) (leOp (.const "0" .int) (.bvar 0 .int))) + +#eval translateBounded' test1 + +--easier to read +#eval (LExpr.eraseTypes (LExprT.toLExpr (translateBounded' test1))) + +/-- info: Lambda.LExpr.quant + (Lambda.QuantifierKind.all) + (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.bvar 0) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Bool.Implies" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Int.Le" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) + (Lambda.LExpr.bvar 0))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Int.Le" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) + (Lambda.LExpr.bvar 0))) +-/ +#guard_msgs in +#eval (LExprT.toLExpr (translateBounded' test1)) + +-- 2. λ(x: Nat), if 0 <= x then 1 else 0 (assumption inside ite) + +def test2 : LExprT String := .abs (.ite (leOp (.const "0" .int) (.bvar 0 .int)) (.const "1" .int) (.const "2" .int) .int) (.arrow natTy .int) + +#eval translateBounded' test2 + +#eval (LExpr.eraseTypes (LExprT.toLExpr (translateBounded' test2))) + +/-- info: Lambda.LExpr.abs + (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.ite + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Bool.Implies" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Int.Le" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) + (Lambda.LExpr.bvar 0))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Int.Le" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) + (Lambda.LExpr.bvar 0))) + (Lambda.LExpr.const "1" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.const "2" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) -/ +#guard_msgs in +#eval (LExprT.toLExpr (translateBounded' test2)) + +-- 3. λ(x: int), if foo x >= 0 then 1 else 0 (for foo: int -> Nat) (propagate op/application information) + +def test3 : LExprT String := .abs (.ite (geOp (.app (.op "foo" (.arrow .int natTy)) (.bvar 0 .int) natTy) (.const "0" .int)) (.const "1" .int) (.const "0" .int) .int) (.arrow .int .int) + +#eval translateBounded' test3 + +#eval (LExpr.eraseTypes (LExprT.toLExpr (translateBounded' test3))) + +/-- +info: Lambda.LExpr.abs + (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.ite + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Bool.Implies" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Int.Le" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) + (Lambda.LExpr.app + (Lambda.LExpr.op + "foo" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.bvar 0)))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Int.Ge" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.app + (Lambda.LExpr.op + "foo" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.bvar 0))) + (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata)))))) + (Lambda.LExpr.const "1" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) + -/ +#guard_msgs in +#eval (LExprT.toLExpr (translateBounded' test3)) + +-- 4. (x: Nat) + (y: Nat) >= 0 (free variable bounds) + +def test4 : LExprT String := geOp (addOp (.fvar "x" natTy) (.fvar "y" natTy)) (.const "0" .int) + +#eval translateBounded' test4 + +#eval (LExpr.eraseTypes (LExprT.toLExpr (translateBounded' test4))) + +/-- +info: Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Bool.Implies" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Int.Le" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) + (Lambda.LExpr.fvar "x" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata)))))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Bool.Implies" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Int.Le" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) + (Lambda.LExpr.fvar "y" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata)))))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Int.Ge" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Int.Add" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.fvar "x" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) + (Lambda.LExpr.fvar "y" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata)))))) + (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata)))))) -/ +#guard_msgs in +#eval (LExprT.toLExpr (translateBounded' test4)) + +-- 5. λ (x: Nat). λ (y: Nat). x + y >= 0 (multiple bound vars) + +def test5 : LExprT String := .abs (.abs (geOp (addOp (.bvar 1 .int) (.bvar 0 .int)) (.const "0" .int)) (.arrow natTy .bool)) (.arrow natTy (.arrow natTy .bool)) + +#eval translateBounded' test5 + +#eval (LExpr.eraseTypes (LExprT.toLExpr (translateBounded' test5))) + +/-- +info: Lambda.LExpr.abs + (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.abs + (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Bool.Implies" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Int.Le" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) + (Lambda.LExpr.bvar 0))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Bool.Implies" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Int.Le" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) + (Lambda.LExpr.bvar 1))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Int.Ge" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Int.Add" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.bvar 1)) + (Lambda.LExpr.bvar 0))) + (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata)))))))) +-/ +#guard_msgs in +#eval (LExprT.toLExpr (translateBounded' test5)) + +-- 6. λ (x: Nat), if foo then λ (y: Nat). not (x = -1) else λ (y: Nat). x + y >= 0 (propagate inside branches of if-then-else) +def test6 : LExprT String := .abs (.ite (.op "foo" .bool) (.abs (notOp (.eq (.bvar 1 .int) (.const "-1" .int) .bool)) (.arrow natTy .bool)) (.abs (geOp (addOp (.bvar 1 .int) (.bvar 0 .int)) (.const "0" .int)) (.arrow natTy .bool)) (.arrow natTy .bool)) (.arrow natTy (.arrow natTy .bool)) +#eval translateBounded' test6 + +#eval (LExpr.eraseTypes (LExprT.toLExpr (translateBounded' test6))) + +/-- +info: Lambda.LExpr.abs + (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.ite + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Bool.Implies" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Int.Le" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) + (Lambda.LExpr.bvar 0))) + (Lambda.LExpr.op "foo" (some (Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata))))) + (Lambda.LExpr.abs + (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Bool.Implies" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Int.Le" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) + (Lambda.LExpr.bvar 0))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Bool.Implies" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Int.Le" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) + (Lambda.LExpr.bvar 1))) + (Lambda.LExpr.app + (Lambda.LExpr.op + "Bool.Not" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.eq + (Lambda.LExpr.bvar 1) + (Lambda.LExpr.const "-1" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))))))) + (Lambda.LExpr.abs + (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Bool.Implies" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Int.Le" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) + (Lambda.LExpr.bvar 0))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Bool.Implies" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Int.Le" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) + (Lambda.LExpr.bvar 1))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Int.Ge" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op + "Int.Add" + (some (Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons + "arrow" + [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), + Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)] + (Lambda.LTyRestrict.nodata)))) + (Lambda.LExpr.bvar 1)) + (Lambda.LExpr.bvar 0))) + (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))))))) + -/ +#guard_msgs in +#eval (LExprT.toLExpr (translateBounded' test6)) /-- diff --git a/Strata/DL/Lambda/LTy.lean b/Strata/DL/Lambda/LTy.lean index 20ce9d22..89e651f9 100644 --- a/Strata/DL/Lambda/LTy.lean +++ b/Strata/DL/Lambda/LTy.lean @@ -25,6 +25,30 @@ abbrev TyIdentifier := String instance : Coe String TyIdentifier where coe := id +/-- +Bounded integer expressions: the grammar allows expressions like +a <= b /\ b < c, where a, b, and c are integer constants or a (single) variable. This lets us express types like {x | 0 <= x < 2^32} +-/ +inductive BoundVal : Type where + /-- The bounded variable -/ + | bvar + /-- Integer constants -/ + | bconst (val: Int) +deriving Inhabited, Repr, DecidableEq + +inductive BoundExpr : Type where + /-- b1 = b2 -/ + | beq (b1 b2: BoundVal) + /-- b1 < b2 -/ + | blt (e1 e2: BoundVal) + /-- b1 <= b2 -/ + | ble (e1 e2: BoundVal) + /-- e1 /\ e2 -/ + | band (e1 e2: BoundExpr) + /-- not e -/ + | bnot (e: BoundExpr) +deriving Inhabited, Repr, DecidableEq + /-- An additional parameter that restricts the values of a type or attaches additional metadata to a type. @@ -33,6 +57,8 @@ For now, only bitvectors are supported. inductive LTyRestrict : Type where /-- Special support for bitvector types of every size. -/ | bitvecdata (size: Nat) + /-- Bounded integers -/ + | bounddata (b: BoundExpr) | nodata deriving Inhabited, Repr, DecidableEq @@ -64,6 +90,10 @@ def LMonoTy.int : LMonoTy := def LMonoTy.real : LMonoTy := .tcons "real" [] +@[match_pattern] +def LMonoTy.bounded (b: BoundExpr) : LMonoTy := + .tcons "int" [] (LTyRestrict.bounddata b) + @[match_pattern] def LMonoTy.bv1 : LMonoTy := bitvec 1 @@ -271,6 +301,7 @@ Get all type constructors in monotype `mty`. def LMonoTy.getTyConstructors (mty : LMonoTy) : List LMonoTy := match mty with | .ftvar _ => [] + | bitvec _ => [] | .tcons name mtys _ => let typeargs := List.replicate mtys.length "_dummy" let args := typeargs.mapIdx (fun i elem => LMonoTy.ftvar (elem ++ toString i)) From 6aba2ebf5ffcd7e7ec4ea4fa12b3f0fc00b673e0 Mon Sep 17 00:00:00 2001 From: Josh Cohen Date: Thu, 25 Sep 2025 17:12:26 -0400 Subject: [PATCH 06/19] Add lots of test cases, clean up code Test cases for wf generation Fix a few bugs, refactor and improve comments in code Some AI-generated tests for more bound cases --- Strata/DL/Bounded/BExpr.lean | 431 +++++++++++++++++++++++++++++------ 1 file changed, 361 insertions(+), 70 deletions(-) diff --git a/Strata/DL/Bounded/BExpr.lean b/Strata/DL/Bounded/BExpr.lean index 53a57006..7e658ad7 100644 --- a/Strata/DL/Bounded/BExpr.lean +++ b/Strata/DL/Bounded/BExpr.lean @@ -106,6 +106,39 @@ def increaseBVar (e: LExprT Identifier) : LExprT Identifier := def increaseBVars (l: List (LExprT Identifier)) : List (LExprT Identifier) := List.map increaseBVar l +/- Do translation and wf generation in 1 go - need same assumptions, expesive to compute both-/ +structure translationRes Identifier where +(translate : LExprT Identifier) +(wfCond : List (LExprT Identifier)) +(assume : List (LExprT Identifier)) + +-- more aux functions + +-- evaluate expression at bound if type is bounded +def boundExprIfType [Coe String Identifier] (ty: LMonoTy) (e: LExprT Identifier) : List (LExprT Identifier) := + ((isBounded ty).map (fun b => BoundExprToLExprT b e)).toList + +-- Generate WF condition for calling e1 with argument e2 if bounded type expected +def wfCallCondition [Coe String Identifier] (assume : List (LExprT Identifier)) (e1 e2: LExprT Identifier) := + match LExprT.toLMonoTy e1 with + | .arrow (.bounded b) _ => + -- check that translated e2 satisfies precondition under assumptions + [addAssumptions assume (BoundExprToLExprT b e2)] + | _ => [] + +--TODO: change to avoid quantifier +-- NOTE: l1 includes assumptions, l2 does not +def addBoundedWf [Coe String Identifier] (assume: List (LExprT Identifier)) (l1 l2: List (LExprT Identifier)) : List (LExprT Identifier) := + List.map (fun e => .quant .all .int (.bvar 0 .int) (addAssumptions assume e)) l1 ++ List.map (.quant .all .int (.bvar 0 .int)) l2 + +def addBoundedWfAssume [Coe String Identifier] (assume: List (LExprT Identifier)) (l: List (LExprT Identifier)) := + addBoundedWf assume l [] + +-- for bounded-valued terms, need to generate condition that body satisfies bound +-- def wfBoundedBody [Coe String Identifier] (ty: LMonoTy) (e: LExprT Identifier) := +-- ((isBounded ty).map (fun b => BoundExprToLExprT b e)).toList + + /-- Translate expression to one without bounded ints. This is a non-trivial translation, as we want to preserve the semantics of the resulting term. @@ -128,50 +161,74 @@ Note that we do NOT have to worry about the fact that the new functions have a l Invariant: assumptions must not have bounded types (TODO enforce this), same for inputs Invariant (I think): All assumptions are of form: b(bvar #x) -/ -def translateBounded [Coe String Identifier] (e: LExprT Identifier) (assumptions: List (LExprT Identifier)) : LExprT Identifier × List (LExprT Identifier) := +def translateBounded [Coe String Identifier] (e: LExprT Identifier) (assume: List (LExprT Identifier)) : translationRes Identifier := match e with - | .const c ty => (.const c (removeBound ty), []) - -- an op adds an assumption if it has bounded type (otherwise, its application may add an assumption further up the chain) - | .op o ty => let e1 := .op o (removeBound ty); - (if isBool ty then addAssumptions assumptions e1 else e1, ((isBounded ty).bind (fun b => BoundExprToLExprT b e1)).toList) - -- bvars should already have been handled if needed - | .bvar b ty => (.bvar b (removeBound ty), []) - -- fvars should be given back to the caller as a potential assumption - | .fvar f ty => let e1 := .fvar f (removeBound ty); (e1, ((isBounded ty).bind (fun b => BoundExprToLExprT b e1)).toList) - -- application: if the entire application has bool type, add assumptions (and remove from recursive call) - -- need to collect assumptions from both recursive calls and add if needed + -- constants do not need assumptions; they produce a wf goal if bounded + | .const c ty => + let res := .const c (removeBound ty); + ⟨res, boundExprIfType ty res, []⟩ + -- an op adds a bottom-up assumption if it has bounded type; its wf is assumed + | .op o ty => + let res := .op o (removeBound ty); + ⟨if isBool ty then addAssumptions assume res else res, [], boundExprIfType ty res⟩ + -- bvars are handled when bound + | .bvar b ty => ⟨ .bvar b (removeBound ty), [], [] ⟩ + -- fvars generate bottom-up assumptions if bounded + | .fvar f ty => + let res := .fvar f (removeBound ty); + ⟨res, [], boundExprIfType ty res⟩ + /- + Application has several cases: + 1. If the entire application has boolean type, assumptions can be added + 2. Otherwise, if the application has bounded type, it produces a bottom-up assumption. There is a subtle case. If the function has type (t -> int), then we must generate a wf condition. Otherwise, the bound will be assumed (for external operators) or checked (for abstraction/if-then-else/etc) + 3. In either case, we produce a wf condition if the argument should have bounded type + -/ | .app e1 e2 ty => if isBool ty then - let e1' := translateBounded e1 []; - let e2' := translateBounded e2 []; - -- safe to add all assumptions, including recursive ones, nothing to propagate up - (addAssumptions (assumptions ++ e1'.2 ++ e2'.2) (.app e1'.1 e2'.1 (removeBound ty)), []) + let e1' := translateBounded e1 []; + let e2' := translateBounded e2 []; + let all_assumes := assume ++ e1'.assume ++ e2'.assume; + let res := addAssumptions all_assumes (.app e1'.translate e2'.translate (removeBound ty)); + ⟨res, (wfCallCondition (assume ++ e2'.assume) e1 e2'.translate) ++ e1'.wfCond ++ e2'.wfCond, []⟩ else - -- if the application has bounded type, need to add assumption further up the chain - --just recursive, may need assumptions in both but do not need assumptions from 1 in another - let e1' := translateBounded e1 assumptions; - let e2' := translateBounded e2 assumptions; - let e' := .app e1'.1 e2'.1 (removeBound ty); - (e', e1'.2 ++ e2'.2 ++ ((isBounded ty).bind (fun b => BoundExprToLExprT b e')).toList) - -- abstraction: probably the trickiest one. - -- 1. if the body has bool type, easier, add assumptions (including new one) and translate - -- 2. Otherwise, need to add assumption and increase bvars of all in "assumptions" list (as they are passing through binder) + let e1' := translateBounded e1 assume; + let e2' := translateBounded e2 assume; + let e' := .app e1'.translate e2'.translate (removeBound ty); + let extraWf := + match LExprT.toLMonoTy e1, ty with + | .arrow _ .int, .bounded _ => + boundExprIfType ty e' + | _, _ => []; + ⟨e', (wfCallCondition (assume ++ e2'.assume) e1 e2'.translate) ++ extraWf ++ e1'.wfCond ++ e2'.wfCond, e1'.assume ++ e2'.assume ++ boundExprIfType ty e'⟩ + /- + Abstraction is the most complex case: + 1. If the argument is bounded, add as top-down assumption + 2. If the body has type bool, add assumptions and translate + 3. Otherwise, add assumptions and increase bvars of all in "assume" list, as they are passing through a binder + 3. WF: prove body satisfies bound if needed with same assumptions (but without new binder) + -/ | .abs e ty => - --get new assumption (if variable is bounded) let newAssumption := match ty with | .arrow (.bounded b) _ => [BoundExprToLExprT b (.bvar 0 .int)] | _ => []; - -- case 1: body has bool type match ty with | .arrow _ .bool => let e' := translateBounded e []; - (.abs (addAssumptions (newAssumption ++ (increaseBVars assumptions) ++ e'.2) e'.1) (removeBound ty), []) - -- otherwise, propagate the assumptions, adding a new one - | _ => let e' := translateBounded e (newAssumption ++ (increaseBVars assumptions)); - (.abs e'.1 (removeBound ty), e'.2) - -- quantifiers are complex but a bit simpler because we know they are boolean-valued + let all_assume := newAssumption ++ (increaseBVars assume) ++ e'.assume; + -- TODO: change + ⟨.abs (addAssumptions all_assume e'.translate) (removeBound ty), addBoundedWfAssume all_assume e'.wfCond, []⟩ + | .arrow _ ty1 => + let all_assume := newAssumption ++ (increaseBVars assume); + let e' := translateBounded e all_assume; + let e'' := .abs e'.translate (removeBound ty); + -- Note: don't add assumptions to e'.wfCond, already included + ⟨e'', addBoundedWf all_assume (boundExprIfType ty1 e'') e'.wfCond, e'.assume⟩ + | _ => ⟨.const "0" .int, [], []⟩ -- error case + /- + Quantifiers are simpler because they are boolean-valued. ∀ (x : nat). e adds an assumption x >= 0 -> ..., while ∃ (x: nat). e results in x >= 0 ∧ .. + -/ | .quant .all ty tr e => let newAssumption := match ty with @@ -180,8 +237,9 @@ def translateBounded [Coe String Identifier] (e: LExprT Identifier) (assumptions | _ => []; let e' := translateBounded e []; let tr' := translateBounded tr []; --TODO: need "clean" one here - (.quant .all (removeBound ty) tr'.1 (addAssumptions (newAssumption ++ (increaseBVars assumptions) ++ e'.2) e'.1), []) - -- only difference: need exists x, b(x) /\ (assumptions -> e) + let all_assume := (newAssumption ++ (increaseBVars assume) ++ e'.assume); + -- TODO: factor out quant part + ⟨.quant .all (removeBound ty) tr'.translate (addAssumptions all_assume e'.translate), addBoundedWfAssume all_assume e'.wfCond, []⟩ | .quant .exist ty tr e => let newAssumption := match ty with @@ -190,36 +248,46 @@ def translateBounded [Coe String Identifier] (e: LExprT Identifier) (assumptions | _ => none; let e' := translateBounded e []; let tr' := translateBounded tr []; --TODO: need "clean" one here - let add_and x : LExprT Identifier := match newAssumption with + let add_and x : LExprT Identifier := + match newAssumption with | .some f => (.app (.app (LFunc.opExprT boolAndFunc) f (LMonoTy.arrow LMonoTy.bool LMonoTy.bool)) x LMonoTy.bool) | .none => x; - (.quant .exist (removeBound ty) tr'.1 (add_and (addAssumptions ((increaseBVars assumptions) ++ e'.2) e'.1)), []) - -- if-then-else is recursive, but we can add the assumptions to the condition. Likewise, we can test if the result is boolean-valued to see if we add to result + ⟨.quant .exist (removeBound ty) tr'.translate (add_and (addAssumptions ((increaseBVars assume) ++ e'.assume) e'.translate)), addBoundedWfAssume (newAssumption.toList ++ (increaseBVars assume) ++ e'.assume) e'.wfCond, []⟩ + /- + For if-then-else, we add assumptions to the condition, which is always bool-valued. For a bool-valued result, we can add the conditions freely. For a bounded-valued term, we produce a wf condition proving this. + -/ | .ite c t f .bool => let c' := translateBounded c []; let t' := translateBounded t []; let f' := translateBounded f []; - (.ite (addAssumptions (assumptions ++ c'.2) c'.1) (addAssumptions (assumptions ++ t'.2) t'.1) (addAssumptions (assumptions ++ f'.2) f'.1) .bool, []) + ⟨.ite (addAssumptions (assume ++ c'.assume) c'.translate) (addAssumptions (assume ++ t'.assume) t'.translate) (addAssumptions (assume ++ f'.assume) f'.translate) .bool, c'.wfCond ++ t'.wfCond ++ f'.wfCond ,[]⟩ | .ite c t f ty => let c' := translateBounded c []; - let t' := translateBounded t assumptions; - let f' := translateBounded f assumptions; - (.ite (addAssumptions (assumptions ++ c'.2) c'.1) t'.1 f'.1 ty, t'.2 ++ f'.2) - -- eq is bool so easy - | .eq e1 e2 ty => --shouldnt' need to check bool + let t' := translateBounded t assume; + let f' := translateBounded f assume; + let e' := .ite (addAssumptions (assume ++ c'.assume) c'.translate) t'.translate f'.translate ty; + ⟨e', (boundExprIfType ty e) ++ c'.wfCond ++ t'.wfCond ++ f'.wfCond, t'.assume ++ f'.assume⟩ + -- Equality is always bool-valued, so we can add assumptions + | .eq e1 e2 ty => let e1' := translateBounded e1 []; let e2' := translateBounded e2 []; - (addAssumptions (assumptions ++ e1'.2 ++ e2'.2) (.eq e1'.1 e2'.1 ty), []) - -- metadata just recursive + ⟨addAssumptions (assume ++ e1'.assume ++ e2'.assume) (.eq e1'.translate e2'.translate ty), e1'.wfCond ++ e2'.wfCond, []⟩ | .mdata m e => - let e' := translateBounded e assumptions; - (.mdata m e'.1, e'.2) + let e' := translateBounded e assume; + ⟨.mdata m e'.translate, e'.wfCond, e'.assume⟩ def translateBounded' [Coe String Identifier] (e: LExprT Identifier) : LExprT Identifier := - (translateBounded e []).1 + (translateBounded e []).translate + +def boundedWfConditions [Coe String Identifier] (e: LExprT Identifier) : List (LExprT Identifier) := + (translateBounded e []).wfCond + +-- NOTE: the assumptions are useful: they show us the "axioms" that we depend on (assumptions about external ops and free variables) -- Tests +namespace Test + -- NOTE: with a semantics for LExpr/LExprT, we could prove the equivalences mentioned above def natTy : LMonoTy := LMonoTy.bounded (.ble (.bconst 0) .bvar) @@ -229,8 +297,18 @@ def geOp (e1 e2: LExprT String) : LExprT String := .app (.app (LFunc.opExprT int def addOp (e1 e2: LExprT String) : LExprT String := .app (.app (LFunc.opExprT intAddFunc) e1 (.arrow .int .int)) e2 .int +def mulOp (e1 e2: LExprT String) : LExprT String := .app (.app (LFunc.opExprT intMulFunc) e1 (.arrow .int .int)) e2 .int + def notOp (e: LExprT String) : LExprT String := .app (LFunc.opExprT boolNotFunc) e .bool +-- easier to read +def eraseTy (x: LExprT String) := + LExpr.eraseTypes (LExprT.toLExpr x) + +def eraseTys l := List.map eraseTy l + +namespace TranslateTest + -- 1. ∀ (x: Nat), 0 <= x (quantified assumption) def test1 := (@LExprT.quant String .all natTy (.bvar 0 natTy) (leOp (.const "0" .int) (.bvar 0 .int))) @@ -238,7 +316,7 @@ def test1 := (@LExprT.quant String .all natTy (.bvar 0 natTy) (leOp (.const "0" #eval translateBounded' test1 --easier to read -#eval (LExpr.eraseTypes (LExprT.toLExpr (translateBounded' test1))) +#eval (eraseTy (translateBounded' test1) ) /-- info: Lambda.LExpr.quant (Lambda.QuantifierKind.all) @@ -828,29 +906,242 @@ info: Lambda.LExpr.abs #guard_msgs in #eval (LExprT.toLExpr (translateBounded' test6)) +end TranslateTest -/-- -Generate the constraints for a bounded lambda expression. They are: -1. For any int constant c : {x | b (x)}, generate constraint b(c) -2. For .app e1 e2, if e1 has type {x : b(x)} -> t, generate constraint b(e2) -3. For fvar -We also get information (NOTE: BOTH for translation and constraints): -3. For forall {x | b(x)} tr e have forall x, b(x) -> e -4. For exists {x | b(x)} tr e, have exists x, b(x) /\ e -For any free variables in the term, we add the assumption from their bound -(for example, suppose we have forall (x: Nat), x * (y: Nat) >= 0, - this results in y >= 0 -> forall (x: int), x * y >= 0) --/ -def bounded_constraints [Coe String Identifier] -(e: Lambda.LExpr BMonoTy Identifier) : List (Lambda.LExpr Lambda.LMonoTy Identifier) := - match e with - | .const x (some (.boundint b)) => - -- Constant must be int - match x.toInt? with - | some i => [bound_to_expr b (.const x LMonoTy.int )] - | none => [] --this will be caught during typechecking - | .app e1 e2 => - --problem: need to be at given type but dont have types yet +-- Tests for wf conditions + +namespace WFTest + +-- 1. constant: (1: Nat) +-- Expected: 0 <= 1 + +def test1 : LExprT String := .const "1" natTy + +/--info: [Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.const "1" none)] + -/ +#guard_msgs in +#eval eraseTys (boundedWfConditions test1) + +-- 2. application: (λ x: Nat. x) 1 +-- Expected: 0 <= 1 + +def test2 : LExprT String := .app (.abs (.bvar 0 .int) (.arrow natTy .int)) (.const "1" .int) .int + +#eval eraseTys (boundedWfConditions test2) + +-- 3. application with assumption (bottom up): (λ x: Nat. x) (foo : Nat) +-- Expected: 0 <= foo -> 0 <= foo + +def test3 : LExprT String := .app (.abs (.bvar 0 .int) (.arrow natTy .int)) (.op "foo" natTy) .int + +#eval eraseTys (boundedWfConditions test3) + +-- 4. application with assumption (top down): (λ x: Nat. (λ y: Nat. y) x) +-- Expected: 0 <= x -> 0 <= x + +def test4 : LExprT String := .abs (.app (.abs (.bvar 0 .int) (.arrow natTy .int)) (.bvar 0 .int) .int) (.arrow natTy .int) + +#eval eraseTys (boundedWfConditions test4) + +-- 5. abstraction with assumption: (λ x: Nat. foo (x + 1)) (foo: Nat -> int) +-- Expected: 0 <= x -> 0 <= x + 1 + +def test5 : LExprT String := .abs (.app (.op "foo" (.arrow natTy .int)) (addOp (.bvar 0 .int) (.const "1" .int)) .int) (.arrow natTy .int) + +#eval eraseTys (boundedWfConditions test5) + +-- 6. quantified assumption: (∃ x: Nat. foo x) where (foo: Nat -> int) +-- Expected: 0 <= x -> 0 <= x + +def test6 : LExprT String := .quant .exist natTy (.bvar 0 .int) (.app (.op "foo" (.arrow natTy .int)) (.bvar 0 .int) .int) + +#eval eraseTys (boundedWfConditions test6) + +-- 7. Lambda with bounded body (λ x: Nat. (x + 1: Nat)) +-- Expected: 0 <= x -> 0 <= x + 1 + +def test7 : LExprT String := .abs (addOp (.bvar 0 .int) (.const "1" .int)) (.arrow natTy natTy) + +#eval eraseTys (boundedWfConditions test7) + + +-- 8. Application with bounded body: (foo x) : Nat where foo: int -> Nat +-- Expected: [] (foo assumed) + +def test8 : LExprT String := .app (.op "foo" (.arrow .int natTy)) (.fvar "x" .int) natTy + +#eval eraseTys (boundedWfConditions test8) + +-- 9. Application with bounded body, no assumption: (λ (x: int) . x * x) 1 : Nat +-- Expected: 0 <= (λ (x: int) . x * x) 1 + +def test9 : LExprT String := .app (.abs (mulOp (.bvar 0 .int) (.bvar 0 .int)) (.arrow .int .int)) (.const "1" .int) natTy + +#eval eraseTys (boundedWfConditions test9) + +-- 10. If-then-else with bounded body: if b then 1 else 0 : Nat +-- Expected: 0 <= if b then 1 else 0 + +def test10 : LExprT String := .ite (.const "b" .bool) (.const "1" .int) (.const "0" .int) natTy + +#eval eraseTys (boundedWfConditions test10) + +-- 11. If-then-else with bound functions: if b then λ (x: int). x * x : Nat else λ (y: int). 0 : Nat (whole type int -> Nat) +-- Expected: 0 <= x * x; 0 <= 0 + +def test11 : LExprT String := .ite (.const "b" .bool) (.abs (mulOp (.bvar 0 .int) (.bvar 0 .int)) (.arrow .int natTy)) (.abs (.const "0" .int) (.arrow .int natTy)) (.arrow .int natTy) + +#eval eraseTys (boundedWfConditions test11) + +end WFTest + +--tests with more sophisticated bounded types (mostly AI generated) + +namespace OtherTest + +-- Test 1: Nested bounded function applications +-- Input: add : {x < 10} → {y < 5} → {z < 15}, applied to (3 : {x < 10}) and (2 : {x < 5}) +-- Expected: translate = add 3 2, wfCond = [3 < 10, 2 < 5] (note: there are duplicates, should fix - change to Set) + +def testNestedBoundedApps : LExprT String := + .app (.app (.op "add" (.arrow (.bounded (.blt (.bvar) (.bconst 10))) + (.arrow (.bounded (.blt (.bvar) (.bconst 5))) + (.bounded (.blt (.bvar) (.bconst 15)))))) + (.const "3" (.bounded (.blt (.bvar) (.bconst 10)))) + (.arrow (.bounded (.blt (.bvar) (.bconst 5))) (.bounded (.blt (.bvar) (.bconst 15))))) + (.const "2" (.bounded (.blt (.bvar) (.bconst 5)))) + (.bounded (.blt (.bvar) (.bconst 15))) + +#eval eraseTys (boundedWfConditions testNestedBoundedApps) +#eval eraseTy (translateBounded' testNestedBoundedApps) + +-- Test 2: Bounded variable in quantifier with complex bound expression +-- Input: ∀ (x : {x < 100 ∧ 0 ≤ x}). x = 42 +-- Expected: translate = ∀ x:int. (x < 100 ∧ 0 ≤ x) → (x = 42), wfCond = [] + +def testComplexBoundInQuantifier : LExprT String := + .quant .all (.bounded (.band (.blt (.bvar) (.bconst 100)) + (.ble (.bconst 0) (.bvar)))) + (.bvar 0 (.bounded (.band (.blt (.bvar) (.bconst 100)) + (.ble (.bconst 0) (.bvar))))) + (.eq (.bvar 0 (.bounded (.band (.blt (.bvar) (.bconst 100)) + (.ble (.bconst 0) (.bvar))))) + (.const "42" .int) .bool) + +#eval eraseTys (boundedWfConditions testComplexBoundInQuantifier) +#eval eraseTy (translateBounded' testComplexBoundInQuantifier) + +-- Test 3: If-then-else with bounded branches and boolean condition +-- Input: if ((getValue 5 : {0 ≤ x}) > 0) then (1 : {x < 10}) else (0 : {x < 10}) : {x < 10} +-- Expected: translate = if (0 ≤ getValue 5) → (getValue 5 > 0) then 1 else 0, wfCond = [1 < 10, 0 < 10, (if (getValue 5) then 1 else 0) < 10] +def testBoundedIte : LExprT String := + .ite (.app (.app (LFunc.opExprT intLtFunc) + (.const "0" .int) + (.arrow .int .bool)) + (.app (.op "getValue" (.arrow .int (.bounded (.ble (.bconst 0) (.bvar))))) + (.const "5" .int) + (.bounded (.ble (.bconst 0) (.bvar)))) + .bool) + (.const "1" (.bounded (.blt (.bvar) (.bconst 10)))) + (.const "0" (.bounded (.blt (.bvar) (.bconst 10)))) + (.bounded (.blt (.bvar) (.bconst 10))) + +#eval eraseTys (boundedWfConditions testBoundedIte) +#eval eraseTy (translateBounded' testBoundedIte) + +-- Test 4: Lambda with bounded parameter and bounded return type +-- Input: λ (x : {0 ≤ x}). increment x : {y < 100} +-- Expected: translate = λ x:int. increment x, wfCond = [∀ x:int. 0 ≤ x → increment x < 100; 0 <= x -> 0 <= x] +def testBoundedLambda : LExprT String := + .abs (.app (.op "increment" (.arrow (.bounded (.ble (.bconst 0) (.bvar))) + (.bounded (.blt (.bvar) (.bconst 100))))) + (.bvar 0 (.bounded (.ble (.bconst 0) (.bvar)))) + (.bounded (.blt (.bvar) (.bconst 100)))) + (.arrow (.bounded (.ble (.bconst 0) (.bvar))) + (.bounded (.blt (.bvar) (.bconst 100)))) + +#eval eraseTys (boundedWfConditions testBoundedLambda) +#eval eraseTy (translateBounded' testBoundedLambda) + +-- Test 5: Equality between bounded expressions +-- Input: (square (3 : {-10 ≤ x ≤ 10}) : {0 ≤ y}) = (9 : {0 ≤ z}) +-- Expected: translate = square 3 = 9, wfCond = [-10 ≤ 3 ≤ 10, 0 ≤ 9] +def testBoundedEquality : LExprT String := + .eq (.app (.op "square" (.arrow (.bounded (.band (.ble (.bconst (-10)) (.bvar)) + (.ble (.bvar) (.bconst 10)))) + (.bounded (.ble (.bconst 0) (.bvar))))) + (.const "3" (.bounded (.band (.ble (.bconst (-10)) (.bvar)) + (.ble (.bvar) (.bconst 10))))) + (.bounded (.ble (.bconst 0) (.bvar)))) + (.const "9" (.bounded (.ble (.bconst 0) (.bvar)))) + .bool + +#eval eraseTys (boundedWfConditions testBoundedEquality) +#eval eraseTy (translateBounded' testBoundedEquality) + +-- Test 6: Free variable with bounded type in assumptions +-- Input: compare (x : {x < 50}) 25, with assumption [x < 50] +-- Expected: translate = (x < 50) → compare x 25, wfCond = [] +def testFreeVarWithAssumptions : LExprT String := + .app (.app (.op "compare" (.arrow .int (.arrow .int .bool))) + (.fvar "x" (.bounded (.blt (.bvar) (.bconst 50)))) + (.arrow .int .bool)) + (.const "25" .int) + .bool + +#eval eraseTys (boundedWfConditions testFreeVarWithAssumptions) +#eval eraseTy (translateBounded' testFreeVarWithAssumptions) + +-- Test 7: Metadata preservation with bounded types +-- Input: @metadata(42 : {0 ≤ x < 100}) +-- Expected: translate = @metadata(42), wfCond = [0 ≤ 42 < 100] +def testMetadataWithBounds : LExprT String := + .mdata (Info.mk "test_info") + (.const "42" (.bounded (.band (.ble (.bconst 0) (.bvar)) + (.blt (.bvar) (.bconst 100))))) + +#eval eraseTys (boundedWfConditions testMetadataWithBounds) +#eval eraseTy (translateBounded' testMetadataWithBounds) + +-- Test 8: Chain of bounded function applications +-- Input: f3 (f2 (f1 5 : {x < 10}) : {x < 20}) : {x < 30} +-- Expected: translate = f3 (f2 (f1 5)), wfCond = [f1 5 < 10 → f2 (f1 5) < 20 -> f2 (f1 5) < 20, f1 5 < 10 -> f1 5 < 10] +def testBoundedChain : LExprT String := + .app (.op "f3" (.arrow (.bounded (.blt (.bvar) (.bconst 20))) + (.bounded (.blt (.bvar) (.bconst 30))))) + (.app (.op "f2" (.arrow (.bounded (.blt (.bvar) (.bconst 10))) + (.bounded (.blt (.bvar) (.bconst 20))))) + (.app (.op "f1" (.arrow .int (.bounded (.blt (.bvar) (.bconst 10))))) + (.const "5" .int) + (.bounded (.blt (.bvar) (.bconst 10)))) + (.bounded (.blt (.bvar) (.bconst 20)))) + (.bounded (.blt (.bvar) (.bconst 30))) + +#eval eraseTys (boundedWfConditions testBoundedChain) +#eval eraseTy (translateBounded' testBoundedChain) + +-- Test 9: Bounded type in boolean context with negation +-- Input: ¬((getValue 10 : {0 ≤ x}) < 5) +-- Expected: translate = (0 ≤ getValue 10) → ¬(getValue 10 < 5), wfCond = [] +def testBoundedInBoolContext : LExprT String := + .app (LFunc.opExprT boolNotFunc) + (.app (.app (LFunc.opExprT intLtFunc) + (.app (.op "getValue" (.arrow .int (.bounded (.ble (.bconst 0) (.bvar))))) + (.const "10" .int) + (.bounded (.ble (.bconst 0) (.bvar)))) + (.arrow .int .bool)) + (.const "5" .int) + .bool) + .bool + +#eval eraseTys (boundedWfConditions testBoundedInBoolContext) +#eval eraseTy (translateBounded' testBoundedInBoolContext) + +end OtherTest + +end Test end Bounded From 184211f7d7b6cc4f73bf361e7dbd2b4ff14d9a0f Mon Sep 17 00:00:00 2001 From: Josh Cohen Date: Fri, 26 Sep 2025 13:59:58 -0400 Subject: [PATCH 07/19] Some improvements to code and tests Change List -> Set to avoid duplicates Remove quantifiers from output wf conditions when possible Add guard_msgs for tests Split tests into separate file Fix bug in abstraction case --- Strata/DL/Bounded/BExpr.lean | 944 +++---------------------------- Strata/DL/Bounded/BExprTest.lean | 738 ++++++++++++++++++++++++ Strata/DL/Bounded/BTy.lean | 61 -- 3 files changed, 806 insertions(+), 937 deletions(-) create mode 100644 Strata/DL/Bounded/BExprTest.lean delete mode 100644 Strata/DL/Bounded/BTy.lean diff --git a/Strata/DL/Bounded/BExpr.lean b/Strata/DL/Bounded/BExpr.lean index 7e658ad7..40d75c71 100644 --- a/Strata/DL/Bounded/BExpr.lean +++ b/Strata/DL/Bounded/BExpr.lean @@ -5,7 +5,6 @@ -/ import Strata.DL.Lambda.LTy --- import Strata.DL.Bounded.BTy import Strata.DL.Lambda.LExprT import Strata.DL.Lambda.IntBoolFactory @@ -27,6 +26,8 @@ namespace Bounded open Std (ToFormat Format format) open Lambda +variable {Identifier : Type} [ToString Identifier] [DecidableEq Identifier] [ToFormat Identifier] [HasGen Identifier] + /- There are two parts to desugaring BLambda to Lambda: translating the terms @@ -106,10 +107,28 @@ def increaseBVar (e: LExprT Identifier) : LExprT Identifier := def increaseBVars (l: List (LExprT Identifier)) : List (LExprT Identifier) := List.map increaseBVar l +-- TODO: see what to use here +def ListSet α := List α +-- #print Membership +-- instance : Membership α (ListSet α) := +-- Membership.mk (fun l x => by unfold ListSet at l; exact (x ∈ l)) + +def ListSet.contains {α} [DecidableEq α] (l: ListSet α) (x: α) : Bool := + List.foldr (fun y acc => x == y || acc) false l + +def ListSet.insert {α} [DecidableEq α] (l: ListSet α) (x: α) : ListSet α := + if ListSet.contains l x then l else x :: l + +def ListSet.insertAll {α} [DecidableEq α] (l: ListSet α) (xs: List α) : ListSet α := + List.foldr (fun x acc => ListSet.insert acc x) l xs + +def ListSet.union {α} [DecidableEq α] (l: List (ListSet α)) : ListSet α := + List.foldr ListSet.insertAll [] l + /- Do translation and wf generation in 1 go - need same assumptions, expesive to compute both-/ structure translationRes Identifier where (translate : LExprT Identifier) -(wfCond : List (LExprT Identifier)) +(wfCond : ListSet (LExprT Identifier)) (assume : List (LExprT Identifier)) -- more aux functions @@ -126,7 +145,6 @@ def wfCallCondition [Coe String Identifier] (assume : List (LExprT Identifier)) [addAssumptions assume (BoundExprToLExprT b e2)] | _ => [] ---TODO: change to avoid quantifier -- NOTE: l1 includes assumptions, l2 does not def addBoundedWf [Coe String Identifier] (assume: List (LExprT Identifier)) (l1 l2: List (LExprT Identifier)) : List (LExprT Identifier) := List.map (fun e => .quant .all .int (.bvar 0 .int) (addAssumptions assume e)) l1 ++ List.map (.quant .all .int (.bvar 0 .int)) l2 @@ -134,11 +152,6 @@ def addBoundedWf [Coe String Identifier] (assume: List (LExprT Identifier)) (l1 def addBoundedWfAssume [Coe String Identifier] (assume: List (LExprT Identifier)) (l: List (LExprT Identifier)) := addBoundedWf assume l [] --- for bounded-valued terms, need to generate condition that body satisfies bound --- def wfBoundedBody [Coe String Identifier] (ty: LMonoTy) (e: LExprT Identifier) := --- ((isBounded ty).map (fun b => BoundExprToLExprT b e)).toList - - /-- Translate expression to one without bounded ints. This is a non-trivial translation, as we want to preserve the semantics of the resulting term. @@ -161,7 +174,7 @@ Note that we do NOT have to worry about the fact that the new functions have a l Invariant: assumptions must not have bounded types (TODO enforce this), same for inputs Invariant (I think): All assumptions are of form: b(bvar #x) -/ -def translateBounded [Coe String Identifier] (e: LExprT Identifier) (assume: List (LExprT Identifier)) : translationRes Identifier := +def translateBounded [Coe String Identifier] [DecidableEq Identifier] (e: LExprT Identifier) (assume: List (LExprT Identifier)) : translationRes Identifier := match e with -- constants do not need assumptions; they produce a wf goal if bounded | .const c ty => @@ -189,7 +202,7 @@ def translateBounded [Coe String Identifier] (e: LExprT Identifier) (assume: Lis let e2' := translateBounded e2 []; let all_assumes := assume ++ e1'.assume ++ e2'.assume; let res := addAssumptions all_assumes (.app e1'.translate e2'.translate (removeBound ty)); - ⟨res, (wfCallCondition (assume ++ e2'.assume) e1 e2'.translate) ++ e1'.wfCond ++ e2'.wfCond, []⟩ + ⟨res, ListSet.union [wfCallCondition (assume ++ e2'.assume) e1 e2'.translate, e1'.wfCond, e2'.wfCond], []⟩ else let e1' := translateBounded e1 assume; let e2' := translateBounded e2 assume; @@ -199,7 +212,7 @@ def translateBounded [Coe String Identifier] (e: LExprT Identifier) (assume: Lis | .arrow _ .int, .bounded _ => boundExprIfType ty e' | _, _ => []; - ⟨e', (wfCallCondition (assume ++ e2'.assume) e1 e2'.translate) ++ extraWf ++ e1'.wfCond ++ e2'.wfCond, e1'.assume ++ e2'.assume ++ boundExprIfType ty e'⟩ + ⟨e', ListSet.union [wfCallCondition (assume ++ e2'.assume) e1 e2'.translate, extraWf, e1'.wfCond, e2'.wfCond], boundExprIfType ty e' ++ e1'.assume ++ e2'.assume⟩ /- Abstraction is the most complex case: 1. If the argument is bounded, add as top-down assumption @@ -224,7 +237,7 @@ def translateBounded [Coe String Identifier] (e: LExprT Identifier) (assume: Lis let e' := translateBounded e all_assume; let e'' := .abs e'.translate (removeBound ty); -- Note: don't add assumptions to e'.wfCond, already included - ⟨e'', addBoundedWf all_assume (boundExprIfType ty1 e'') e'.wfCond, e'.assume⟩ + ⟨e'', addBoundedWf all_assume (boundExprIfType ty1 e'.translate) e'.wfCond, e'.assume⟩ | _ => ⟨.const "0" .int, [], []⟩ -- error case /- Quantifiers are simpler because they are boolean-valued. ∀ (x : nat). e adds an assumption x >= 0 -> ..., while ∃ (x: nat). e results in x >= 0 ∧ .. @@ -260,888 +273,67 @@ def translateBounded [Coe String Identifier] (e: LExprT Identifier) (assume: Lis let c' := translateBounded c []; let t' := translateBounded t []; let f' := translateBounded f []; - ⟨.ite (addAssumptions (assume ++ c'.assume) c'.translate) (addAssumptions (assume ++ t'.assume) t'.translate) (addAssumptions (assume ++ f'.assume) f'.translate) .bool, c'.wfCond ++ t'.wfCond ++ f'.wfCond ,[]⟩ + ⟨.ite (addAssumptions (assume ++ c'.assume) c'.translate) (addAssumptions (assume ++ t'.assume) t'.translate) (addAssumptions (assume ++ f'.assume) f'.translate) .bool, ListSet.union [c'.wfCond, t'.wfCond, f'.wfCond] ,[]⟩ | .ite c t f ty => let c' := translateBounded c []; let t' := translateBounded t assume; let f' := translateBounded f assume; let e' := .ite (addAssumptions (assume ++ c'.assume) c'.translate) t'.translate f'.translate ty; - ⟨e', (boundExprIfType ty e) ++ c'.wfCond ++ t'.wfCond ++ f'.wfCond, t'.assume ++ f'.assume⟩ + ⟨e', ListSet.union [boundExprIfType ty e, c'.wfCond, t'.wfCond, f'.wfCond], t'.assume ++ f'.assume⟩ -- Equality is always bool-valued, so we can add assumptions | .eq e1 e2 ty => let e1' := translateBounded e1 []; let e2' := translateBounded e2 []; - ⟨addAssumptions (assume ++ e1'.assume ++ e2'.assume) (.eq e1'.translate e2'.translate ty), e1'.wfCond ++ e2'.wfCond, []⟩ + ⟨addAssumptions (assume ++ e1'.assume ++ e2'.assume) (.eq e1'.translate e2'.translate ty), ListSet.union [e1'.wfCond, e2'.wfCond], []⟩ | .mdata m e => let e' := translateBounded e assume; ⟨.mdata m e'.translate, e'.wfCond, e'.assume⟩ -def translateBounded' [Coe String Identifier] (e: LExprT Identifier) : LExprT Identifier := +-- TODO: name +def translateBounded' [Coe String Identifier] (e: LExprT Identifier) : LExprT Identifier := (translateBounded e []).translate -def boundedWfConditions [Coe String Identifier] (e: LExprT Identifier) : List (LExprT Identifier) := - (translateBounded e []).wfCond - --- NOTE: the assumptions are useful: they show us the "axioms" that we depend on (assumptions about external ops and free variables) - --- Tests - -namespace Test - --- NOTE: with a semantics for LExpr/LExprT, we could prove the equivalences mentioned above - -def natTy : LMonoTy := LMonoTy.bounded (.ble (.bconst 0) .bvar) -def leOp (e1 e2: LExprT String) : LExprT String := .app (.app (LFunc.opExprT intLeFunc) e1 (.arrow .int .bool)) e2 .bool - -def geOp (e1 e2: LExprT String) : LExprT String := .app (.app (LFunc.opExprT intGeFunc) e1 (.arrow .int .bool)) e2 .bool - -def addOp (e1 e2: LExprT String) : LExprT String := .app (.app (LFunc.opExprT intAddFunc) e1 (.arrow .int .int)) e2 .int - -def mulOp (e1 e2: LExprT String) : LExprT String := .app (.app (LFunc.opExprT intMulFunc) e1 (.arrow .int .int)) e2 .int - -def notOp (e: LExprT String) : LExprT String := .app (LFunc.opExprT boolNotFunc) e .bool - --- easier to read -def eraseTy (x: LExprT String) := - LExpr.eraseTypes (LExprT.toLExpr x) - -def eraseTys l := List.map eraseTy l - -namespace TranslateTest - --- 1. ∀ (x: Nat), 0 <= x (quantified assumption) - -def test1 := (@LExprT.quant String .all natTy (.bvar 0 natTy) (leOp (.const "0" .int) (.bvar 0 .int))) - -#eval translateBounded' test1 - ---easier to read -#eval (eraseTy (translateBounded' test1) ) - -/-- info: Lambda.LExpr.quant - (Lambda.QuantifierKind.all) - (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) - (Lambda.LExpr.bvar 0) - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Bool.Implies" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Int.Le" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) - (Lambda.LExpr.bvar 0))) - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Int.Le" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) - (Lambda.LExpr.bvar 0))) --/ -#guard_msgs in -#eval (LExprT.toLExpr (translateBounded' test1)) - --- 2. λ(x: Nat), if 0 <= x then 1 else 0 (assumption inside ite) - -def test2 : LExprT String := .abs (.ite (leOp (.const "0" .int) (.bvar 0 .int)) (.const "1" .int) (.const "2" .int) .int) (.arrow natTy .int) - -#eval translateBounded' test2 - -#eval (LExpr.eraseTypes (LExprT.toLExpr (translateBounded' test2))) - -/-- info: Lambda.LExpr.abs - (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) - (Lambda.LExpr.ite - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Bool.Implies" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Int.Le" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) - (Lambda.LExpr.bvar 0))) - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Int.Le" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) - (Lambda.LExpr.bvar 0))) - (Lambda.LExpr.const "1" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.const "2" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) -/ -#guard_msgs in -#eval (LExprT.toLExpr (translateBounded' test2)) - --- 3. λ(x: int), if foo x >= 0 then 1 else 0 (for foo: int -> Nat) (propagate op/application information) - -def test3 : LExprT String := .abs (.ite (geOp (.app (.op "foo" (.arrow .int natTy)) (.bvar 0 .int) natTy) (.const "0" .int)) (.const "1" .int) (.const "0" .int) .int) (.arrow .int .int) - -#eval translateBounded' test3 - -#eval (LExpr.eraseTypes (LExprT.toLExpr (translateBounded' test3))) - -/-- -info: Lambda.LExpr.abs - (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) - (Lambda.LExpr.ite - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Bool.Implies" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Int.Le" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) - (Lambda.LExpr.app - (Lambda.LExpr.op - "foo" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.bvar 0)))) - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Int.Ge" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.app - (Lambda.LExpr.op - "foo" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.bvar 0))) - (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata)))))) - (Lambda.LExpr.const "1" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) - -/ -#guard_msgs in -#eval (LExprT.toLExpr (translateBounded' test3)) - --- 4. (x: Nat) + (y: Nat) >= 0 (free variable bounds) - -def test4 : LExprT String := geOp (addOp (.fvar "x" natTy) (.fvar "y" natTy)) (.const "0" .int) - -#eval translateBounded' test4 - -#eval (LExpr.eraseTypes (LExprT.toLExpr (translateBounded' test4))) - -/-- -info: Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Bool.Implies" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Int.Le" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) - (Lambda.LExpr.fvar "x" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata)))))) - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Bool.Implies" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Int.Le" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) - (Lambda.LExpr.fvar "y" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata)))))) - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Int.Ge" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Int.Add" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.fvar "x" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) - (Lambda.LExpr.fvar "y" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata)))))) - (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata)))))) -/ -#guard_msgs in -#eval (LExprT.toLExpr (translateBounded' test4)) - --- 5. λ (x: Nat). λ (y: Nat). x + y >= 0 (multiple bound vars) - -def test5 : LExprT String := .abs (.abs (geOp (addOp (.bvar 1 .int) (.bvar 0 .int)) (.const "0" .int)) (.arrow natTy .bool)) (.arrow natTy (.arrow natTy .bool)) - -#eval translateBounded' test5 - -#eval (LExpr.eraseTypes (LExprT.toLExpr (translateBounded' test5))) +-- TEMPORARY until LExpr/LExprT are unified +def LExprT.substK (k : Nat) (s : LExprT Identifier) (e : LExprT Identifier) : LExprT Identifier := + match e with + | .const c ty => .const c ty + | .op o ty => .op o ty + | .bvar i ty => if (i == k) then s else .bvar i ty + | .fvar y ty => .fvar y ty + | .mdata info e' => .mdata info (substK k s e') + | .abs e' ty => .abs (substK (k + 1) s e') ty + | .quant qk ty tr' e' => .quant qk ty (substK (k + 1) s tr') (substK (k + 1) s e') + | .app e1 e2 ty => .app (substK k s e1) (substK k s e2) ty + | .ite c t e ty => .ite (substK k s c) (substK k s t) (substK k s e) ty + | .eq e1 e2 ty => .eq (substK k s e1) (substK k s e2) ty + +def LExprT.varOpen (k : Nat) (x : Identifier × LMonoTy) (e : LExprT Identifier) : LExprT Identifier := + LExprT.substK k (.fvar x.fst x.snd) e -/-- -info: Lambda.LExpr.abs - (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) - (Lambda.LExpr.abs - (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Bool.Implies" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Int.Le" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) - (Lambda.LExpr.bvar 0))) - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Bool.Implies" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Int.Le" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) - (Lambda.LExpr.bvar 1))) - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Int.Ge" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Int.Add" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.bvar 1)) - (Lambda.LExpr.bvar 0))) - (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata)))))))) +/- +Many of the wf conditions have the form: forall x, f(x). We remove the quantifiers to make the SMT formulas easier to solve. +Must be stateful because we need to generate fresh variables -/ -#guard_msgs in -#eval (LExprT.toLExpr (translateBounded' test5)) - - --- 6. λ (x: Nat), if foo then λ (y: Nat). not (x = -1) else λ (y: Nat). x + y >= 0 (propagate inside branches of if-then-else) - -def test6 : LExprT String := .abs (.ite (.op "foo" .bool) (.abs (notOp (.eq (.bvar 1 .int) (.const "-1" .int) .bool)) (.arrow natTy .bool)) (.abs (geOp (addOp (.bvar 1 .int) (.bvar 0 .int)) (.const "0" .int)) (.arrow natTy .bool)) (.arrow natTy .bool)) (.arrow natTy (.arrow natTy .bool)) - -#eval translateBounded' test6 - -#eval (LExpr.eraseTypes (LExprT.toLExpr (translateBounded' test6))) - -/-- -info: Lambda.LExpr.abs - (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) - (Lambda.LExpr.ite - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Bool.Implies" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Int.Le" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) - (Lambda.LExpr.bvar 0))) - (Lambda.LExpr.op "foo" (some (Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata))))) - (Lambda.LExpr.abs - (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Bool.Implies" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Int.Le" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) - (Lambda.LExpr.bvar 0))) - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Bool.Implies" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Int.Le" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) - (Lambda.LExpr.bvar 1))) - (Lambda.LExpr.app - (Lambda.LExpr.op - "Bool.Not" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.eq - (Lambda.LExpr.bvar 1) - (Lambda.LExpr.const "-1" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))))))) - (Lambda.LExpr.abs - (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Bool.Implies" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Int.Le" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) - (Lambda.LExpr.bvar 0))) - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Bool.Implies" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Int.Le" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))) - (Lambda.LExpr.bvar 1))) - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Int.Ge" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.app - (Lambda.LExpr.app - (Lambda.LExpr.op - "Int.Add" - (some (Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons - "arrow" - [Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata), - Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)] - (Lambda.LTyRestrict.nodata)))) - (Lambda.LExpr.bvar 1)) - (Lambda.LExpr.bvar 0))) - (Lambda.LExpr.const "0" (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))))))))) - -/ -#guard_msgs in -#eval (LExprT.toLExpr (translateBounded' test6)) - -end TranslateTest - --- Tests for wf conditions - -namespace WFTest - --- 1. constant: (1: Nat) --- Expected: 0 <= 1 - -def test1 : LExprT String := .const "1" natTy - -/--info: [Lambda.LExpr.app - (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) - (Lambda.LExpr.const "1" none)] - -/ -#guard_msgs in -#eval eraseTys (boundedWfConditions test1) - --- 2. application: (λ x: Nat. x) 1 --- Expected: 0 <= 1 - -def test2 : LExprT String := .app (.abs (.bvar 0 .int) (.arrow natTy .int)) (.const "1" .int) .int - -#eval eraseTys (boundedWfConditions test2) - --- 3. application with assumption (bottom up): (λ x: Nat. x) (foo : Nat) --- Expected: 0 <= foo -> 0 <= foo - -def test3 : LExprT String := .app (.abs (.bvar 0 .int) (.arrow natTy .int)) (.op "foo" natTy) .int - -#eval eraseTys (boundedWfConditions test3) - --- 4. application with assumption (top down): (λ x: Nat. (λ y: Nat. y) x) --- Expected: 0 <= x -> 0 <= x - -def test4 : LExprT String := .abs (.app (.abs (.bvar 0 .int) (.arrow natTy .int)) (.bvar 0 .int) .int) (.arrow natTy .int) - -#eval eraseTys (boundedWfConditions test4) - --- 5. abstraction with assumption: (λ x: Nat. foo (x + 1)) (foo: Nat -> int) --- Expected: 0 <= x -> 0 <= x + 1 - -def test5 : LExprT String := .abs (.app (.op "foo" (.arrow natTy .int)) (addOp (.bvar 0 .int) (.const "1" .int)) .int) (.arrow natTy .int) - -#eval eraseTys (boundedWfConditions test5) - --- 6. quantified assumption: (∃ x: Nat. foo x) where (foo: Nat -> int) --- Expected: 0 <= x -> 0 <= x - -def test6 : LExprT String := .quant .exist natTy (.bvar 0 .int) (.app (.op "foo" (.arrow natTy .int)) (.bvar 0 .int) .int) - -#eval eraseTys (boundedWfConditions test6) - --- 7. Lambda with bounded body (λ x: Nat. (x + 1: Nat)) --- Expected: 0 <= x -> 0 <= x + 1 - -def test7 : LExprT String := .abs (addOp (.bvar 0 .int) (.const "1" .int)) (.arrow natTy natTy) - -#eval eraseTys (boundedWfConditions test7) - - --- 8. Application with bounded body: (foo x) : Nat where foo: int -> Nat --- Expected: [] (foo assumed) - -def test8 : LExprT String := .app (.op "foo" (.arrow .int natTy)) (.fvar "x" .int) natTy - -#eval eraseTys (boundedWfConditions test8) - --- 9. Application with bounded body, no assumption: (λ (x: int) . x * x) 1 : Nat --- Expected: 0 <= (λ (x: int) . x * x) 1 - -def test9 : LExprT String := .app (.abs (mulOp (.bvar 0 .int) (.bvar 0 .int)) (.arrow .int .int)) (.const "1" .int) natTy - -#eval eraseTys (boundedWfConditions test9) - --- 10. If-then-else with bounded body: if b then 1 else 0 : Nat --- Expected: 0 <= if b then 1 else 0 - -def test10 : LExprT String := .ite (.const "b" .bool) (.const "1" .int) (.const "0" .int) natTy - -#eval eraseTys (boundedWfConditions test10) - --- 11. If-then-else with bound functions: if b then λ (x: int). x * x : Nat else λ (y: int). 0 : Nat (whole type int -> Nat) --- Expected: 0 <= x * x; 0 <= 0 - -def test11 : LExprT String := .ite (.const "b" .bool) (.abs (mulOp (.bvar 0 .int) (.bvar 0 .int)) (.arrow .int natTy)) (.abs (.const "0" .int) (.arrow .int natTy)) (.arrow .int natTy) - -#eval eraseTys (boundedWfConditions test11) - -end WFTest - ---tests with more sophisticated bounded types (mostly AI generated) - -namespace OtherTest - --- Test 1: Nested bounded function applications --- Input: add : {x < 10} → {y < 5} → {z < 15}, applied to (3 : {x < 10}) and (2 : {x < 5}) --- Expected: translate = add 3 2, wfCond = [3 < 10, 2 < 5] (note: there are duplicates, should fix - change to Set) - -def testNestedBoundedApps : LExprT String := - .app (.app (.op "add" (.arrow (.bounded (.blt (.bvar) (.bconst 10))) - (.arrow (.bounded (.blt (.bvar) (.bconst 5))) - (.bounded (.blt (.bvar) (.bconst 15)))))) - (.const "3" (.bounded (.blt (.bvar) (.bconst 10)))) - (.arrow (.bounded (.blt (.bvar) (.bconst 5))) (.bounded (.blt (.bvar) (.bconst 15))))) - (.const "2" (.bounded (.blt (.bvar) (.bconst 5)))) - (.bounded (.blt (.bvar) (.bconst 15))) - -#eval eraseTys (boundedWfConditions testNestedBoundedApps) -#eval eraseTy (translateBounded' testNestedBoundedApps) - --- Test 2: Bounded variable in quantifier with complex bound expression --- Input: ∀ (x : {x < 100 ∧ 0 ≤ x}). x = 42 --- Expected: translate = ∀ x:int. (x < 100 ∧ 0 ≤ x) → (x = 42), wfCond = [] - -def testComplexBoundInQuantifier : LExprT String := - .quant .all (.bounded (.band (.blt (.bvar) (.bconst 100)) - (.ble (.bconst 0) (.bvar)))) - (.bvar 0 (.bounded (.band (.blt (.bvar) (.bconst 100)) - (.ble (.bconst 0) (.bvar))))) - (.eq (.bvar 0 (.bounded (.band (.blt (.bvar) (.bconst 100)) - (.ble (.bconst 0) (.bvar))))) - (.const "42" .int) .bool) - -#eval eraseTys (boundedWfConditions testComplexBoundInQuantifier) -#eval eraseTy (translateBounded' testComplexBoundInQuantifier) - --- Test 3: If-then-else with bounded branches and boolean condition --- Input: if ((getValue 5 : {0 ≤ x}) > 0) then (1 : {x < 10}) else (0 : {x < 10}) : {x < 10} --- Expected: translate = if (0 ≤ getValue 5) → (getValue 5 > 0) then 1 else 0, wfCond = [1 < 10, 0 < 10, (if (getValue 5) then 1 else 0) < 10] -def testBoundedIte : LExprT String := - .ite (.app (.app (LFunc.opExprT intLtFunc) - (.const "0" .int) - (.arrow .int .bool)) - (.app (.op "getValue" (.arrow .int (.bounded (.ble (.bconst 0) (.bvar))))) - (.const "5" .int) - (.bounded (.ble (.bconst 0) (.bvar)))) - .bool) - (.const "1" (.bounded (.blt (.bvar) (.bconst 10)))) - (.const "0" (.bounded (.blt (.bvar) (.bconst 10)))) - (.bounded (.blt (.bvar) (.bconst 10))) - -#eval eraseTys (boundedWfConditions testBoundedIte) -#eval eraseTy (translateBounded' testBoundedIte) - --- Test 4: Lambda with bounded parameter and bounded return type --- Input: λ (x : {0 ≤ x}). increment x : {y < 100} --- Expected: translate = λ x:int. increment x, wfCond = [∀ x:int. 0 ≤ x → increment x < 100; 0 <= x -> 0 <= x] -def testBoundedLambda : LExprT String := - .abs (.app (.op "increment" (.arrow (.bounded (.ble (.bconst 0) (.bvar))) - (.bounded (.blt (.bvar) (.bconst 100))))) - (.bvar 0 (.bounded (.ble (.bconst 0) (.bvar)))) - (.bounded (.blt (.bvar) (.bconst 100)))) - (.arrow (.bounded (.ble (.bconst 0) (.bvar))) - (.bounded (.blt (.bvar) (.bconst 100)))) - -#eval eraseTys (boundedWfConditions testBoundedLambda) -#eval eraseTy (translateBounded' testBoundedLambda) - --- Test 5: Equality between bounded expressions --- Input: (square (3 : {-10 ≤ x ≤ 10}) : {0 ≤ y}) = (9 : {0 ≤ z}) --- Expected: translate = square 3 = 9, wfCond = [-10 ≤ 3 ≤ 10, 0 ≤ 9] -def testBoundedEquality : LExprT String := - .eq (.app (.op "square" (.arrow (.bounded (.band (.ble (.bconst (-10)) (.bvar)) - (.ble (.bvar) (.bconst 10)))) - (.bounded (.ble (.bconst 0) (.bvar))))) - (.const "3" (.bounded (.band (.ble (.bconst (-10)) (.bvar)) - (.ble (.bvar) (.bconst 10))))) - (.bounded (.ble (.bconst 0) (.bvar)))) - (.const "9" (.bounded (.ble (.bconst 0) (.bvar)))) - .bool - -#eval eraseTys (boundedWfConditions testBoundedEquality) -#eval eraseTy (translateBounded' testBoundedEquality) - --- Test 6: Free variable with bounded type in assumptions --- Input: compare (x : {x < 50}) 25, with assumption [x < 50] --- Expected: translate = (x < 50) → compare x 25, wfCond = [] -def testFreeVarWithAssumptions : LExprT String := - .app (.app (.op "compare" (.arrow .int (.arrow .int .bool))) - (.fvar "x" (.bounded (.blt (.bvar) (.bconst 50)))) - (.arrow .int .bool)) - (.const "25" .int) - .bool - -#eval eraseTys (boundedWfConditions testFreeVarWithAssumptions) -#eval eraseTy (translateBounded' testFreeVarWithAssumptions) - --- Test 7: Metadata preservation with bounded types --- Input: @metadata(42 : {0 ≤ x < 100}) --- Expected: translate = @metadata(42), wfCond = [0 ≤ 42 < 100] -def testMetadataWithBounds : LExprT String := - .mdata (Info.mk "test_info") - (.const "42" (.bounded (.band (.ble (.bconst 0) (.bvar)) - (.blt (.bvar) (.bconst 100))))) - -#eval eraseTys (boundedWfConditions testMetadataWithBounds) -#eval eraseTy (translateBounded' testMetadataWithBounds) - --- Test 8: Chain of bounded function applications --- Input: f3 (f2 (f1 5 : {x < 10}) : {x < 20}) : {x < 30} --- Expected: translate = f3 (f2 (f1 5)), wfCond = [f1 5 < 10 → f2 (f1 5) < 20 -> f2 (f1 5) < 20, f1 5 < 10 -> f1 5 < 10] -def testBoundedChain : LExprT String := - .app (.op "f3" (.arrow (.bounded (.blt (.bvar) (.bconst 20))) - (.bounded (.blt (.bvar) (.bconst 30))))) - (.app (.op "f2" (.arrow (.bounded (.blt (.bvar) (.bconst 10))) - (.bounded (.blt (.bvar) (.bconst 20))))) - (.app (.op "f1" (.arrow .int (.bounded (.blt (.bvar) (.bconst 10))))) - (.const "5" .int) - (.bounded (.blt (.bvar) (.bconst 10)))) - (.bounded (.blt (.bvar) (.bconst 20)))) - (.bounded (.blt (.bvar) (.bconst 30))) - -#eval eraseTys (boundedWfConditions testBoundedChain) -#eval eraseTy (translateBounded' testBoundedChain) - --- Test 9: Bounded type in boolean context with negation --- Input: ¬((getValue 10 : {0 ≤ x}) < 5) --- Expected: translate = (0 ≤ getValue 10) → ¬(getValue 10 < 5), wfCond = [] -def testBoundedInBoolContext : LExprT String := - .app (LFunc.opExprT boolNotFunc) - (.app (.app (LFunc.opExprT intLtFunc) - (.app (.op "getValue" (.arrow .int (.bounded (.ble (.bconst 0) (.bvar))))) - (.const "10" .int) - (.bounded (.ble (.bconst 0) (.bvar)))) - (.arrow .int .bool)) - (.const "5" .int) - .bool) - .bool - -#eval eraseTys (boundedWfConditions testBoundedInBoolContext) -#eval eraseTy (translateBounded' testBoundedInBoolContext) - -end OtherTest +def removeLeadingForall (T : TEnv Identifier) (e: LExprT Identifier) : Except Format (LExprT Identifier × TEnv Identifier) := + match e with + | .quant .all ty _ e => do + let (xv, xty, T) ← Lambda.LExprT.typeBoundVar T ty; + .ok (LExprT.varOpen 0 (xv, xty) e, T) + | _ => .ok (e, T) + +def removeLeadingForalls (T : TEnv Identifier) (l: List (LExprT Identifier)) : Except Format (List (LExprT Identifier) × TEnv Identifier) := + List.foldlM (fun (l, T) e => + do + let (e, T') ← removeLeadingForall T e; + .ok (e :: l, T')) ([], T) l + +-- Functional version with extra quantifiers +def boundedWfConditions' [Coe String Identifier] (e: LExprT Identifier) : List (LExprT Identifier) := + (translateBounded e []).wfCond -end Test +-- Stateful version without extra quantifiers +def boundedWfConditions [Coe String Identifier] (T : TEnv Identifier) (e: LExprT Identifier) : Except Format (List (LExprT Identifier) × TEnv Identifier) := removeLeadingForalls T (boundedWfConditions' e) +-- NOTE: the assumptions are useful: they show us the "axioms" that we depend on (assumptions about external ops and free variables) end Bounded diff --git a/Strata/DL/Bounded/BExprTest.lean b/Strata/DL/Bounded/BExprTest.lean new file mode 100644 index 00000000..b0c5a926 --- /dev/null +++ b/Strata/DL/Bounded/BExprTest.lean @@ -0,0 +1,738 @@ +import Strata.DL.Bounded.BExpr +import Strata.DL.Lambda.LExprT +import Strata.DL.Lambda.LTy + +-- Tests + +namespace Test +open Std (ToFormat Format format) +open Lambda +open Bounded + +-- NOTE: with a semantics for LExpr/LExprT, we could prove the equivalences mentioned above + +def natTy : LMonoTy := LMonoTy.bounded (.ble (.bconst 0) .bvar) +def leOp (e1 e2: LExprT String) : LExprT String := .app (.app (LFunc.opExprT intLeFunc) e1 (.arrow .int .bool)) e2 .bool + +def geOp (e1 e2: LExprT String) : LExprT String := .app (.app (LFunc.opExprT intGeFunc) e1 (.arrow .int .bool)) e2 .bool + +def addOp (e1 e2: LExprT String) : LExprT String := .app (.app (LFunc.opExprT intAddFunc) e1 (.arrow .int .int)) e2 .int + +def mulOp (e1 e2: LExprT String) : LExprT String := .app (.app (LFunc.opExprT intMulFunc) e1 (.arrow .int .int)) e2 .int + +def notOp (e: LExprT String) : LExprT String := .app (LFunc.opExprT boolNotFunc) e .bool + +-- easier to read +def eraseTy (x: LExprT String) := + LExpr.eraseTypes (LExprT.toLExpr x) + +def eraseTys l := List.map eraseTy l + +namespace TranslateTest + +-- 1. ∀ (x: Nat), 0 <= x (quantified assumption) +-- Expected: ∀ (x: int), 0 <= x -> 0 <= x + +def test1 := (@LExprT.quant String .all natTy (.bvar 0 natTy) (leOp (.const "0" .int) (.bvar 0 .int))) + +/-- info: Lambda.LExpr.quant + (Lambda.QuantifierKind.all) + (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.bvar 0) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.bvar 0))) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.bvar 0))) + -/ +#guard_msgs in +#eval (eraseTy (translateBounded' test1) ) + +-- 2. λ(x: Nat), if 0 <= x then 1 else 2 (assumption inside ite) +-- Expected: λ (x: int), if 0 <= x -> 0 <= x then 1 else 2 + +def test2 : LExprT String := .abs (.ite (leOp (.const "0" .int) (.bvar 0 .int)) (.const "1" .int) (.const "2" .int) .int) (.arrow natTy .int) + +/-- info: Lambda.LExpr.abs + (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.ite + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.bvar 0))) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.bvar 0))) + (Lambda.LExpr.const "1" none) + (Lambda.LExpr.const "2" none))-/ +#guard_msgs in +#eval (eraseTy (translateBounded' test2)) + +-- 3. λ(x: int), if foo x >= 0 then 1 else 0 (for foo: int -> Nat) (propagate op/application information) +-- Expected: λ (x: int), if 0 <= foo x -> foo x >= 0 then 1 else 0 + +def test3 : LExprT String := .abs (.ite (geOp (.app (.op "foo" (.arrow .int natTy)) (.bvar 0 .int) natTy) (.const "0" .int)) (.const "1" .int) (.const "0" .int) .int) (.arrow .int .int) + +/-- info: Lambda.LExpr.abs + (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.ite + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.bvar 0)))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Int.Ge" none) + (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.bvar 0))) + (Lambda.LExpr.const "0" none))) + (Lambda.LExpr.const "1" none) + (Lambda.LExpr.const "0" none))-/ +#guard_msgs in +#eval (eraseTy (translateBounded' test3)) + +-- 4. (x: Nat) + (y: Nat) >= 0 (free variable bounds) +-- Expected: 0 <= (x: int) -> 0 <= (y : int) -> x + y >= 0 + +def test4 : LExprT String := geOp (addOp (.fvar "x" natTy) (.fvar "y" natTy)) (.const "0" .int) + +/-- info: Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.fvar "x" none))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.fvar "y" none))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Int.Ge" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Add" none) (Lambda.LExpr.fvar "x" none)) + (Lambda.LExpr.fvar "y" none))) + (Lambda.LExpr.const "0" none)))-/ +#guard_msgs in +#eval eraseTy (translateBounded' test4) + +-- 5. λ (x: Nat). λ (y: Nat). x + y >= 0 (multiple bound vars) +-- Expected: λ (x: int). λ (y: int). 0 <= y -> 0 <= x -> x + y >= 0 + +def test5 : LExprT String := .abs (.abs (geOp (addOp (.bvar 1 .int) (.bvar 0 .int)) (.const "0" .int)) (.arrow natTy .bool)) (.arrow natTy (.arrow natTy .bool)) + +/-- info: Lambda.LExpr.abs + (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.abs + (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.bvar 0))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.bvar 1))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Int.Ge" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Add" none) (Lambda.LExpr.bvar 1)) + (Lambda.LExpr.bvar 0))) + (Lambda.LExpr.const "0" none)))))-/ +#guard_msgs in +#eval eraseTy (translateBounded' test5) + +-- 6. λ (x: Nat), if foo then λ (y: Nat). not (x = -1) else λ (y: Nat). x + y >= 0 (propagate inside branches of if-then-else) +--Expected: λ (x: int), if 0 <= x -> foo then λ (y: int), 0 <= y -> 0 <= x -> not (x = -1) else λ (y: int). 0 <= y -> 0 <= x -> x + y >= 0 + +def test6 : LExprT String := .abs (.ite (.op "foo" .bool) (.abs (notOp (.eq (.bvar 1 .int) (.const "-1" .int) .bool)) (.arrow natTy .bool)) (.abs (geOp (addOp (.bvar 1 .int) (.bvar 0 .int)) (.const "0" .int)) (.arrow natTy .bool)) (.arrow natTy .bool)) (.arrow natTy (.arrow natTy .bool)) + +/-- info: Lambda.LExpr.abs + (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.ite + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.bvar 0))) + (Lambda.LExpr.op "foo" none)) + (Lambda.LExpr.abs + (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.bvar 0))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.bvar 1))) + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Not" none) + (Lambda.LExpr.eq (Lambda.LExpr.bvar 1) (Lambda.LExpr.const "-1" none)))))) + (Lambda.LExpr.abs + (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.bvar 0))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.bvar 1))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Int.Ge" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Add" none) (Lambda.LExpr.bvar 1)) + (Lambda.LExpr.bvar 0))) + (Lambda.LExpr.const "0" none))))))-/ +#guard_msgs in +#eval eraseTy (translateBounded' test6) + +end TranslateTest + +-- Tests for wf conditions + +-- For these tests, we initialize with an empty TEnv +def runWFTest (e: LExprT String) := do + let (l, _) ← boundedWfConditions TEnv.default e; + .ok (eraseTys l) + +namespace WFTest + +-- 1. constant: (1: Nat) +-- Expected: 0 <= 1 + +def test1 : LExprT String := .const "1" natTy + +/--info: ok: [Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.const "1" none)] + -/ +#guard_msgs in +#eval runWFTest test1 + +-- 2. application: (λ x: Nat. x) 1 +-- Expected: 0 <= 1 + +def test2 : LExprT String := .app (.abs (.bvar 0 .int) (.arrow natTy .int)) (.const "1" .int) .int + +/-- info: ok: [Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.const "1" none)]-/ +#guard_msgs in +#eval runWFTest test2 + +-- 3. application with assumption (bottom up): (λ x: Nat. x) (foo : Nat) +-- Expected: 0 <= foo -> 0 <= foo + +def test3 : LExprT String := .app (.abs (.bvar 0 .int) (.arrow natTy .int)) (.op "foo" natTy) .int + +/-- info: ok: [Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.op "foo" none))) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.op "foo" none))]-/ +#guard_msgs in +#eval runWFTest test3 + +-- 4. application with assumption (top down): (λ x: Nat. (λ y: Nat. y) x) +-- Expected: 0 <= x -> 0 <= x (no quantifiers) + +def test4 : LExprT String := .abs (.app (.abs (.bvar 0 .int) (.arrow natTy .int)) (.bvar 0 .int) .int) (.arrow natTy .int) + +/-- info: ok: [Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.fvar "$__var0" none))) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.fvar "$__var0" none))]-/ +#guard_msgs in +#eval runWFTest test4 + +-- 5. abstraction with assumption: (λ x: Nat. foo (x + 1)) (foo: Nat -> int) +-- Expected: 0 <= x -> 0 <= x + 1 + +def test5 : LExprT String := .abs (.app (.op "foo" (.arrow natTy .int)) (addOp (.bvar 0 .int) (.const "1" .int)) .int) (.arrow natTy .int) + +/--info: ok: [Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.fvar "$__var0" none))) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Add" none) (Lambda.LExpr.fvar "$__var0" none)) + (Lambda.LExpr.const "1" none)))]-/ +#guard_msgs in +#eval runWFTest test5 + +-- 6. quantified assumption: (∃ x: Nat. foo x) where (foo: Nat -> int) +-- Expected: 0 <= x -> 0 <= x + +def test6 : LExprT String := .quant .exist natTy (.bvar 0 .int) (.app (.op "foo" (.arrow natTy .int)) (.bvar 0 .int) .int) + +/-- info: ok: [Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.fvar "$__var0" none))) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.fvar "$__var0" none))]-/ +#guard_msgs in +#eval runWFTest test6 + +-- 7. Lambda with bounded body (λ x: Nat. (x + 1: Nat)) +-- Expected: 0 <= x -> 0 <= x + 1 + +def test7 : LExprT String := .abs (addOp (.bvar 0 .int) (.const "1" .int)) (.arrow natTy natTy) + +/-- info: ok: [Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.fvar "$__var0" none))) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Add" none) (Lambda.LExpr.fvar "$__var0" none)) + (Lambda.LExpr.const "1" none)))]-/ +#guard_msgs in +#eval runWFTest test7 + + +-- 8. Application with bounded body: (foo x) : Nat where foo: int -> Nat +-- Expected: [] (foo assumed) + +def test8 : LExprT String := .app (.op "foo" (.arrow .int natTy)) (.fvar "x" .int) natTy + +/-- info: ok: []-/ +#guard_msgs in +#eval runWFTest test8 + +-- 9. Application with bounded body, no assumption: (λ (x: int) . x * x) 1 : Nat +-- Expected: 0 <= (λ (x: int) . x * x) 1 + +def test9 : LExprT String := .app (.abs (mulOp (.bvar 0 .int) (.bvar 0 .int)) (.arrow .int .int)) (.const "1" .int) natTy + +/-- info: ok: [Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.app + (Lambda.LExpr.abs + (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Mul" none) (Lambda.LExpr.bvar 0)) + (Lambda.LExpr.bvar 0))) + (Lambda.LExpr.const "1" none))]-/ +#guard_msgs in +#eval runWFTest test9 + +-- 10. If-then-else with bounded body: if b then 1 else 0 : Nat +-- Expected: 0 <= if b then 1 else 0 + +def test10 : LExprT String := .ite (.const "b" .bool) (.const "1" .int) (.const "0" .int) natTy + +/-- info: ok: [Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.ite (Lambda.LExpr.const "b" none) (Lambda.LExpr.const "1" none) (Lambda.LExpr.const "0" none))]-/ +#guard_msgs in +#eval runWFTest test10 + +-- 11. If-then-else with bound functions: if b then λ (x: int). x * x : Nat else λ (y: int). 0 : Nat (whole type int -> Nat) +-- Expected: 0 <= x * x; 0 <= 0 + +def test11 : LExprT String := .ite (.const "b" .bool) (.abs (mulOp (.bvar 0 .int) (.bvar 0 .int)) (.arrow .int natTy)) (.abs (.const "0" .int) (.arrow .int natTy)) (.arrow .int natTy) + +/-- info: ok: [Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Mul" none) (Lambda.LExpr.fvar "$__var1" none)) + (Lambda.LExpr.fvar "$__var1" none)), Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.const "0" none)]-/ +#guard_msgs in +#eval runWFTest test11 + +end WFTest + +--tests with more sophisticated bounded types (mostly AI generated) + +namespace OtherTest + +-- Test 1: Nested bounded function applications +-- Input: add : {x < 10} → {y < 5} → {z < 15}, applied to (3 : {x < 10}) and (2 : {x < 5}) +-- Expected: translate = add 3 2, wfCond = [3 < 10, 2 < 5] + +def testNestedBoundedApps : LExprT String := + .app (.app (.op "add" (.arrow (.bounded (.blt (.bvar) (.bconst 10))) + (.arrow (.bounded (.blt (.bvar) (.bconst 5))) + (.bounded (.blt (.bvar) (.bconst 15)))))) + (.const "3" (.bounded (.blt (.bvar) (.bconst 10)))) + (.arrow (.bounded (.blt (.bvar) (.bconst 5))) (.bounded (.blt (.bvar) (.bconst 15))))) + (.const "2" (.bounded (.blt (.bvar) (.bconst 5)))) + (.bounded (.blt (.bvar) (.bconst 15))) + +/-- info: ok: [Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Lt" none) (Lambda.LExpr.const "2" none)) + (Lambda.LExpr.const "5" none), Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Lt" none) (Lambda.LExpr.const "3" none)) + (Lambda.LExpr.const "10" none)]-/ +#guard_msgs in +#eval runWFTest testNestedBoundedApps + +/-- info: Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "add" none) (Lambda.LExpr.const "3" none)) + (Lambda.LExpr.const "2" none)-/ +#guard_msgs in +#eval eraseTy (translateBounded' testNestedBoundedApps) + +-- Test 2: Bounded variable in quantifier with complex bound expression +-- Input: ∀ (x : {x < 100 ∧ 0 ≤ x}). x = 42 +-- Expected: translate = ∀ x:int. (x < 100 ∧ 0 ≤ x) → (x = 42), wfCond = [] + +def testComplexBoundInQuantifier : LExprT String := + .quant .all (.bounded (.band (.blt (.bvar) (.bconst 100)) + (.ble (.bconst 0) (.bvar)))) + (.bvar 0 (.bounded (.band (.blt (.bvar) (.bconst 100)) + (.ble (.bconst 0) (.bvar))))) + (.eq (.bvar 0 (.bounded (.band (.blt (.bvar) (.bconst 100)) + (.ble (.bconst 0) (.bvar))))) + (.const "42" .int) .bool) + +/-- info: ok: []-/ +#guard_msgs in +#eval runWFTest testComplexBoundInQuantifier + +/-- info: Lambda.LExpr.quant + (Lambda.QuantifierKind.all) + (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.bvar 0) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.And" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Lt" none) (Lambda.LExpr.bvar 0)) + (Lambda.LExpr.const "100" none))) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.bvar 0)))) + (Lambda.LExpr.eq (Lambda.LExpr.bvar 0) (Lambda.LExpr.const "42" none))) +-/ +#guard_msgs in +#eval eraseTy (translateBounded' testComplexBoundInQuantifier) + +-- Test 3: If-then-else with bounded branches and boolean condition +-- Input: if (0 < (getValue 5 : {0 ≤ x})) then (1 : {x < 10}) else (0 : {x < 10}) : {x < 10} +-- Expected: translate = if (0 ≤ getValue 5) → (0 < getValue 5) then 1 else 0, wfCond = [1 < 10, 0 < 10, (if (getValue 5) then 1 else 0) < 10] +def testBoundedIte : LExprT String := + .ite (.app (.app (LFunc.opExprT intLtFunc) + (.const "0" .int) + (.arrow .int .bool)) + (.app (.op "getValue" (.arrow .int (.bounded (.ble (.bconst 0) (.bvar))))) + (.const "5" .int) + (.bounded (.ble (.bconst 0) (.bvar)))) + .bool) + (.const "1" (.bounded (.blt (.bvar) (.bconst 10)))) + (.const "0" (.bounded (.blt (.bvar) (.bconst 10)))) + (.bounded (.blt (.bvar) (.bconst 10))) + +/-- info: ok: [Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Int.Lt" none) + (Lambda.LExpr.ite + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Lt" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.app (Lambda.LExpr.op "getValue" none) (Lambda.LExpr.const "5" none))) + (Lambda.LExpr.const "1" none) + (Lambda.LExpr.const "0" none))) + (Lambda.LExpr.const "10" none), Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Lt" none) (Lambda.LExpr.const "1" none)) + (Lambda.LExpr.const "10" none), Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Lt" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.const "10" none)]-/ +#guard_msgs in +#eval runWFTest testBoundedIte + +/-- info: Lambda.LExpr.ite + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.app (Lambda.LExpr.op "getValue" none) (Lambda.LExpr.const "5" none)))) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Lt" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.app (Lambda.LExpr.op "getValue" none) (Lambda.LExpr.const "5" none)))) + (Lambda.LExpr.const "1" none) + (Lambda.LExpr.const "0" none)-/ +#guard_msgs in +#eval eraseTy (translateBounded' testBoundedIte) + +-- Test 4: Lambda with bounded parameter and bounded return type +-- Input: λ (x : {0 ≤ x}). increment x : {y < 100} +-- Expected: translate = λ x:int. increment x, wfCond = [0 ≤ x → increment x < 100; 0 <= x -> 0 <= x] +def testBoundedLambda : LExprT String := + .abs (.app (.op "increment" (.arrow (.bounded (.ble (.bconst 0) (.bvar))) + (.bounded (.blt (.bvar) (.bconst 100))))) + (.bvar 0 (.bounded (.ble (.bconst 0) (.bvar)))) + (.bounded (.blt (.bvar) (.bconst 100)))) + (.arrow (.bounded (.ble (.bconst 0) (.bvar))) + (.bounded (.blt (.bvar) (.bconst 100)))) + +/-- info: ok: [Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.fvar "$__var1" none))) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.fvar "$__var1" none)), Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.fvar "$__var0" none))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Int.Lt" none) + (Lambda.LExpr.app (Lambda.LExpr.op "increment" none) (Lambda.LExpr.fvar "$__var0" none))) + (Lambda.LExpr.const "100" none))]-/ +#guard_msgs in +#eval runWFTest testBoundedLambda + +/-- info: Lambda.LExpr.abs + (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.app (Lambda.LExpr.op "increment" none) (Lambda.LExpr.bvar 0))-/ +#guard_msgs in +#eval eraseTy (translateBounded' testBoundedLambda) + +-- Test 5: Equality between bounded expressions +-- Input: (square (3 : {-10 ≤ x ≤ 10}) : {0 ≤ y}) = (9 : {0 ≤ z}) +-- Expected: translate = 0 <= square 3 -> square 3 = 9, wfCond = [-10 ≤ 3 ≤ 10, 0 ≤ 9] +def testBoundedEquality : LExprT String := + .eq (.app (.op "square" (.arrow (.bounded (.band (.ble (.bconst (-10)) (.bvar)) + (.ble (.bvar) (.bconst 10)))) + (.bounded (.ble (.bconst 0) (.bvar))))) + (.const "3" (.bounded (.band (.ble (.bconst (-10)) (.bvar)) + (.ble (.bvar) (.bconst 10))))) + (.bounded (.ble (.bconst 0) (.bvar)))) + (.const "9" (.bounded (.ble (.bconst 0) (.bvar)))) + .bool + +/-- info: ok: [Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.And" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "-10" none)) + (Lambda.LExpr.const "3" none))) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "3" none)) + (Lambda.LExpr.const "10" none)), Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.const "9" none)]-/ +#guard_msgs in +#eval runWFTest testBoundedEquality + +/-- info: Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.app (Lambda.LExpr.op "square" none) (Lambda.LExpr.const "3" none)))) + (Lambda.LExpr.eq + (Lambda.LExpr.app (Lambda.LExpr.op "square" none) (Lambda.LExpr.const "3" none)) + (Lambda.LExpr.const "9" none))-/ +#guard_msgs in +#eval eraseTy (translateBounded' testBoundedEquality) + +-- Test 6: Free variable with bounded type in assumptions +-- Input: compare (x : {x < 50}) 25, with assumption [x < 50] +-- Expected: translate = (x < 50) → compare x 25, wfCond = [] +def testFreeVarWithAssumptions : LExprT String := + .app (.app (.op "compare" (.arrow .int (.arrow .int .bool))) + (.fvar "x" (.bounded (.blt (.bvar) (.bconst 50)))) + (.arrow .int .bool)) + (.const "25" .int) + .bool + +/-- info: ok: []-/ +#guard_msgs in +#eval runWFTest testFreeVarWithAssumptions + +/-- info: Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Lt" none) (Lambda.LExpr.fvar "x" none)) + (Lambda.LExpr.const "50" none))) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "compare" none) (Lambda.LExpr.fvar "x" none)) + (Lambda.LExpr.const "25" none))-/ +#guard_msgs in +#eval eraseTy (translateBounded' testFreeVarWithAssumptions) + +-- Test 7: Metadata preservation with bounded types +-- Input: @metadata(42 : {0 ≤ x < 100}) +-- Expected: translate = @metadata(42), wfCond = [0 ≤ 42 < 100] +def testMetadataWithBounds : LExprT String := + .mdata (Info.mk "test_info") + (.const "42" (.bounded (.band (.ble (.bconst 0) (.bvar)) + (.blt (.bvar) (.bconst 100))))) + +/-- info: ok: [Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.And" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.const "42" none))) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Lt" none) (Lambda.LExpr.const "42" none)) + (Lambda.LExpr.const "100" none))]-/ +#guard_msgs in +#eval runWFTest testMetadataWithBounds + +/-- info: Lambda.LExpr.mdata { value := "test_info" } (Lambda.LExpr.const "42" none)-/ +#guard_msgs in +#eval eraseTy (translateBounded' testMetadataWithBounds) + +-- Test 8: Chain of bounded function applications +-- Input: f3 (f2 (f1 5 : {x < 10}) : {x < 20}) : {x < 30} +-- Expected: translate = f3 (f2 (f1 5)), wfCond = [f2 (f1 5) < 20 -> f1 5 < 10 → f2 (f1 5) < 20, f1 5 < 10 -> f1 5 < 10] +def testBoundedChain : LExprT String := + .app (.op "f3" (.arrow (.bounded (.blt (.bvar) (.bconst 20))) + (.bounded (.blt (.bvar) (.bconst 30))))) + (.app (.op "f2" (.arrow (.bounded (.blt (.bvar) (.bconst 10))) + (.bounded (.blt (.bvar) (.bconst 20))))) + (.app (.op "f1" (.arrow .int (.bounded (.blt (.bvar) (.bconst 10))))) + (.const "5" .int) + (.bounded (.blt (.bvar) (.bconst 10)))) + (.bounded (.blt (.bvar) (.bconst 20)))) + (.bounded (.blt (.bvar) (.bconst 30))) + +/--info: ok: [Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Int.Lt" none) + (Lambda.LExpr.app + (Lambda.LExpr.op "f2" none) + (Lambda.LExpr.app (Lambda.LExpr.op "f1" none) (Lambda.LExpr.const "5" none)))) + (Lambda.LExpr.const "20" none))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Int.Lt" none) + (Lambda.LExpr.app (Lambda.LExpr.op "f1" none) (Lambda.LExpr.const "5" none))) + (Lambda.LExpr.const "10" none))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Int.Lt" none) + (Lambda.LExpr.app + (Lambda.LExpr.op "f2" none) + (Lambda.LExpr.app (Lambda.LExpr.op "f1" none) (Lambda.LExpr.const "5" none)))) + (Lambda.LExpr.const "20" none))), Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Int.Lt" none) + (Lambda.LExpr.app (Lambda.LExpr.op "f1" none) (Lambda.LExpr.const "5" none))) + (Lambda.LExpr.const "10" none))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Int.Lt" none) + (Lambda.LExpr.app (Lambda.LExpr.op "f1" none) (Lambda.LExpr.const "5" none))) + (Lambda.LExpr.const "10" none))]-/ +#guard_msgs in +#eval runWFTest testBoundedChain + +/-- info: Lambda.LExpr.app + (Lambda.LExpr.op "f3" none) + (Lambda.LExpr.app + (Lambda.LExpr.op "f2" none) + (Lambda.LExpr.app (Lambda.LExpr.op "f1" none) (Lambda.LExpr.const "5" none)))-/ +#guard_msgs in +#eval eraseTy (translateBounded' testBoundedChain) + +-- Test 9: Bounded type in boolean context with negation +-- Input: ¬((getValue 10 : {0 ≤ x}) < 5) +-- Expected: translate = (0 ≤ getValue 10) → ¬(getValue 10 < 5), wfCond = [] +def testBoundedInBoolContext : LExprT String := + .app (LFunc.opExprT boolNotFunc) + (.app (.app (LFunc.opExprT intLtFunc) + (.app (.op "getValue" (.arrow .int (.bounded (.ble (.bconst 0) (.bvar))))) + (.const "10" .int) + (.bounded (.ble (.bconst 0) (.bvar)))) + (.arrow .int .bool)) + (.const "5" .int) + .bool) + .bool + +def foo : LExprT String := (.app (.app (LFunc.opExprT intLtFunc) + (.app (.op "getValue" (.arrow .int (.bounded (.ble (.bconst 0) (.bvar))))) + (.const "10" .int) + (.bounded (.ble (.bconst 0) (.bvar)))) + (.arrow .int .bool)) + (.const "5" .int) + .bool) + +#eval (eraseTys (translateBounded foo []).assume) + +-- NOTE: all bottom-up assumptions need to be kept, even in boolean-valued terms, since they have to go to the top level + +--ex: not (foo x < 0) for foo: int -> nat +-- want (I think): 0 <= foo x -> not (foo x < 0) - TRUE +-- we do NOT want: not (0 <= foo x -> foo x < 0) - this is FALSE +-- think we need to collect bottom-up assumptions differently + +/-- info: ok: []-/ +#guard_msgs in +#eval runWFTest testBoundedInBoolContext +#eval (translateBounded testBoundedInBoolContext []).assume +#eval eraseTy (translateBounded' testBoundedInBoolContext) + +end OtherTest + +end Test diff --git a/Strata/DL/Bounded/BTy.lean b/Strata/DL/Bounded/BTy.lean deleted file mode 100644 index 8a5575c1..00000000 --- a/Strata/DL/Bounded/BTy.lean +++ /dev/null @@ -1,61 +0,0 @@ -/- - Copyright Strata Contributors - - SPDX-License-Identifier: Apache-2.0 OR MIT --/ - -import Strata.DL.Lambda.LTy - -/-! ## Lambda Types with Bounded Ints --/ - --- The grammar allows expressions like a <= b /\ b < c, where a, b, and c --- are integer constants or a (single) variable. This lets us express e.g. --- {x | 0 <= x < 2^32} - -namespace Bounded -open Std (ToFormat Format format) - - -abbrev TyIdentifier := String - -inductive BVal : Type where - /-- The bounded variable -/ - | bvar - /-- Integer constants -/ - | bconst (val: Int) -deriving Inhabited, Repr, DecidableEq - -inductive BExpr : Type where - /-- b1 = b2 -/ - | beq (b1 b2: BVal) - /-- b1 < b2 -/ - | blt (e1 e2: BVal) - /-- b1 <= b2 -/ - | ble (e1 e2: BVal) - /-- e1 /\ e2 -/ - | band (e1 e2: BExpr) - /-- not e -/ - | bnot (e: BExpr) -deriving Inhabited, Repr, DecidableEq - -inductive BMonoTy : Type where - /-- Type variable. -/ - | ftvar (name : TyIdentifier) - /-- Type constructor. -/ - | tcons (name : String) (args : List BMonoTy) - /-- Special support for bitvector types of every size. -/ - | bitvec (size : Nat) - /-- Added to LMonoTy: bounded ints -/ - | boundint (b: BExpr) -deriving Inhabited, Repr - --- When desugaring to Lambda types, all bounded ints are erased -def BMonoTy_to_LMonoTy (b: BMonoTy) : Lambda.LMonoTy := - match b with - | .ftvar name => Lambda.LMonoTy.ftvar name - | .tcons name args => Lambda.LMonoTy.tcons name (List.map BMonoTy_to_LMonoTy args) - | .bitvec size => Lambda.LMonoTy.bitvec size - | .boundint _ => Lambda.LMonoTy.int - -end Bounded From 8561549dd06a4e7c5635c8f661113e15f2e41a46 Mon Sep 17 00:00:00 2001 From: Josh Cohen Date: Mon, 29 Sep 2025 13:04:54 -0400 Subject: [PATCH 08/19] Fix bug in if-then-else, note about positivity --- Strata/DL/Bounded/BExpr.lean | 17 +++++++++++++++-- Strata/DL/Bounded/BExprTest.lean | 10 ++++++++-- 2 files changed, 23 insertions(+), 4 deletions(-) diff --git a/Strata/DL/Bounded/BExpr.lean b/Strata/DL/Bounded/BExpr.lean index 40d75c71..4655737f 100644 --- a/Strata/DL/Bounded/BExpr.lean +++ b/Strata/DL/Bounded/BExpr.lean @@ -173,8 +173,21 @@ Note that we do NOT have to worry about the fact that the new functions have a l Invariant: assumptions must not have bounded types (TODO enforce this), same for inputs Invariant (I think): All assumptions are of form: b(bvar #x) + +There is one more complication. When propagating bottom-up assumptions, we cannot just add them at the first bool expression encountered. To see why, consider the example: +~ ((x: Nat) < 0) +This must become +0 <= x -> ~ (x < 0) +NOT +~ (0 <= x -> x < 0) +basically, this should occur any time we are on the left of an implication. Thus, we must add assumptions at any top-level boolean expression. To keep track of this, we use the parameter inBool which becomes false once we have passed at least one bool-valued expression. We only add assumptions if inBool is + +We need to keep track of positivity. We cannot add assumptions (e.g) on the LHS of an implication; we need to propagate a level above. This is annoying because connectives are not built in to Lambda. + -/ -def translateBounded [Coe String Identifier] [DecidableEq Identifier] (e: LExprT Identifier) (assume: List (LExprT Identifier)) : translationRes Identifier := +def translateBounded [Coe String Identifier] [DecidableEq Identifier] (e: LExprT Identifier) (assume: List (LExprT Identifier)) +--(pos : Bool) +: translationRes Identifier := match e with -- constants do not need assumptions; they produce a wf goal if bounded | .const c ty => @@ -279,7 +292,7 @@ def translateBounded [Coe String Identifier] [DecidableEq Identifier] (e: LExprT let t' := translateBounded t assume; let f' := translateBounded f assume; let e' := .ite (addAssumptions (assume ++ c'.assume) c'.translate) t'.translate f'.translate ty; - ⟨e', ListSet.union [boundExprIfType ty e, c'.wfCond, t'.wfCond, f'.wfCond], t'.assume ++ f'.assume⟩ + ⟨e', ListSet.union [boundExprIfType ty e', c'.wfCond, t'.wfCond, f'.wfCond], t'.assume ++ f'.assume⟩ -- Equality is always bool-valued, so we can add assumptions | .eq e1 e2 ty => let e1' := translateBounded e1 []; diff --git a/Strata/DL/Bounded/BExprTest.lean b/Strata/DL/Bounded/BExprTest.lean index b0c5a926..0a0b8608 100644 --- a/Strata/DL/Bounded/BExprTest.lean +++ b/Strata/DL/Bounded/BExprTest.lean @@ -479,8 +479,14 @@ def testBoundedIte : LExprT String := (Lambda.LExpr.op "Int.Lt" none) (Lambda.LExpr.ite (Lambda.LExpr.app - (Lambda.LExpr.app (Lambda.LExpr.op "Int.Lt" none) (Lambda.LExpr.const "0" none)) - (Lambda.LExpr.app (Lambda.LExpr.op "getValue" none) (Lambda.LExpr.const "5" none))) + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.app (Lambda.LExpr.op "getValue" none) (Lambda.LExpr.const "5" none)))) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Lt" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.app (Lambda.LExpr.op "getValue" none) (Lambda.LExpr.const "5" none)))) (Lambda.LExpr.const "1" none) (Lambda.LExpr.const "0" none))) (Lambda.LExpr.const "10" none), Lambda.LExpr.app From 43b8d5f4d59283c2a27f41a8b1be5d2d84b8e7b2 Mon Sep 17 00:00:00 2001 From: Josh Cohen Date: Tue, 30 Sep 2025 17:30:24 -0400 Subject: [PATCH 09/19] Fix positivity issue, add tests --- Strata/DL/Bounded/BExpr.lean | 180 ++++++++++---- Strata/DL/Bounded/BExprTest.lean | 394 ++++++++++++++++++++++++++++--- 2 files changed, 502 insertions(+), 72 deletions(-) diff --git a/Strata/DL/Bounded/BExpr.lean b/Strata/DL/Bounded/BExpr.lean index 4655737f..dd7c1e74 100644 --- a/Strata/DL/Bounded/BExpr.lean +++ b/Strata/DL/Bounded/BExpr.lean @@ -137,7 +137,9 @@ structure translationRes Identifier where def boundExprIfType [Coe String Identifier] (ty: LMonoTy) (e: LExprT Identifier) : List (LExprT Identifier) := ((isBounded ty).map (fun b => BoundExprToLExprT b e)).toList --- Generate WF condition for calling e1 with argument e2 if bounded type expected +/-- +Generate WF condition for calling e1 with argument e2 if bounded type expected +-/ def wfCallCondition [Coe String Identifier] (assume : List (LExprT Identifier)) (e1 e2: LExprT Identifier) := match LExprT.toLMonoTy e1 with | .arrow (.bounded b) _ => @@ -152,6 +154,54 @@ def addBoundedWf [Coe String Identifier] (assume: List (LExprT Identifier)) (l1 def addBoundedWfAssume [Coe String Identifier] (assume: List (LExprT Identifier)) (l: List (LExprT Identifier)) := addBoundedWf assume l [] +/-- +Determines if a term uses a particular bound variable at the top level +-/ +def hasBvar (e: LExprT Identifier) (n: Nat) : Bool := + match e with + | .bvar m _ => n == m + | .app e1 e2 _ => hasBvar e1 n || hasBvar e2 n + | .abs e _ => hasBvar e (n + 1) + | .quant _ _ _ e => hasBvar e (n + 1) + | .eq e1 e2 _ => hasBvar e1 n || hasBvar e2 n + | .mdata _ e => hasBvar e n + | .ite e1 e2 e3 _ => hasBvar e1 n || hasBvar e2 n || hasBvar e3 n + | _ => false + +/-- +When passing an assumption through a binder, we need to quantify the bound variable if used. +-/ +def quantifyAssumptions (ty: LMonoTy) (e: LExprT Identifier): LExprT Identifier := + --TODO: triggers? + if hasBvar e 0 then .quant .all ty (.bvar 0 ty) e else e + +-- We don't have pattern matching on operators (yet) because of the abstract metadata, so we resort to the follow workaround +-- def isNot [Coe String Identifier] (e: LExprT Identifier) : Option (LExprT Identifier) := +-- match e with +-- | .app (.op o _) e2 _ => +-- if o == boolNotFunc.name then some e2 else none +-- | _ => none + +-- def isImpl [Coe String Identifier] (e: LExprT Identifier) : Option (LExprT Identifier × LExprT Identifier) := +-- match e with +-- | .app (.app (.op o _) e1 _) e2 _ => +-- if o == boolImpliesFunc.name then some (e1, e2) else none +-- | _ => none + +-- Ugh, need to know how to 1. combine results to produce new term +-- 2. produce wf conditions (should it just be app Bool case) START HERE +-- def boolCase [Coe String Identifier] (assume: List (LExprT Identifier)) (l: List (translationRes Identifier)) (comb: List (LExprT Identifier) -> LExprT Identifier) (pos: Bool) := +-- let all_assumes := assume ++ List.flatten (List.map translationRes.assume l); +-- let res := (if pos then addAssumptions all_assumes else addAssumptions assume) (comb (List.map translationRes.translate l)); +-- (res, ) + + +-- let e1' := translateBounded e1 [] false; +-- let e2' := translateBounded e2 [] false; +-- let all_assumes := assume ++ e1'.assume ++ e2'.assume; +-- let res := addAssumptions all_assumes (.app e1'.translate e2'.translate .bool); +-- ⟨res, ListSet.union [wfCallCondition (assume ++ e2'.assume) e1 e2'.translate, e1'.wfCond, e2'.wfCond], []⟩ + /-- Translate expression to one without bounded ints. This is a non-trivial translation, as we want to preserve the semantics of the resulting term. @@ -184,10 +234,12 @@ basically, this should occur any time we are on the left of an implication. Thus We need to keep track of positivity. We cannot add assumptions (e.g) on the LHS of an implication; we need to propagate a level above. This is annoying because connectives are not built in to Lambda. +Idea: if positive, can add all assumptions, if not, can add top-down but propagae bottom-up. Other complication - if we propagate, look through, if uses unbound bvar 0, add quantifier at beginning (only for lambda, quantifier case) + +Need to make sure - ALWAYS safe to add top-down assumptions + -/ -def translateBounded [Coe String Identifier] [DecidableEq Identifier] (e: LExprT Identifier) (assume: List (LExprT Identifier)) ---(pos : Bool) -: translationRes Identifier := +def translateBounded [Coe String Identifier] [DecidableEq Identifier] (e: LExprT Identifier) (assume: List (LExprT Identifier)) (pos : Bool) : translationRes Identifier := match e with -- constants do not need assumptions; they produce a wf goal if bounded | .const c ty => @@ -208,17 +260,56 @@ def translateBounded [Coe String Identifier] [DecidableEq Identifier] (e: LExprT 1. If the entire application has boolean type, assumptions can be added 2. Otherwise, if the application has bounded type, it produces a bottom-up assumption. There is a subtle case. If the function has type (t -> int), then we must generate a wf condition. Otherwise, the bound will be assumed (for external operators) or checked (for abstraction/if-then-else/etc) 3. In either case, we produce a wf condition if the argument should have bounded type + + TODO: factor out boolean cases - need to say: if !pos, then propagate bottom-up assumptions (do NOT add); we can still add top-down ones (these are always safe to add) otherwise, add assumptions as right now + -/ - | .app e1 e2 ty => - if isBool ty then - let e1' := translateBounded e1 []; - let e2' := translateBounded e2 []; - let all_assumes := assume ++ e1'.assume ++ e2'.assume; - let res := addAssumptions all_assumes (.app e1'.translate e2'.translate (removeBound ty)); - ⟨res, ListSet.union [wfCallCondition (assume ++ e2'.assume) e1 e2'.translate, e1'.wfCond, e2'.wfCond], []⟩ + | .app (.op o ty1) e2 .bool => + if o == boolNotFunc.name then + -- inside "not" - not positive + let e2' := translateBounded e2 [] false; + let all_assumes := if pos then (assume ++ e2'.assume) else assume; + let res := addAssumptions all_assumes (.app (.op o ty1) e2'.translate .bool); + ⟨res, [], if pos then [] else e2'.assume⟩ else - let e1' := translateBounded e1 assume; - let e2' := translateBounded e2 assume; + -- Here, safe to keep same positivity + -- let e1' := translateBounded e1 [] pos; + let e2' := translateBounded e2 [] pos; + let all_assumes := if pos then (assume ++ e2'.assume) else assume; + let res := addAssumptions all_assumes (.app (.op o (removeBound ty1)) e2'.translate .bool); + ⟨res, ListSet.union [wfCallCondition (assume ++ e2'.assume) (.op o ty1) e2'.translate, e2'.wfCond], if pos then [] else e2'.assume⟩ + | .app (.app (.op o ty1) e1 ty2) e2 .bool => + -- Case 1: First argument NOT positive, second keeps positivity + if o == boolImpliesFunc.name then + let e1' := translateBounded e1 [] false; + let e2' := translateBounded e2 [] pos; + let all_assumes := if pos then (assume ++ e1'.assume ++ e2'.assume) else assume; -- TODO: factor out + let res := addAssumptions all_assumes (.app (.app (.op o (removeBound ty1)) e1'.translate (removeBound ty2)) e2'.translate .bool); + ⟨res, ListSet.union [e1'.wfCond, e2'.wfCond], if pos then [] else e1'.assume ++ e2'.assume⟩ + -- Case 2: Both arguments keep same positivity + else if o == boolAndFunc.name || o == boolOrFunc.name then + let e1' := translateBounded e1 [] pos; + let e2' := translateBounded e2 [] pos; + let all_assumes := if pos then (assume ++ e1'.assume ++ e2'.assume) else assume; + let res := addAssumptions all_assumes (.app (.app (.op o (removeBound ty1)) e1'.translate (removeBound ty2)) e2'.translate .bool); + ⟨res, ListSet.union [e1'.wfCond, e2'.wfCond], if pos then [] else e1'.assume ++ e2'.assume⟩ + -- Case 3: Both arguments cannot be assumed positive (equality and everything else - need because someone could define their own implication/equivalence; we know nothing about meaning) + else + let e1' := translateBounded e1 [] false; + let e2' := translateBounded e2 [] false; + let all_assumes := if pos then (assume ++ e1'.assume ++ e2'.assume) else assume; + let res := addAssumptions all_assumes (.app (.app (.op o (removeBound ty1)) e1'.translate (removeBound ty2)) e2'.translate .bool); + ⟨res, ListSet.union [wfCallCondition (assume ++ e2'.assume) e1 e2'.translate, e1'.wfCond, e2'.wfCond], if pos then [] else e1'.assume ++ e2'.assume⟩ + | .app e1 e2 .bool => + --Anything else we cannot assume is positive + let e1' := translateBounded e1 [] false; + let e2' := translateBounded e2 [] false; + let all_assumes := if pos then (assume ++ e1'.assume ++ e2'.assume) else assume; + let res := addAssumptions all_assumes (.app e1'.translate e2'.translate .bool); + ⟨res, ListSet.union [wfCallCondition (assume ++ e2'.assume) e1 e2'.translate, e1'.wfCond, e2'.wfCond], if pos then [] else e1'.assume ++ e2'.assume⟩ + | .app e1 e2 ty => + let e1' := translateBounded e1 assume pos; + let e2' := translateBounded e2 assume pos; let e' := .app e1'.translate e2'.translate (removeBound ty); let extraWf := match LExprT.toLMonoTy e1, ty with @@ -240,17 +331,17 @@ def translateBounded [Coe String Identifier] [DecidableEq Identifier] (e: LExprT [BoundExprToLExprT b (.bvar 0 .int)] | _ => []; match ty with - | .arrow _ .bool => - let e' := translateBounded e []; - let all_assume := newAssumption ++ (increaseBVars assume) ++ e'.assume; + | .arrow ty1 .bool => + let e' := translateBounded e [] pos; + let all_assume := newAssumption ++ (increaseBVars assume) ++ (if pos then e'.assume else []); -- TODO: change - ⟨.abs (addAssumptions all_assume e'.translate) (removeBound ty), addBoundedWfAssume all_assume e'.wfCond, []⟩ - | .arrow _ ty1 => + ⟨.abs (addAssumptions all_assume e'.translate) (removeBound ty), addBoundedWfAssume all_assume e'.wfCond, if pos then [] else List.map (quantifyAssumptions ty1) e'.assume⟩ + | .arrow ty1 ty2 => let all_assume := newAssumption ++ (increaseBVars assume); - let e' := translateBounded e all_assume; + let e' := translateBounded e all_assume pos; let e'' := .abs e'.translate (removeBound ty); -- Note: don't add assumptions to e'.wfCond, already included - ⟨e'', addBoundedWf all_assume (boundExprIfType ty1 e'.translate) e'.wfCond, e'.assume⟩ + ⟨e'', addBoundedWf all_assume (boundExprIfType ty2 e'.translate) e'.wfCond, List.map (quantifyAssumptions ty1) e'.assume⟩ | _ => ⟨.const "0" .int, [], []⟩ -- error case /- Quantifiers are simpler because they are boolean-valued. ∀ (x : nat). e adds an assumption x >= 0 -> ..., while ∃ (x: nat). e results in x >= 0 ∧ .. @@ -261,50 +352,55 @@ def translateBounded [Coe String Identifier] [DecidableEq Identifier] (e: LExprT | .bounded b => [BoundExprToLExprT b (.bvar 0 .int)] | _ => []; - let e' := translateBounded e []; - let tr' := translateBounded tr []; --TODO: need "clean" one here - let all_assume := (newAssumption ++ (increaseBVars assume) ++ e'.assume); + let e' := translateBounded e [] pos; + let tr' := translateBounded tr [] pos; --TODO: need "clean" one here + let all_assume := (newAssumption ++ (increaseBVars assume) ++ (if pos then e'.assume else [])); -- TODO: factor out quant part - ⟨.quant .all (removeBound ty) tr'.translate (addAssumptions all_assume e'.translate), addBoundedWfAssume all_assume e'.wfCond, []⟩ + ⟨.quant .all (removeBound ty) tr'.translate (addAssumptions all_assume e'.translate), addBoundedWfAssume all_assume e'.wfCond, if pos then [] else List.map (quantifyAssumptions ty) e'.assume⟩ | .quant .exist ty tr e => let newAssumption := match ty with | .bounded b => some (BoundExprToLExprT b (.bvar 0 .int)) | _ => none; - let e' := translateBounded e []; - let tr' := translateBounded tr []; --TODO: need "clean" one here + let e' := translateBounded e [] pos; + let tr' := translateBounded tr [] pos; --TODO: need "clean" one here let add_and x : LExprT Identifier := match newAssumption with | .some f => (.app (.app (LFunc.opExprT boolAndFunc) f (LMonoTy.arrow LMonoTy.bool LMonoTy.bool)) x LMonoTy.bool) | .none => x; - ⟨.quant .exist (removeBound ty) tr'.translate (add_and (addAssumptions ((increaseBVars assume) ++ e'.assume) e'.translate)), addBoundedWfAssume (newAssumption.toList ++ (increaseBVars assume) ++ e'.assume) e'.wfCond, []⟩ + let all_assume := (increaseBVars assume) ++ if pos then e'.assume else []; + ⟨.quant .exist (removeBound ty) tr'.translate (add_and (addAssumptions all_assume e'.translate)), addBoundedWfAssume (newAssumption.toList ++ all_assume) e'.wfCond, if pos then [] else List.map (quantifyAssumptions ty) e'.assume⟩ /- For if-then-else, we add assumptions to the condition, which is always bool-valued. For a bool-valued result, we can add the conditions freely. For a bounded-valued term, we produce a wf condition proving this. -/ | .ite c t f .bool => - let c' := translateBounded c []; - let t' := translateBounded t []; - let f' := translateBounded f []; - ⟨.ite (addAssumptions (assume ++ c'.assume) c'.translate) (addAssumptions (assume ++ t'.assume) t'.translate) (addAssumptions (assume ++ f'.assume) f'.translate) .bool, ListSet.union [c'.wfCond, t'.wfCond, f'.wfCond] ,[]⟩ + -- c is NOT positive; equivalent to c -> t /\ ~c -> f + let c' := translateBounded c [] false; + let t' := translateBounded t [] pos; + let f' := translateBounded f [] pos; + -- NOTE: VERY ugly + ⟨(if pos then addAssumptions c'.assume else id) (.ite (addAssumptions assume c'.translate) (addAssumptions (assume ++ if pos then t'.assume else []) t'.translate) (addAssumptions (assume ++ if pos then f'.assume else []) f'.translate) .bool), ListSet.union [c'.wfCond, t'.wfCond, f'.wfCond] , c'.assume ++ if pos then [] else t'.assume ++ f'.assume⟩ | .ite c t f ty => - let c' := translateBounded c []; - let t' := translateBounded t assume; - let f' := translateBounded f assume; - let e' := .ite (addAssumptions (assume ++ c'.assume) c'.translate) t'.translate f'.translate ty; - ⟨e', ListSet.union [boundExprIfType ty e', c'.wfCond, t'.wfCond, f'.wfCond], t'.assume ++ f'.assume⟩ + let c' := translateBounded c [] pos; + let t' := translateBounded t assume pos; + let f' := translateBounded f assume pos; + -- here can add inside if positive, never add outside + let e' := .ite (addAssumptions (assume ++ if pos then c'.assume else []) c'.translate) t'.translate f'.translate ty; + ⟨e', ListSet.union [boundExprIfType ty e', c'.wfCond, t'.wfCond, f'.wfCond], if pos then [] else c'.assume ++ t'.assume ++ f'.assume⟩ -- Equality is always bool-valued, so we can add assumptions + -- equality and iff are equivalent, NOT positive | .eq e1 e2 ty => - let e1' := translateBounded e1 []; - let e2' := translateBounded e2 []; - ⟨addAssumptions (assume ++ e1'.assume ++ e2'.assume) (.eq e1'.translate e2'.translate ty), ListSet.union [e1'.wfCond, e2'.wfCond], []⟩ + let e1' := translateBounded e1 [] false; + let e2' := translateBounded e2 [] false; + ⟨addAssumptions (assume ++ (if pos then e1'.assume ++ e2'.assume else [])) (.eq e1'.translate e2'.translate ty), ListSet.union [e1'.wfCond, e2'.wfCond], if pos then [] else e1'.assume ++ e2'.assume⟩ | .mdata m e => - let e' := translateBounded e assume; + let e' := translateBounded e assume pos; ⟨.mdata m e'.translate, e'.wfCond, e'.assume⟩ -- TODO: name def translateBounded' [Coe String Identifier] (e: LExprT Identifier) : LExprT Identifier := - (translateBounded e []).translate + (translateBounded e [] true).translate -- TEMPORARY until LExpr/LExprT are unified def LExprT.substK (k : Nat) (s : LExprT Identifier) (e : LExprT Identifier) : LExprT Identifier := @@ -342,7 +438,7 @@ def removeLeadingForalls (T : TEnv Identifier) (l: List (LExprT Identifier)) : E -- Functional version with extra quantifiers def boundedWfConditions' [Coe String Identifier] (e: LExprT Identifier) : List (LExprT Identifier) := - (translateBounded e []).wfCond + (translateBounded e [] true).wfCond -- Stateful version without extra quantifiers def boundedWfConditions [Coe String Identifier] (T : TEnv Identifier) (e: LExprT Identifier) : Except Format (List (LExprT Identifier) × TEnv Identifier) := removeLeadingForalls T (boundedWfConditions' e) diff --git a/Strata/DL/Bounded/BExprTest.lean b/Strata/DL/Bounded/BExprTest.lean index 0a0b8608..1a6ed2b3 100644 --- a/Strata/DL/Bounded/BExprTest.lean +++ b/Strata/DL/Bounded/BExprTest.lean @@ -702,43 +702,377 @@ def testBoundedChain : LExprT String := #guard_msgs in #eval eraseTy (translateBounded' testBoundedChain) --- Test 9: Bounded type in boolean context with negation --- Input: ¬((getValue 10 : {0 ≤ x}) < 5) --- Expected: translate = (0 ≤ getValue 10) → ¬(getValue 10 < 5), wfCond = [] -def testBoundedInBoolContext : LExprT String := - .app (LFunc.opExprT boolNotFunc) - (.app (.app (LFunc.opExprT intLtFunc) - (.app (.op "getValue" (.arrow .int (.bounded (.ble (.bconst 0) (.bvar))))) - (.const "10" .int) - (.bounded (.ble (.bconst 0) (.bvar)))) - (.arrow .int .bool)) - (.const "5" .int) - .bool) +end OtherTest + +-- Tests relating to positivity + +-- section PositiveTest + +-- -- Test 9: Bounded type in boolean context with negation +-- -- Input: ¬((getValue 10 : {0 ≤ x}) < 5) +-- -- Expected: translate = (0 ≤ getValue 10) → ¬(getValue 10 < 5), wfCond = [] +-- def testBoundedInBoolContext : LExprT String := +-- .app (LFunc.opExprT boolNotFunc) +-- (.app (.app (LFunc.opExprT intLtFunc) +-- (.app (.op "getValue" (.arrow .int (.bounded (.ble (.bconst 0) (.bvar))))) +-- (.const "10" .int) +-- (.bounded (.ble (.bconst 0) (.bvar)))) +-- (.arrow .int .bool)) +-- (.const "5" .int) +-- .bool) +-- .bool + +-- def foo : LExprT String := (.app (.app (LFunc.opExprT intLtFunc) +-- (.app (.op "getValue" (.arrow .int (.bounded (.ble (.bconst 0) (.bvar))))) +-- (.const "10" .int) +-- (.bounded (.ble (.bconst 0) (.bvar)))) +-- (.arrow .int .bool)) +-- (.const "5" .int) +-- .bool) + +-- #eval (eraseTys (translateBounded foo [] true).assume) + + +-- -- NOTE: all bottom-up assumptions need to be kept, even in boolean-valued terms, since they have to go to the top level + +-- --ex: not (foo x < 0) for foo: int -> nat +-- -- want (I think): 0 <= foo x -> not (foo x < 0) - TRUE +-- -- we do NOT want: not (0 <= foo x -> foo x < 0) - this is FALSE +-- -- think we need to collect bottom-up assumptions differently + +-- /-- info: ok: []-/ +-- #guard_msgs in +-- #eval runWFTest testBoundedInBoolContext +-- #eval (translateBounded testBoundedInBoolContext [] true).assume +-- #eval eraseTy (translateBounded' testBoundedInBoolContext) + +-- test: (foo x < 0 -> False) should be same as above + +-- not (forall (x: int), foo x < 0) + +-- (λ x y. x -> y) (foo x < 0) false +-- should be (forall x, foo x >= 0) -> (λ x y. x -> y) (foo x < 0) false + +/- +Here we have several tests for positivity. For many, we have the main test (e.g. Test 1) as well as an auxilliary test (Test 1a) showing the ordinary (i.e. no positivity complications) version. This demonstrates that we can create better translations when we know that there are no positivity restrictions. +-/ + +namespace PositiveTest + +def ltOp (e1 e2: LExprT String) : LExprT String := .app (.app (LFunc.opExprT intLtFunc) e1 (.arrow .int .bool)) e2 .bool +def impliesOp (e1 e2: LExprT String) : LExprT String := .app (.app (LFunc.opExprT boolImpliesFunc) e1 (.arrow .bool .bool)) e2 .bool + +-- 1. not (foo 1 < 0) where foo: int -> nat +-- Expected: 0 <= foo 1 -> not (foo 1 < 0) +def test1 : LExprT String := + notOp (ltOp (.app (.op "foo" (.arrow .int natTy)) (.const "1" .int) natTy) (.const "0" .int)) + +/-- info: Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.const "1" none)))) + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Not" none) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Int.Lt" none) + (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.const "1" none))) + (Lambda.LExpr.const "0" none))) -/ +#guard_msgs in +#eval eraseTy (translateBounded' test1) + +-- 1a. (foo 1 < 0) where foo: int -> nat +-- Expected: 0 <= foo 1 -> foo 1 < 0 +def test1a : LExprT String := + ltOp (.app (.op "foo" (.arrow .int natTy)) (.const "1" .int) natTy) (.const "0" .int) + +/-- +info: Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.const "1" none)))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Int.Lt" none) + (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.const "1" none))) + (Lambda.LExpr.const "0" none))-/ +#guard_msgs in +#eval eraseTy (translateBounded' test1a) + + +-- 2. not (forall (x: int), foo x < 0) for foo of same type +-- Expected: (forall (x: int), 0 <= foo x) -> not (forall (x: int), foo x < 0) +def test2 : LExprT String := + notOp (.quant .all .int (.bvar 0 .int) + (ltOp (.app (.op "foo" (.arrow .int natTy)) (.bvar 0 .int) natTy) (.const "0" .int))) + + +/-- info: Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.quant + (Lambda.QuantifierKind.all) + (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.bvar 0) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.bvar 0))))) + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Not" none) + (Lambda.LExpr.quant + (Lambda.QuantifierKind.all) + (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.bvar 0) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Int.Lt" none) + (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.bvar 0))) + (Lambda.LExpr.const "0" none)))) -/ +#guard_msgs in +#eval eraseTy (translateBounded' test2) + +-- 2a. (forall (x: int), foo x < 0) for foo: int -> nat +-- Expected: forall (x: int), 0 <= foo x -> foo x < 0 +def test2a : LExprT String := + .quant .all .int (.bvar 0 .int) + (ltOp (.app (.op "foo" (.arrow .int natTy)) (.bvar 0 .int) natTy) (.const "0" .int)) + +/-- info: Lambda.LExpr.quant + (Lambda.QuantifierKind.all) + (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.bvar 0) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.bvar 0)))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Int.Lt" none) + (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.bvar 0))) + (Lambda.LExpr.const "0" none)))-/ +#guard_msgs in +#eval eraseTy (translateBounded' test2a) + +-- 3. (foo 1 < 0 -> false) +-- Expected: 0 <= foo 1 -> (foo 1 < 0 -> false) +def test3 : LExprT String := + impliesOp (ltOp (.app (.op "foo" (.arrow .int natTy)) (.const "1" .int) natTy) (.const "0" .int)) + (.const "false" .bool) + +/-- info: Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.const "1" none)))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Int.Lt" none) + (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.const "1" none))) + (Lambda.LExpr.const "0" none))) + (Lambda.LExpr.const "false" none))-/ +#guard_msgs in +#eval eraseTy (translateBounded' test3) + +-- 3a. (b -> foo 1 < 0) for boolean constant b where foo: int -> nat +-- Expected: b -> 0 <= foo 1 -> foo 1 < 0 +def test3a : LExprT String := + impliesOp (.const "b" .bool) + (ltOp (.app (.op "foo" (.arrow .int natTy)) (.const "1" .int) natTy) (.const "0" .int)) + +/-- info: Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Bool.Implies" none) (Lambda.LExpr.const "b" none)) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.const "1" none)))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Int.Lt" none) + (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.const "1" none))) + (Lambda.LExpr.const "0" none)))-/ +#guard_msgs in +#eval eraseTy (translateBounded' test3a) + + +-- 4. (\x y. x -> y) (foo x < 0) false +-- Expected: (forall x, 0 <= foo x) -> (\x y: bool. x -> y) (foo x < 0) false +def test4 : LExprT String := + .app (.app (.abs (.abs (impliesOp (.bvar 1 .bool) (.bvar 0 .bool)) (.arrow .bool .bool)) (.arrow .bool (.arrow .bool .bool))) + (ltOp (.app (.op "foo" (.arrow .int natTy)) (.fvar "x" .int) natTy) (.const "0" .int)) + (.arrow .bool .bool)) + (.const "false" .bool) .bool -def foo : LExprT String := (.app (.app (LFunc.opExprT intLtFunc) - (.app (.op "getValue" (.arrow .int (.bounded (.ble (.bconst 0) (.bvar))))) - (.const "10" .int) - (.bounded (.ble (.bconst 0) (.bvar)))) - (.arrow .int .bool)) - (.const "5" .int) - .bool) +/-- info: Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.fvar "x" none)))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.abs + (some (Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.abs + (some (Lambda.LMonoTy.tcons "bool" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Bool.Implies" none) (Lambda.LExpr.bvar 1)) + (Lambda.LExpr.bvar 0)))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Int.Lt" none) + (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.fvar "x" none))) + (Lambda.LExpr.const "0" none))) + (Lambda.LExpr.const "false" none))-/ +#guard_msgs in +#eval eraseTy (translateBounded' test4) -#eval (eraseTys (translateBounded foo []).assume) +-- 5. (\x. foo x < 0) 1 +-- Expected: (forall x, 0 <= foo x) -> (\x. foo x < 0) 1 +def test5 : LExprT String := + .app (.abs (ltOp (.app (.op "foo" (.arrow .int natTy)) (.bvar 0 .int) natTy) (.const "0" .int)) (.arrow .int .bool)) + (.const "1" .int) + .bool --- NOTE: all bottom-up assumptions need to be kept, even in boolean-valued terms, since they have to go to the top level +/-- info: Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.quant + (Lambda.QuantifierKind.all) + (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.bvar 0) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.bvar 0))))) + (Lambda.LExpr.app + (Lambda.LExpr.abs + (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Int.Lt" none) + (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.bvar 0))) + (Lambda.LExpr.const "0" none))) + (Lambda.LExpr.const "1" none)) -/ +#guard_msgs in +#eval eraseTy (translateBounded' test5) ---ex: not (foo x < 0) for foo: int -> nat --- want (I think): 0 <= foo x -> not (foo x < 0) - TRUE --- we do NOT want: not (0 <= foo x -> foo x < 0) - this is FALSE --- think we need to collect bottom-up assumptions differently +-- 5a. (\x. foo y < 0) 1 where foo: int -> nat and y is a free variable +-- Expected: 0 <= foo y -> (\x. foo y < 0) 1 +def test5a : LExprT String := + .app (.abs (ltOp (.app (.op "foo" (.arrow .int natTy)) (.fvar "y" .int) natTy) (.const "0" .int)) (.arrow .int .bool)) + (.const "1" .int) + .bool -/-- info: ok: []-/ +/-- info: Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.fvar "y" none)))) + (Lambda.LExpr.app + (Lambda.LExpr.abs + (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Int.Lt" none) + (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.fvar "y" none))) + (Lambda.LExpr.const "0" none))) + (Lambda.LExpr.const "1" none))-/ #guard_msgs in -#eval runWFTest testBoundedInBoolContext -#eval (translateBounded testBoundedInBoolContext []).assume -#eval eraseTy (translateBounded' testBoundedInBoolContext) +#eval eraseTy (translateBounded' test5a) -end OtherTest + +-- 6. (\x. foo x) 1 >= 0 +-- Expected: 0 <= (\x. foo x) 1 -> (forall x, 0 <= foo x) -> (\x. foo x) 1 >= 0 +def test6 : LExprT String := + geOp (.app (.abs (.app (.op "foo" (.arrow .int natTy)) (.bvar 0 .int) natTy) (.arrow .int natTy)) + (.const "1" .int) + natTy) + (.const "0" .int) + +/-- info: Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.app + (Lambda.LExpr.abs + (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.bvar 0))) + (Lambda.LExpr.const "1" none)))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.quant + (Lambda.QuantifierKind.all) + (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.bvar 0) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.bvar 0))))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Int.Ge" none) + (Lambda.LExpr.app + (Lambda.LExpr.abs + (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) + (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.bvar 0))) + (Lambda.LExpr.const "1" none))) + (Lambda.LExpr.const "0" none)))-/ +#guard_msgs in +#eval eraseTy (translateBounded' test6) + +-- 7. (if foo x < 0 then bar == 1 else baz > 0) where bar and baz have type nat +-- Expected: 0 <= foo x -> if foo x < 0 then 0 <= bar -> bar == 1 else 0 <= baz -> baz > 0 +def test7 : LExprT String := + .ite (ltOp (.app (.op "foo" (.arrow .int natTy)) (.fvar "x" .int) natTy) (.const "0" .int)) + (.eq (.fvar "bar" natTy) (.const "1" .int) .bool) + (.app (.app (LFunc.opExprT intGtFunc) (.fvar "baz" natTy) (.arrow .int .bool)) (.const "0" .int) .bool) + .bool + +/-- info: Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.fvar "x" none)))) + (Lambda.LExpr.ite + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Int.Lt" none) + (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.fvar "x" none))) + (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.fvar "bar" none))) + (Lambda.LExpr.eq (Lambda.LExpr.fvar "bar" none) (Lambda.LExpr.const "1" none))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.fvar "baz" none))) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Gt" none) (Lambda.LExpr.fvar "baz" none)) + (Lambda.LExpr.const "0" none))))-/ +#guard_msgs in +#eval eraseTy (translateBounded' test7) + +end PositiveTest end Test From 0520d4f7681731c4275eb5bed9908b2ddbc75f3b Mon Sep 17 00:00:00 2001 From: Josh Cohen Date: Wed, 1 Oct 2025 15:31:46 -0400 Subject: [PATCH 10/19] Clean up application case a bit --- Strata/DL/Bounded/BExpr.lean | 90 +++++++++++++++----------------- Strata/DL/Bounded/BExprTest.lean | 2 + 2 files changed, 45 insertions(+), 47 deletions(-) diff --git a/Strata/DL/Bounded/BExpr.lean b/Strata/DL/Bounded/BExpr.lean index dd7c1e74..5f65d144 100644 --- a/Strata/DL/Bounded/BExpr.lean +++ b/Strata/DL/Bounded/BExpr.lean @@ -175,6 +175,20 @@ def quantifyAssumptions (ty: LMonoTy) (e: LExprT Identifier): LExprT Identifier --TODO: triggers? if hasBvar e 0 then .quant .all ty (.bvar 0 ty) e else e +def conditionalAdd (pos : Bool) (base extra : List α) : List α := + if pos then base ++ extra else base + +def conditionalReturn (pos : Bool) (value : List α) : List α := + if pos then [] else value + +def combineWfConditions (conditions : List (List (LExprT Identifier))) : ListSet (LExprT Identifier) := + ListSet.union conditions + +def makeBoundAssumption [Coe String Identifier] (ty : LMonoTy) : List (LExprT Identifier) := + match ty with + | .bounded b => [BoundExprToLExprT b (.bvar 0 .int)] + | _ => [] + -- We don't have pattern matching on operators (yet) because of the abstract metadata, so we resort to the follow workaround -- def isNot [Coe String Identifier] (e: LExprT Identifier) : Option (LExprT Identifier) := -- match e with @@ -265,58 +279,40 @@ def translateBounded [Coe String Identifier] [DecidableEq Identifier] (e: LExprT -/ | .app (.op o ty1) e2 .bool => - if o == boolNotFunc.name then + let notCase := o == boolNotFunc.name; -- inside "not" - not positive - let e2' := translateBounded e2 [] false; - let all_assumes := if pos then (assume ++ e2'.assume) else assume; - let res := addAssumptions all_assumes (.app (.op o ty1) e2'.translate .bool); - ⟨res, [], if pos then [] else e2'.assume⟩ - else - -- Here, safe to keep same positivity - -- let e1' := translateBounded e1 [] pos; - let e2' := translateBounded e2 [] pos; - let all_assumes := if pos then (assume ++ e2'.assume) else assume; - let res := addAssumptions all_assumes (.app (.op o (removeBound ty1)) e2'.translate .bool); - ⟨res, ListSet.union [wfCallCondition (assume ++ e2'.assume) (.op o ty1) e2'.translate, e2'.wfCond], if pos then [] else e2'.assume⟩ + let e2' := translateBounded e2 [] (not notCase && pos); + let all_assumes := conditionalAdd pos assume e2'.assume; + let res := addAssumptions all_assumes (.app (.op o (removeBound ty1)) e2'.translate .bool); + ⟨res, ListSet.union [wfCallCondition (assume ++ e2'.assume) (.op o ty1) e2'.translate, e2'.wfCond], conditionalReturn pos e2'.assume⟩ | .app (.app (.op o ty1) e1 ty2) e2 .bool => - -- Case 1: First argument NOT positive, second keeps positivity - if o == boolImpliesFunc.name then - let e1' := translateBounded e1 [] false; - let e2' := translateBounded e2 [] pos; - let all_assumes := if pos then (assume ++ e1'.assume ++ e2'.assume) else assume; -- TODO: factor out - let res := addAssumptions all_assumes (.app (.app (.op o (removeBound ty1)) e1'.translate (removeBound ty2)) e2'.translate .bool); - ⟨res, ListSet.union [e1'.wfCond, e2'.wfCond], if pos then [] else e1'.assume ++ e2'.assume⟩ - -- Case 2: Both arguments keep same positivity - else if o == boolAndFunc.name || o == boolOrFunc.name then - let e1' := translateBounded e1 [] pos; - let e2' := translateBounded e2 [] pos; - let all_assumes := if pos then (assume ++ e1'.assume ++ e2'.assume) else assume; - let res := addAssumptions all_assumes (.app (.app (.op o (removeBound ty1)) e1'.translate (removeBound ty2)) e2'.translate .bool); - ⟨res, ListSet.union [e1'.wfCond, e2'.wfCond], if pos then [] else e1'.assume ++ e2'.assume⟩ - -- Case 3: Both arguments cannot be assumed positive (equality and everything else - need because someone could define their own implication/equivalence; we know nothing about meaning) - else - let e1' := translateBounded e1 [] false; - let e2' := translateBounded e2 [] false; - let all_assumes := if pos then (assume ++ e1'.assume ++ e2'.assume) else assume; - let res := addAssumptions all_assumes (.app (.app (.op o (removeBound ty1)) e1'.translate (removeBound ty2)) e2'.translate .bool); - ⟨res, ListSet.union [wfCallCondition (assume ++ e2'.assume) e1 e2'.translate, e1'.wfCond, e2'.wfCond], if pos then [] else e1'.assume ++ e2'.assume⟩ + -- The first argument has positivity pos if the operator is "and" or "or" otherwise false + let first := (o == boolAndFunc.name || o == boolOrFunc.name) && pos; + -- The second also includes the RHS of implication + let second := (o == boolAndFunc.name || o == boolOrFunc.name || o == boolImpliesFunc.name) && pos; + let e1' := translateBounded e1 [] first; + let e2' := translateBounded e2 [] second; + let all_assumes := conditionalAdd pos assume (e1'.assume ++ e2'.assume); + let res := addAssumptions all_assumes (.app (.app (.op o (removeBound ty1)) e1'.translate (removeBound ty2)) e2'.translate .bool); + ⟨res, ListSet.union [wfCallCondition (assume ++ e2'.assume) e1 e2'.translate, e1'.wfCond, e2'.wfCond], conditionalReturn pos (e1'.assume ++ e2'.assume)⟩ | .app e1 e2 .bool => --Anything else we cannot assume is positive - let e1' := translateBounded e1 [] false; - let e2' := translateBounded e2 [] false; - let all_assumes := if pos then (assume ++ e1'.assume ++ e2'.assume) else assume; - let res := addAssumptions all_assumes (.app e1'.translate e2'.translate .bool); - ⟨res, ListSet.union [wfCallCondition (assume ++ e2'.assume) e1 e2'.translate, e1'.wfCond, e2'.wfCond], if pos then [] else e1'.assume ++ e2'.assume⟩ + let e1' := translateBounded e1 [] false; + let e2' := translateBounded e2 [] false; + let all_assumes := conditionalAdd pos assume (e1'.assume ++ e2'.assume); + let res := addAssumptions all_assumes (.app e1'.translate e2'.translate .bool); + ⟨res, ListSet.union [wfCallCondition (assume ++ e2'.assume) e1 e2'.translate, e1'.wfCond, e2'.wfCond], conditionalReturn pos (e1'.assume ++ e2'.assume)⟩ | .app e1 e2 ty => - let e1' := translateBounded e1 assume pos; - let e2' := translateBounded e2 assume pos; - let e' := .app e1'.translate e2'.translate (removeBound ty); - let extraWf := - match LExprT.toLMonoTy e1, ty with - | .arrow _ .int, .bounded _ => - boundExprIfType ty e' - | _, _ => []; - ⟨e', ListSet.union [wfCallCondition (assume ++ e2'.assume) e1 e2'.translate, extraWf, e1'.wfCond, e2'.wfCond], boundExprIfType ty e' ++ e1'.assume ++ e2'.assume⟩ + let e1' := translateBounded e1 assume pos; + let e2' := translateBounded e2 assume pos; + let e' := .app e1'.translate e2'.translate (removeBound ty); + -- If we have application f x where f : _ -> int and f x has bounded type, need wf condition that application result is bounded (cannot add in general because some bounds are assumed) + let extraWf := + match LExprT.toLMonoTy e1, ty with + | .arrow _ .int, .bounded _ => + boundExprIfType ty e' + | _, _ => []; + ⟨e', ListSet.union [wfCallCondition (assume ++ e2'.assume) e1 e2'.translate, extraWf, e1'.wfCond, e2'.wfCond], boundExprIfType ty e' ++ e1'.assume ++ e2'.assume⟩ /- Abstraction is the most complex case: 1. If the argument is bounded, add as top-down assumption diff --git a/Strata/DL/Bounded/BExprTest.lean b/Strata/DL/Bounded/BExprTest.lean index 1a6ed2b3..ca164f60 100644 --- a/Strata/DL/Bounded/BExprTest.lean +++ b/Strata/DL/Bounded/BExprTest.lean @@ -1073,6 +1073,8 @@ def test7 : LExprT String := #guard_msgs in #eval eraseTy (translateBounded' test7) +--TODO: test and/or + end PositiveTest end Test From 29a35a5f9748526d8e85a9421df506be970951b9 Mon Sep 17 00:00:00 2001 From: Josh Cohen Date: Fri, 3 Oct 2025 11:10:44 -0400 Subject: [PATCH 11/19] Clean up rest of translateBounded --- Strata/DL/Bounded/BExpr.lean | 95 ++++++++++++++++-------------------- 1 file changed, 43 insertions(+), 52 deletions(-) diff --git a/Strata/DL/Bounded/BExpr.lean b/Strata/DL/Bounded/BExpr.lean index 5f65d144..39fff142 100644 --- a/Strata/DL/Bounded/BExpr.lean +++ b/Strata/DL/Bounded/BExpr.lean @@ -68,6 +68,18 @@ def LFunc.opExprT (f: LFunc Identifier) : LExprT Identifier := | ity :: irest => Lambda.LMonoTy.mkArrow ity (irest ++ output_tys) .op f.name ty +/-- +Constructs the LExprT (.op o) e1 e2, where o has type bool -> bool -> bool +-/ +def boolOpExprT (o: LFunc Identifier) (e1 e2: LExprT Identifier) : LExprT Identifier := + .app (.app (LFunc.opExprT o) e1 (.arrow .bool .bool)) e2 .bool + +def andExprT [Coe String Identifier] (e1 e2: LExprT Identifier) : LExprT Identifier := + boolOpExprT boolAndFunc e1 e2 + +def implExprT [Coe String Identifier] (e1 e2: LExprT Identifier) : LExprT Identifier := + boolOpExprT boolImpliesFunc e1 e2 + --This is VERY messy,need to think about where this should take place --MUCH nicer to do on untyped terms but need type for function application -- Invariant, e must have type int @@ -80,15 +92,16 @@ LExprT Identifier := .app (.app (LFunc.opExprT intLtFunc) (BoundValToLExprT b1 e) (LMonoTy.arrow LMonoTy.int LMonoTy.bool)) (BoundValToLExprT b2 e) LMonoTy.bool | .ble b1 b2 => .app (.app (LFunc.opExprT intLeFunc) (BoundValToLExprT b1 e) (LMonoTy.arrow LMonoTy.int LMonoTy.bool)) (BoundValToLExprT b2 e) LMonoTy.bool - | .band e1 e2 => - .app (.app (LFunc.opExprT boolAndFunc) (BoundExprToLExprT e1 e) - (LMonoTy.arrow LMonoTy.bool LMonoTy.bool)) (BoundExprToLExprT e2 e) LMonoTy.bool + | .band e1 e2 => andExprT (BoundExprToLExprT e1 e) (BoundExprToLExprT e2 e) | .bnot e1 => .app (LFunc.opExprT boolNotFunc) (BoundExprToLExprT e1 e) LMonoTy.bool -- e must have type bool def addAssumptions [Coe String Identifier] (l: List (LExprT Identifier)) (e: LExprT Identifier) : LExprT Identifier := - List.foldr (fun x acc => .app (.app (LFunc.opExprT boolImpliesFunc) x LMonoTy.bool) acc LMonoTy.bool) e l + List.foldr implExprT e l + +def addAndAssumptions [Coe String Identifier] (l: List (LExprT Identifier)) (e: LExprT Identifier) : LExprT Identifier := + List.foldr andExprT e l def isBool (t: LMonoTy) : Bool := match t with @@ -175,15 +188,15 @@ def quantifyAssumptions (ty: LMonoTy) (e: LExprT Identifier): LExprT Identifier --TODO: triggers? if hasBvar e 0 then .quant .all ty (.bvar 0 ty) e else e -def conditionalAdd (pos : Bool) (base extra : List α) : List α := - if pos then base ++ extra else base +/-- +Add extra to base iff b holds +-/ +def conditionalAdd (b : Bool) (base extra : List α) : List α := + if b then base ++ extra else base def conditionalReturn (pos : Bool) (value : List α) : List α := if pos then [] else value -def combineWfConditions (conditions : List (List (LExprT Identifier))) : ListSet (LExprT Identifier) := - ListSet.union conditions - def makeBoundAssumption [Coe String Identifier] (ty : LMonoTy) : List (LExprT Identifier) := match ty with | .bounded b => [BoundExprToLExprT b (.bvar 0 .int)] @@ -320,53 +333,31 @@ def translateBounded [Coe String Identifier] [DecidableEq Identifier] (e: LExprT 3. Otherwise, add assumptions and increase bvars of all in "assume" list, as they are passing through a binder 3. WF: prove body satisfies bound if needed with same assumptions (but without new binder) -/ - | .abs e ty => - let newAssumption := - match ty with - | .arrow (.bounded b) _ => - [BoundExprToLExprT b (.bvar 0 .int)] - | _ => []; - match ty with - | .arrow ty1 .bool => - let e' := translateBounded e [] pos; - let all_assume := newAssumption ++ (increaseBVars assume) ++ (if pos then e'.assume else []); - -- TODO: change - ⟨.abs (addAssumptions all_assume e'.translate) (removeBound ty), addBoundedWfAssume all_assume e'.wfCond, if pos then [] else List.map (quantifyAssumptions ty1) e'.assume⟩ - | .arrow ty1 ty2 => - let all_assume := newAssumption ++ (increaseBVars assume); - let e' := translateBounded e all_assume pos; - let e'' := .abs e'.translate (removeBound ty); - -- Note: don't add assumptions to e'.wfCond, already included - ⟨e'', addBoundedWf all_assume (boundExprIfType ty2 e'.translate) e'.wfCond, List.map (quantifyAssumptions ty1) e'.assume⟩ - | _ => ⟨.const "0" .int, [], []⟩ -- error case + | .abs e (.arrow ty .bool) => + let e' := translateBounded e [] pos; + let all_assume := conditionalAdd pos (makeBoundAssumption ty ++ (increaseBVars assume)) e'.assume; + ⟨.abs (addAssumptions all_assume e'.translate) (removeBound (.arrow ty .bool)), addBoundedWfAssume all_assume e'.wfCond, conditionalReturn pos (List.map (quantifyAssumptions ty) e'.assume)⟩ + | .abs e (.arrow ty1 ty2) => + let all_assume := makeBoundAssumption ty1 ++ (increaseBVars assume); + let e' := translateBounded e all_assume pos; + let e'' := .abs e'.translate (removeBound (.arrow ty1 ty2)); + -- Note: don't add assumptions to e'.wfCond, already included + ⟨e'', addBoundedWf all_assume (boundExprIfType ty2 e'.translate) e'.wfCond, List.map (quantifyAssumptions ty1) e'.assume⟩ + | .abs _ _ => ⟨.const "0" .int, [], []⟩ -- error case /- Quantifiers are simpler because they are boolean-valued. ∀ (x : nat). e adds an assumption x >= 0 -> ..., while ∃ (x: nat). e results in x >= 0 ∧ .. -/ | .quant .all ty tr e => - let newAssumption := - match ty with - | .bounded b => - [BoundExprToLExprT b (.bvar 0 .int)] - | _ => []; let e' := translateBounded e [] pos; let tr' := translateBounded tr [] pos; --TODO: need "clean" one here - let all_assume := (newAssumption ++ (increaseBVars assume) ++ (if pos then e'.assume else [])); - -- TODO: factor out quant part - ⟨.quant .all (removeBound ty) tr'.translate (addAssumptions all_assume e'.translate), addBoundedWfAssume all_assume e'.wfCond, if pos then [] else List.map (quantifyAssumptions ty) e'.assume⟩ + let all_assume := conditionalAdd pos (makeBoundAssumption ty ++ (increaseBVars assume)) e'.assume; + ⟨.quant .all (removeBound ty) tr'.translate (addAssumptions all_assume e'.translate), addBoundedWfAssume all_assume e'.wfCond, conditionalReturn pos (List.map (quantifyAssumptions ty) e'.assume)⟩ | .quant .exist ty tr e => - let newAssumption := - match ty with - | .bounded b => - some (BoundExprToLExprT b (.bvar 0 .int)) - | _ => none; + let newAssumption := makeBoundAssumption ty; let e' := translateBounded e [] pos; let tr' := translateBounded tr [] pos; --TODO: need "clean" one here - let add_and x : LExprT Identifier := - match newAssumption with - | .some f => (.app (.app (LFunc.opExprT boolAndFunc) f (LMonoTy.arrow LMonoTy.bool LMonoTy.bool)) x LMonoTy.bool) - | .none => x; - let all_assume := (increaseBVars assume) ++ if pos then e'.assume else []; - ⟨.quant .exist (removeBound ty) tr'.translate (add_and (addAssumptions all_assume e'.translate)), addBoundedWfAssume (newAssumption.toList ++ all_assume) e'.wfCond, if pos then [] else List.map (quantifyAssumptions ty) e'.assume⟩ + let all_assume := conditionalAdd pos (increaseBVars assume) e'.assume; + ⟨.quant .exist (removeBound ty) tr'.translate (addAssumptions all_assume (addAndAssumptions newAssumption e'.translate)), addBoundedWfAssume (newAssumption ++ all_assume) e'.wfCond, conditionalReturn pos (List.map (quantifyAssumptions ty) e'.assume)⟩ /- For if-then-else, we add assumptions to the condition, which is always bool-valued. For a bool-valued result, we can add the conditions freely. For a bounded-valued term, we produce a wf condition proving this. -/ @@ -375,21 +366,21 @@ def translateBounded [Coe String Identifier] [DecidableEq Identifier] (e: LExprT let c' := translateBounded c [] false; let t' := translateBounded t [] pos; let f' := translateBounded f [] pos; - -- NOTE: VERY ugly - ⟨(if pos then addAssumptions c'.assume else id) (.ite (addAssumptions assume c'.translate) (addAssumptions (assume ++ if pos then t'.assume else []) t'.translate) (addAssumptions (assume ++ if pos then f'.assume else []) f'.translate) .bool), ListSet.union [c'.wfCond, t'.wfCond, f'.wfCond] , c'.assume ++ if pos then [] else t'.assume ++ f'.assume⟩ + let res := .ite (addAssumptions assume c'.translate) (addAssumptions (conditionalAdd pos assume t'.assume) t'.translate) (addAssumptions (conditionalAdd pos assume f'.assume) f'.translate) .bool; + ⟨(if pos then addAssumptions c'.assume else id) res, ListSet.union [c'.wfCond, t'.wfCond, f'.wfCond], c'.assume ++ conditionalReturn pos (t'.assume ++ f'.assume)⟩ | .ite c t f ty => let c' := translateBounded c [] pos; let t' := translateBounded t assume pos; let f' := translateBounded f assume pos; -- here can add inside if positive, never add outside - let e' := .ite (addAssumptions (assume ++ if pos then c'.assume else []) c'.translate) t'.translate f'.translate ty; - ⟨e', ListSet.union [boundExprIfType ty e', c'.wfCond, t'.wfCond, f'.wfCond], if pos then [] else c'.assume ++ t'.assume ++ f'.assume⟩ + let e' := .ite (addAssumptions (conditionalAdd pos assume c'.assume) c'.translate) t'.translate f'.translate ty; + ⟨e', ListSet.union [boundExprIfType ty e', c'.wfCond, t'.wfCond, f'.wfCond], conditionalReturn pos (c'.assume ++ t'.assume ++ f'.assume)⟩ -- Equality is always bool-valued, so we can add assumptions -- equality and iff are equivalent, NOT positive | .eq e1 e2 ty => let e1' := translateBounded e1 [] false; let e2' := translateBounded e2 [] false; - ⟨addAssumptions (assume ++ (if pos then e1'.assume ++ e2'.assume else [])) (.eq e1'.translate e2'.translate ty), ListSet.union [e1'.wfCond, e2'.wfCond], if pos then [] else e1'.assume ++ e2'.assume⟩ + ⟨addAssumptions (conditionalAdd pos assume (e1'.assume ++ e2'.assume)) (.eq e1'.translate e2'.translate ty), ListSet.union [e1'.wfCond, e2'.wfCond], conditionalReturn pos (e1'.assume ++ e2'.assume)⟩ | .mdata m e => let e' := translateBounded e assume pos; ⟨.mdata m e'.translate, e'.wfCond, e'.assume⟩ From 6eb071382924ff90ebdbaada528125a40dfbc352 Mon Sep 17 00:00:00 2001 From: Josh Cohen Date: Fri, 3 Oct 2025 14:23:31 -0400 Subject: [PATCH 12/19] Clean up code, add comments Better description of overall algorithm Mark definitions private Comments for most functions Other minor code quality improvements --- Strata/DL/Bounded/BExpr.lean | 338 ++++++++++++++----------------- Strata/DL/Bounded/BExprTest.lean | 55 +++-- 2 files changed, 180 insertions(+), 213 deletions(-) diff --git a/Strata/DL/Bounded/BExpr.lean b/Strata/DL/Bounded/BExpr.lean index 39fff142..65e54c32 100644 --- a/Strata/DL/Bounded/BExpr.lean +++ b/Strata/DL/Bounded/BExpr.lean @@ -10,16 +10,31 @@ import Strata.DL.Lambda.IntBoolFactory /-! ## Lambda Expressions with Bounded Ints -Here, we parameterize LExprs with bounded int types. These augmented expressions -can be desugared into regular Lambda expressions by erasing the int bounds. To make -the bounds meaningful, we also generate constraints. -All bounds must be explicitly stated; they will not be inferred. +Here, we parameterize LExprTs with bounded int types. These expressions can be desugared into regular Lambda expressions by performing two steps: +1. Translating the term, keeping track of assumptions implicit in the bounded types +2. Generating well-formedness conditions e.g. for function calls +Most bounds must be explicitly stated; the only inferred/checked ones occur when calling a function requiring a bounded argument on a int-typed argument. +This process -/ -/-! -In general, each input bound gives rise to an assumption in the expression body -as well as an obligation for any arguments, whereas an output bound is the opposite. +/- +The process of keeping track of assumptions and constraints is complicated. Here we give the key ideas (using nat = {x | 0 ≤ x} as an example): +1. There are two kinds of assumptions: context-dependent (e.g. for ∀ (x: nat). e, we assume 0 ≤ x when translating e) and context-independent (e.g. for external operator foo: int → nat, we are always free to assume 0 ≤ foo 1). +2. The context-dependent assumptions must be propagated top-down (when recursively translating the body of a quantifier or lambda), while the context-independent assumptions must be propagated bottom-up (as we do not know which external symbols appear in a subterm without inspecting it). +3. Assumptions can only be added to boolean-valued terms. Therefore, the basic algorithm is as follows: mantain a set of the top-down assumptions. Recursively traverse the term, removing bounded types. Whenever a bool-valued expression is reached, recursively translate the body and collect the bottom-up assumptions. Then add all assumptions to the translated body. +4. However, we cannot freely add assumptions in non-strictly-positive positions without changing the semantics. Therefore, if in such a position, we do not add the bottom-up assumptions but rather keep propagating upward until we reach a strictly positive position (which must exist at least at the outer layer). Since it is not possible to always determine whether a position is positive without β-reduction, we are conservative and assume all unknown positions are non-positive. +5. We must be careful about binders and quantify bottom-up assumptions that escape the scope of a quantified variable due to the above positivity restriction. For example, (∃ (x: int). foo x < 0) → False is translated to (∀ (x: int). 0 ≤ foo' x) → ((∃ (x: int). foo' x < 0) → False) where foo' has type int → int + +Generating well-formedness conditions is relatively simpler. The main rules are: +1. Constants that claim to have a bounded type are checked (e.g. 3: nat) +2. For application e₁e₂ where e₁ has type b → T, we check that b(tr(e₂)) +3. For application e₁e₂ : b where e₁ has type T -> int, we check that the translated application satisfies b. (Note: we cannot do this in general because some bounds e.g. for external operators are assumed) +4. For λ (x : T). e : T -> b, check that e satisfies bound b (possibly assuming any bounds for x) +Proving these well-formedness conditions may require both the top-down and bottom-up assumptions, so we compute all at the same time. + +See "BExprTest.lean" for test cases that further illustrate the expected translations and well-formedness conditions. + -/ namespace Bounded @@ -28,38 +43,18 @@ open Lambda variable {Identifier : Type} [ToString Identifier] [DecidableEq Identifier] [ToFormat Identifier] [HasGen Identifier] -/- - -There are two parts to desugaring BLambda to Lambda: translating the terms -by (almost) simply removing the integer bounds, and generating the -corresponding well-formedness conditions. Well-formedness is more general than -just bounded ints, though that is all we have for now. --/ - -def isBounded (t: LMonoTy) : Option BoundExpr := - match t with - | LMonoTy.bounded b => .some b - | _ => .none - -def removeBound (t: LMonoTy) : LMonoTy := - match t with - | LMonoTy.bounded _ => LMonoTy.int - | .tcons name args m => .tcons name (List.map removeBound args) m - | _ => t - +-- Translate bounded integer expressions to LExprT --- generate a single constraint with a body --- we rely on Lambda's existing substitution for this, simply constructing --- an expression with a single free variable (called "x", but it will always --- be substiuted so the name is irrelevant) --- ugh, free var is Identifier, do substitution directly ---not great, unsafe maybe? def BoundValToLExprT (b: BoundVal) (e: LExprT Identifier) : LExprT Identifier := match b with | .bvar => e | .bconst val => .const (val.repr) LMonoTy.int --- a hack for now (from Factory.lean) +/- +Utilities to construct common operators on common int and bool functions from IntBoolFactory.lean +-/ + +-- Adapted for LExprT from Factory.lean def LFunc.opExprT (f: LFunc Identifier) : LExprT Identifier := let input_tys := f.inputs.values let output_tys := Lambda.LMonoTy.destructArrow f.output @@ -71,60 +66,61 @@ def LFunc.opExprT (f: LFunc Identifier) : LExprT Identifier := /-- Constructs the LExprT (.op o) e1 e2, where o has type bool -> bool -> bool -/ -def boolOpExprT (o: LFunc Identifier) (e1 e2: LExprT Identifier) : LExprT Identifier := +def boolBinopExprT (o: LFunc Identifier) (e1 e2: LExprT Identifier) : LExprT Identifier := .app (.app (LFunc.opExprT o) e1 (.arrow .bool .bool)) e2 .bool +/-- +Constructs the LExprT (.op o) e1 e2, where o has type int -> int -> bool +-/ +def intCompBinopExprT (o: LFunc Identifier) (e1 e2: LExprT Identifier) : LExprT Identifier := + .app (.app (LFunc.opExprT o) e1 (.arrow .int .bool)) e2 .bool + def andExprT [Coe String Identifier] (e1 e2: LExprT Identifier) : LExprT Identifier := - boolOpExprT boolAndFunc e1 e2 + boolBinopExprT boolAndFunc e1 e2 def implExprT [Coe String Identifier] (e1 e2: LExprT Identifier) : LExprT Identifier := - boolOpExprT boolImpliesFunc e1 e2 + boolBinopExprT boolImpliesFunc e1 e2 + +def notExprT [Coe String Identifier] (e: LExprT Identifier) : LExprT Identifier := + .app (LFunc.opExprT boolNotFunc) e .bool ---This is VERY messy,need to think about where this should take place ---MUCH nicer to do on untyped terms but need type for function application --- Invariant, e must have type int +def intLtExprT [Coe String Identifier] (e1 e2: LExprT Identifier) : LExprT Identifier := + intCompBinopExprT intLtFunc e1 e2 + +def intLeExprT [Coe String Identifier] (e1 e2: LExprT Identifier) : LExprT Identifier := + intCompBinopExprT intLeFunc e1 e2 + +/-- +Translate BoundExpr b to corresponding Lambda expression +-/ def BoundExprToLExprT [Coe String Identifier] (b: BoundExpr) (e: LExprT Identifier) : LExprT Identifier := match b with | .beq b1 b2 => .eq (BoundValToLExprT b1 e) (BoundValToLExprT b2 e) LMonoTy.bool - | .blt b1 b2 => - .app (.app (LFunc.opExprT intLtFunc) (BoundValToLExprT b1 e) (LMonoTy.arrow LMonoTy.int LMonoTy.bool)) (BoundValToLExprT b2 e) LMonoTy.bool - | .ble b1 b2 => - .app (.app (LFunc.opExprT intLeFunc) (BoundValToLExprT b1 e) (LMonoTy.arrow LMonoTy.int LMonoTy.bool)) (BoundValToLExprT b2 e) LMonoTy.bool + | .blt b1 b2 => intLtExprT (BoundValToLExprT b1 e) (BoundValToLExprT b2 e) + | .ble b1 b2 => intLeExprT (BoundValToLExprT b1 e) (BoundValToLExprT b2 e) | .band e1 e2 => andExprT (BoundExprToLExprT e1 e) (BoundExprToLExprT e2 e) - | .bnot e1 => - .app (LFunc.opExprT boolNotFunc) (BoundExprToLExprT e1 e) LMonoTy.bool + | .bnot e1 => notExprT (BoundExprToLExprT e1 e) + +-- Auxilliary functions for translation --- e must have type bool -def addAssumptions [Coe String Identifier] (l: List (LExprT Identifier)) (e: LExprT Identifier) : LExprT Identifier := +/-- +Add list of assumptions l to boolean-valued e +-/ +private def addAssumptions [Coe String Identifier] (l: List (LExprT Identifier)) (e: LExprT Identifier) : LExprT Identifier := List.foldr implExprT e l -def addAndAssumptions [Coe String Identifier] (l: List (LExprT Identifier)) (e: LExprT Identifier) : LExprT Identifier := +/-- +Produce l₁ ∧ ... ∧ lₙ ∧ e +-/ +private def addAndExprs [Coe String Identifier] (l: List (LExprT Identifier)) (e: LExprT Identifier) : LExprT Identifier := List.foldr andExprT e l -def isBool (t: LMonoTy) : Bool := - match t with - | .bool => True - | _ => False - --- Only deal with expressions of form bvar, eq, app (output of BoundExprToLExprT on bvar) -def increaseBVar (e: LExprT Identifier) : LExprT Identifier := - match e with - | .bvar b ty => .bvar (b+1) ty - | .eq e1 e2 ty => .eq (increaseBVar e1) (increaseBVar e2) ty - | .app e1 e2 ty => .app (increaseBVar e1) (increaseBVar e2) ty - | _ => e - - -def increaseBVars (l: List (LExprT Identifier)) : List (LExprT Identifier) := - List.map increaseBVar l - --- TODO: see what to use here +/- +An inefficient method of mantaining a list with no duplicates. Should be replaced with a HashSet or similar +-/ def ListSet α := List α --- #print Membership --- instance : Membership α (ListSet α) := --- Membership.mk (fun l x => by unfold ListSet at l; exact (x ∈ l)) def ListSet.contains {α} [DecidableEq α] (l: ListSet α) (x: α) : Bool := List.foldr (fun y acc => x == y || acc) false l @@ -138,20 +134,45 @@ def ListSet.insertAll {α} [DecidableEq α] (l: ListSet α) (xs: List α) : List def ListSet.union {α} [DecidableEq α] (l: List (ListSet α)) : ListSet α := List.foldr ListSet.insertAll [] l -/- Do translation and wf generation in 1 go - need same assumptions, expesive to compute both-/ -structure translationRes Identifier where -(translate : LExprT Identifier) -(wfCond : ListSet (LExprT Identifier)) -(assume : List (LExprT Identifier)) +-- Auxiliary functions for producing bounds/assumptions --- more aux functions +/- +All top-down assumptions initially start as expressions over (.bvar 0 .int). As we pass through binders, these bvars must be increased. We only care about expressions consisting of bvar, eq, app (output of BoundExprToLExprT) +-/ +private def increaseBVar (e: LExprT Identifier) : LExprT Identifier := + match e with + | .bvar b ty => .bvar (b+1) ty + | .eq e1 e2 ty => .eq (increaseBVar e1) (increaseBVar e2) ty + | .app e1 e2 ty => .app (increaseBVar e1) (increaseBVar e2) ty + | _ => e + +private def increaseBVars (l: List (LExprT Identifier)) : List (LExprT Identifier) := + List.map increaseBVar l + +private def isBounded (t: LMonoTy) : Option BoundExpr := + match t with + | LMonoTy.bounded b => .some b + | _ => .none --- evaluate expression at bound if type is bounded -def boundExprIfType [Coe String Identifier] (ty: LMonoTy) (e: LExprT Identifier) : List (LExprT Identifier) := +private def removeTyBound (t: LMonoTy) : LMonoTy := + match t with + | LMonoTy.bounded _ => LMonoTy.int + | .tcons name args m => .tcons name (List.map removeTyBound args) m + | _ => t + +/-- +If ty is .bounded b, produces expression b(e) +-/ +private def boundExprIfType [Coe String Identifier] (ty: LMonoTy) (e: LExprT Identifier) : List (LExprT Identifier) := ((isBounded ty).map (fun b => BoundExprToLExprT b e)).toList /-- -Generate WF condition for calling e1 with argument e2 if bounded type expected +If ty is .bounded b, produces expression b(.bvar 0 .int) +-/ +def makeBoundAssumption [Coe String Identifier] (ty : LMonoTy) : List (LExprT Identifier) := boundExprIfType ty (.bvar 0 .int) + +/-- +Generates WF condition for calling e1 with argument e2 if bounded type expected -/ def wfCallCondition [Coe String Identifier] (assume : List (LExprT Identifier)) (e1 e2: LExprT Identifier) := match LExprT.toLMonoTy e1 with @@ -160,17 +181,22 @@ def wfCallCondition [Coe String Identifier] (assume : List (LExprT Identifier)) [addAssumptions assume (BoundExprToLExprT b e2)] | _ => [] --- NOTE: l1 includes assumptions, l2 does not -def addBoundedWf [Coe String Identifier] (assume: List (LExprT Identifier)) (l1 l2: List (LExprT Identifier)) : List (LExprT Identifier) := +/-- + Universally quantifies all expressions in lists l1 and l2, additionally adding assumptions in "assume" to l1. These quantifiers will eventually be removed in a postprocessing step +-/ +private def addBoundedWf [Coe String Identifier] (assume: List (LExprT Identifier)) (l1 l2: List (LExprT Identifier)) : List (LExprT Identifier) := List.map (fun e => .quant .all .int (.bvar 0 .int) (addAssumptions assume e)) l1 ++ List.map (.quant .all .int (.bvar 0 .int)) l2 -def addBoundedWfAssume [Coe String Identifier] (assume: List (LExprT Identifier)) (l: List (LExprT Identifier)) := +/-- + Universally quantifies all expressions in list l with additional assumptions in "assume" +-/ +private def addBoundedWfAssume [Coe String Identifier] (assume: List (LExprT Identifier)) (l: List (LExprT Identifier)) := addBoundedWf assume l [] /-- Determines if a term uses a particular bound variable at the top level -/ -def hasBvar (e: LExprT Identifier) (n: Nat) : Bool := +private def hasBvar (e: LExprT Identifier) (n: Nat) : Bool := match e with | .bvar m _ => n == m | .app e1 e2 _ => hasBvar e1 n || hasBvar e2 n @@ -184,7 +210,7 @@ def hasBvar (e: LExprT Identifier) (n: Nat) : Bool := /-- When passing an assumption through a binder, we need to quantify the bound variable if used. -/ -def quantifyAssumptions (ty: LMonoTy) (e: LExprT Identifier): LExprT Identifier := +private def quantifyAssumptions (ty: LMonoTy) (e: LExprT Identifier): LExprT Identifier := --TODO: triggers? if hasBvar e 0 then .quant .all ty (.bvar 0 ty) e else e @@ -197,106 +223,42 @@ def conditionalAdd (b : Bool) (base extra : List α) : List α := def conditionalReturn (pos : Bool) (value : List α) : List α := if pos then [] else value -def makeBoundAssumption [Coe String Identifier] (ty : LMonoTy) : List (LExprT Identifier) := - match ty with - | .bounded b => [BoundExprToLExprT b (.bvar 0 .int)] - | _ => [] - --- We don't have pattern matching on operators (yet) because of the abstract metadata, so we resort to the follow workaround --- def isNot [Coe String Identifier] (e: LExprT Identifier) : Option (LExprT Identifier) := --- match e with --- | .app (.op o _) e2 _ => --- if o == boolNotFunc.name then some e2 else none --- | _ => none - --- def isImpl [Coe String Identifier] (e: LExprT Identifier) : Option (LExprT Identifier × LExprT Identifier) := --- match e with --- | .app (.app (.op o _) e1 _) e2 _ => --- if o == boolImpliesFunc.name then some (e1, e2) else none --- | _ => none - --- Ugh, need to know how to 1. combine results to produce new term --- 2. produce wf conditions (should it just be app Bool case) START HERE --- def boolCase [Coe String Identifier] (assume: List (LExprT Identifier)) (l: List (translationRes Identifier)) (comb: List (LExprT Identifier) -> LExprT Identifier) (pos: Bool) := --- let all_assumes := assume ++ List.flatten (List.map translationRes.assume l); --- let res := (if pos then addAssumptions all_assumes else addAssumptions assume) (comb (List.map translationRes.translate l)); --- (res, ) - - --- let e1' := translateBounded e1 [] false; --- let e2' := translateBounded e2 [] false; --- let all_assumes := assume ++ e1'.assume ++ e2'.assume; --- let res := addAssumptions all_assumes (.app e1'.translate e2'.translate .bool); --- ⟨res, ListSet.union [wfCallCondition (assume ++ e2'.assume) e1 e2'.translate, e1'.wfCond, e2'.wfCond], []⟩ +structure translationRes Identifier where +(translate : LExprT Identifier) +(wfCond : ListSet (LExprT Identifier)) +(assume : List (LExprT Identifier)) /-- -Translate expression to one without bounded ints. -This is a non-trivial translation, as we want to preserve the semantics of the resulting term. -To see why this is tricky, consider the following examples: -1. ∀ (x: Nat), 0 <= x (should result in something semantically equivalent to true) -2. λ(x: Nat), if 0 <= x then 1 else 0 (should result in function that always evaluates to 1) -3. λ(x: int), if foo x >= 0 then 1 else 0 (supposing foo : int -> Nat, should also result in function that always evaluates to 1) -4. (x: Nat) + (y: Nat) >= 0 must be true - -To handle the first case, any bound variables induce constraints that must be -inserted in the resulting formula. The second case shows that adding the constraints is complicated, because the resulting term is not necessarily bool-valued. The third example shows that variables are not enough: any term of bounded type must result in an assumption somehow. The fourth shows that terms may not be closed; therefore, we must collect free variable assumptions as well. - -To deal with all of this, our translation function produces two outputs: the translated term and a set of constraints. Broadly, new constraints are added whenever we see a bounded bound variable or an expression of bounded type. We must propagate them back up to any boolean-valued expressions. There is a subtlety resulting from bound variables, as we must make sure the constraints are scoped appropriately (e.g. ∀ (y: Nat). λ(x: Nat). x + y >= 0, which becomes -∀ Nat. λ Nat. #0 + #1 >= 0 must result in ∀ #0 >= 0 -> λ Nat. #0 >= 0 -> #0 + #1 >= 0). - -Also note that assumptions percolate in two directions: A bound variable (possibly) induces a new assumption somewhere inside a subterm, while a (e.g.) function application induces a new assumption in a parent call. Thus, we need both input (the first kind) and output (the second kind) assumption lists. - -Note that we do NOT have to worry about the fact that the new functions have a larger domain when types are erased. We generate separate well-formedness checks (below) to ensure that any call of the function occurs on an argument satisfying the constraints. - -Invariant: assumptions must not have bounded types (TODO enforce this), same for inputs -Invariant (I think): All assumptions are of form: b(bvar #x) - -There is one more complication. When propagating bottom-up assumptions, we cannot just add them at the first bool expression encountered. To see why, consider the example: -~ ((x: Nat) < 0) -This must become -0 <= x -> ~ (x < 0) -NOT -~ (0 <= x -> x < 0) -basically, this should occur any time we are on the left of an implication. Thus, we must add assumptions at any top-level boolean expression. To keep track of this, we use the parameter inBool which becomes false once we have passed at least one bool-valued expression. We only add assumptions if inBool is - -We need to keep track of positivity. We cannot add assumptions (e.g) on the LHS of an implication; we need to propagate a level above. This is annoying because connectives are not built in to Lambda. - -Idea: if positive, can add all assumptions, if not, can add top-down but propagae bottom-up. Other complication - if we propagate, look through, if uses unbound bvar 0, add quantifier at beginning (only for lambda, quantifier case) - -Need to make sure - ALWAYS safe to add top-down assumptions - + Under top-down assumptions "assume" and positivity "pos", produce the translation, well-formedness conditions, and bottom-up assumptions for expression e -/ def translateBounded [Coe String Identifier] [DecidableEq Identifier] (e: LExprT Identifier) (assume: List (LExprT Identifier)) (pos : Bool) : translationRes Identifier := match e with -- constants do not need assumptions; they produce a wf goal if bounded | .const c ty => - let res := .const c (removeBound ty); + let res := .const c (removeTyBound ty); ⟨res, boundExprIfType ty res, []⟩ -- an op adds a bottom-up assumption if it has bounded type; its wf is assumed | .op o ty => - let res := .op o (removeBound ty); - ⟨if isBool ty then addAssumptions assume res else res, [], boundExprIfType ty res⟩ + let res := .op o (removeTyBound ty); + ⟨if ty == .bool then addAssumptions assume res else res, [], boundExprIfType ty res⟩ -- bvars are handled when bound - | .bvar b ty => ⟨ .bvar b (removeBound ty), [], [] ⟩ + | .bvar b ty => ⟨ .bvar b (removeTyBound ty), [], [] ⟩ -- fvars generate bottom-up assumptions if bounded | .fvar f ty => - let res := .fvar f (removeBound ty); + let res := .fvar f (removeTyBound ty); ⟨res, [], boundExprIfType ty res⟩ /- Application has several cases: - 1. If the entire application has boolean type, assumptions can be added + 1. If the entire application has boolean type, assumptions can be added subject to positivity. 2. Otherwise, if the application has bounded type, it produces a bottom-up assumption. There is a subtle case. If the function has type (t -> int), then we must generate a wf condition. Otherwise, the bound will be assumed (for external operators) or checked (for abstraction/if-then-else/etc) - 3. In either case, we produce a wf condition if the argument should have bounded type - - TODO: factor out boolean cases - need to say: if !pos, then propagate bottom-up assumptions (do NOT add); we can still add top-down ones (these are always safe to add) otherwise, add assumptions as right now - + 3. In either case, we produce a wf condition if the argument should have bounded type. -/ | .app (.op o ty1) e2 .bool => let notCase := o == boolNotFunc.name; -- inside "not" - not positive let e2' := translateBounded e2 [] (not notCase && pos); let all_assumes := conditionalAdd pos assume e2'.assume; - let res := addAssumptions all_assumes (.app (.op o (removeBound ty1)) e2'.translate .bool); + let res := addAssumptions all_assumes (.app (.op o (removeTyBound ty1)) e2'.translate .bool); ⟨res, ListSet.union [wfCallCondition (assume ++ e2'.assume) (.op o ty1) e2'.translate, e2'.wfCond], conditionalReturn pos e2'.assume⟩ | .app (.app (.op o ty1) e1 ty2) e2 .bool => -- The first argument has positivity pos if the operator is "and" or "or" otherwise false @@ -306,7 +268,7 @@ def translateBounded [Coe String Identifier] [DecidableEq Identifier] (e: LExprT let e1' := translateBounded e1 [] first; let e2' := translateBounded e2 [] second; let all_assumes := conditionalAdd pos assume (e1'.assume ++ e2'.assume); - let res := addAssumptions all_assumes (.app (.app (.op o (removeBound ty1)) e1'.translate (removeBound ty2)) e2'.translate .bool); + let res := addAssumptions all_assumes (.app (.app (.op o (removeTyBound ty1)) e1'.translate (removeTyBound ty2)) e2'.translate .bool); ⟨res, ListSet.union [wfCallCondition (assume ++ e2'.assume) e1 e2'.translate, e1'.wfCond, e2'.wfCond], conditionalReturn pos (e1'.assume ++ e2'.assume)⟩ | .app e1 e2 .bool => --Anything else we cannot assume is positive @@ -318,8 +280,8 @@ def translateBounded [Coe String Identifier] [DecidableEq Identifier] (e: LExprT | .app e1 e2 ty => let e1' := translateBounded e1 assume pos; let e2' := translateBounded e2 assume pos; - let e' := .app e1'.translate e2'.translate (removeBound ty); - -- If we have application f x where f : _ -> int and f x has bounded type, need wf condition that application result is bounded (cannot add in general because some bounds are assumed) + let e' := .app e1'.translate e2'.translate (removeTyBound ty); + -- If we have application f x where f : _ -> int and f x has bounded type, need wf condition that application result is bounded let extraWf := match LExprT.toLMonoTy e1, ty with | .arrow _ .int, .bounded _ => @@ -327,20 +289,20 @@ def translateBounded [Coe String Identifier] [DecidableEq Identifier] (e: LExprT | _, _ => []; ⟨e', ListSet.union [wfCallCondition (assume ++ e2'.assume) e1 e2'.translate, extraWf, e1'.wfCond, e2'.wfCond], boundExprIfType ty e' ++ e1'.assume ++ e2'.assume⟩ /- - Abstraction is the most complex case: + Lambda abstraction: 1. If the argument is bounded, add as top-down assumption 2. If the body has type bool, add assumptions and translate 3. Otherwise, add assumptions and increase bvars of all in "assume" list, as they are passing through a binder - 3. WF: prove body satisfies bound if needed with same assumptions (but without new binder) + 4. WF: if body bounded, prove body satisfies bound with same assumptions (but without new binder) -/ | .abs e (.arrow ty .bool) => let e' := translateBounded e [] pos; let all_assume := conditionalAdd pos (makeBoundAssumption ty ++ (increaseBVars assume)) e'.assume; - ⟨.abs (addAssumptions all_assume e'.translate) (removeBound (.arrow ty .bool)), addBoundedWfAssume all_assume e'.wfCond, conditionalReturn pos (List.map (quantifyAssumptions ty) e'.assume)⟩ + ⟨.abs (addAssumptions all_assume e'.translate) (removeTyBound (.arrow ty .bool)), addBoundedWfAssume all_assume e'.wfCond, conditionalReturn pos (List.map (quantifyAssumptions ty) e'.assume)⟩ | .abs e (.arrow ty1 ty2) => let all_assume := makeBoundAssumption ty1 ++ (increaseBVars assume); let e' := translateBounded e all_assume pos; - let e'' := .abs e'.translate (removeBound (.arrow ty1 ty2)); + let e'' := .abs e'.translate (removeTyBound (.arrow ty1 ty2)); -- Note: don't add assumptions to e'.wfCond, already included ⟨e'', addBoundedWf all_assume (boundExprIfType ty2 e'.translate) e'.wfCond, List.map (quantifyAssumptions ty1) e'.assume⟩ | .abs _ _ => ⟨.const "0" .int, [], []⟩ -- error case @@ -351,15 +313,15 @@ def translateBounded [Coe String Identifier] [DecidableEq Identifier] (e: LExprT let e' := translateBounded e [] pos; let tr' := translateBounded tr [] pos; --TODO: need "clean" one here let all_assume := conditionalAdd pos (makeBoundAssumption ty ++ (increaseBVars assume)) e'.assume; - ⟨.quant .all (removeBound ty) tr'.translate (addAssumptions all_assume e'.translate), addBoundedWfAssume all_assume e'.wfCond, conditionalReturn pos (List.map (quantifyAssumptions ty) e'.assume)⟩ + ⟨.quant .all (removeTyBound ty) tr'.translate (addAssumptions all_assume e'.translate), addBoundedWfAssume all_assume e'.wfCond, conditionalReturn pos (List.map (quantifyAssumptions ty) e'.assume)⟩ | .quant .exist ty tr e => let newAssumption := makeBoundAssumption ty; let e' := translateBounded e [] pos; let tr' := translateBounded tr [] pos; --TODO: need "clean" one here let all_assume := conditionalAdd pos (increaseBVars assume) e'.assume; - ⟨.quant .exist (removeBound ty) tr'.translate (addAssumptions all_assume (addAndAssumptions newAssumption e'.translate)), addBoundedWfAssume (newAssumption ++ all_assume) e'.wfCond, conditionalReturn pos (List.map (quantifyAssumptions ty) e'.assume)⟩ + ⟨.quant .exist (removeTyBound ty) tr'.translate (addAssumptions all_assume (addAndExprs newAssumption e'.translate)), addBoundedWfAssume (newAssumption ++ all_assume) e'.wfCond, conditionalReturn pos (List.map (quantifyAssumptions ty) e'.assume)⟩ /- - For if-then-else, we add assumptions to the condition, which is always bool-valued. For a bool-valued result, we can add the conditions freely. For a bounded-valued term, we produce a wf condition proving this. + For if-then-else, add assumptions to the condition, which is always bool-valued. For a bool-valued result, can add the conditions freely. For a bounded-valued term, produce a wf condition proving this. -/ | .ite c t f .bool => -- c is NOT positive; equivalent to c -> t /\ ~c -> f @@ -375,8 +337,10 @@ def translateBounded [Coe String Identifier] [DecidableEq Identifier] (e: LExprT -- here can add inside if positive, never add outside let e' := .ite (addAssumptions (conditionalAdd pos assume c'.assume) c'.translate) t'.translate f'.translate ty; ⟨e', ListSet.union [boundExprIfType ty e', c'.wfCond, t'.wfCond, f'.wfCond], conditionalReturn pos (c'.assume ++ t'.assume ++ f'.assume)⟩ - -- Equality is always bool-valued, so we can add assumptions - -- equality and iff are equivalent, NOT positive + /- + Equality is always bool-valued, so can add assumptions + Equality/iff are equivalent, NOT positive + -/ | .eq e1 e2 ty => let e1' := translateBounded e1 [] false; let e2' := translateBounded e2 [] false; @@ -385,12 +349,19 @@ def translateBounded [Coe String Identifier] [DecidableEq Identifier] (e: LExprT let e' := translateBounded e assume pos; ⟨.mdata m e'.translate, e'.wfCond, e'.assume⟩ --- TODO: name -def translateBounded' [Coe String Identifier] (e: LExprT Identifier) : LExprT Identifier := +/-- +Translated an expression with bounded types to one without, preserving semantics of bool-valued terms by adding assumptions +-/ +def translateBoundedExpr [Coe String Identifier] (e: LExprT Identifier) : LExprT Identifier := (translateBounded e [] true).translate --- TEMPORARY until LExpr/LExprT are unified -def LExprT.substK (k : Nat) (s : LExprT Identifier) (e : LExprT Identifier) : LExprT Identifier := +/- +Many of the wf conditions have the form: forall x, f(x). We remove the quantifiers to make the SMT formulas easier to solve. +Must be stateful because we need to generate fresh variables +-/ + +-- temporary until LExpr/LExprT are unified +private def LExprT.substK (k : Nat) (s : LExprT Identifier) (e : LExprT Identifier) : LExprT Identifier := match e with | .const c ty => .const c ty | .op o ty => .op o ty @@ -406,30 +377,27 @@ def LExprT.substK (k : Nat) (s : LExprT Identifier) (e : LExprT Identifier) : LE def LExprT.varOpen (k : Nat) (x : Identifier × LMonoTy) (e : LExprT Identifier) : LExprT Identifier := LExprT.substK k (.fvar x.fst x.snd) e -/- -Many of the wf conditions have the form: forall x, f(x). We remove the quantifiers to make the SMT formulas easier to solve. -Must be stateful because we need to generate fresh variables +/-- +Remove leading "forall" quantifiers statefully -/ -def removeLeadingForall (T : TEnv Identifier) (e: LExprT Identifier) : Except Format (LExprT Identifier × TEnv Identifier) := +private def removeLeadingForall (T : TEnv Identifier) (e: LExprT Identifier) : Except Format (LExprT Identifier × TEnv Identifier) := match e with | .quant .all ty _ e => do let (xv, xty, T) ← Lambda.LExprT.typeBoundVar T ty; .ok (LExprT.varOpen 0 (xv, xty) e, T) | _ => .ok (e, T) -def removeLeadingForalls (T : TEnv Identifier) (l: List (LExprT Identifier)) : Except Format (List (LExprT Identifier) × TEnv Identifier) := +private def removeLeadingForalls (T : TEnv Identifier) (l: List (LExprT Identifier)) : Except Format (List (LExprT Identifier) × TEnv Identifier) := List.foldlM (fun (l, T) e => do let (e, T') ← removeLeadingForall T e; .ok (e :: l, T')) ([], T) l -- Functional version with extra quantifiers -def boundedWfConditions' [Coe String Identifier] (e: LExprT Identifier) : List (LExprT Identifier) := +def boundedWfConditionsQuant [Coe String Identifier] (e: LExprT Identifier) : List (LExprT Identifier) := (translateBounded e [] true).wfCond -- Stateful version without extra quantifiers -def boundedWfConditions [Coe String Identifier] (T : TEnv Identifier) (e: LExprT Identifier) : Except Format (List (LExprT Identifier) × TEnv Identifier) := removeLeadingForalls T (boundedWfConditions' e) - --- NOTE: the assumptions are useful: they show us the "axioms" that we depend on (assumptions about external ops and free variables) +def boundedWfConditions [Coe String Identifier] (T : TEnv Identifier) (e: LExprT Identifier) : Except Format (List (LExprT Identifier) × TEnv Identifier) := removeLeadingForalls T (boundedWfConditionsQuant e) end Bounded diff --git a/Strata/DL/Bounded/BExprTest.lean b/Strata/DL/Bounded/BExprTest.lean index ca164f60..d1eb0d02 100644 --- a/Strata/DL/Bounded/BExprTest.lean +++ b/Strata/DL/Bounded/BExprTest.lean @@ -9,9 +9,8 @@ open Std (ToFormat Format format) open Lambda open Bounded --- NOTE: with a semantics for LExpr/LExprT, we could prove the equivalences mentioned above - def natTy : LMonoTy := LMonoTy.bounded (.ble (.bconst 0) .bvar) + def leOp (e1 e2: LExprT String) : LExprT String := .app (.app (LFunc.opExprT intLeFunc) e1 (.arrow .int .bool)) e2 .bool def geOp (e1 e2: LExprT String) : LExprT String := .app (.app (LFunc.opExprT intGeFunc) e1 (.arrow .int .bool)) e2 .bool @@ -50,7 +49,7 @@ def test1 := (@LExprT.quant String .all natTy (.bvar 0 natTy) (leOp (.const "0" (Lambda.LExpr.bvar 0))) -/ #guard_msgs in -#eval (eraseTy (translateBounded' test1) ) +#eval (eraseTy (translateBoundedExpr test1) ) -- 2. λ(x: Nat), if 0 <= x then 1 else 2 (assumption inside ite) -- Expected: λ (x: int), if 0 <= x -> 0 <= x then 1 else 2 @@ -72,7 +71,7 @@ def test2 : LExprT String := .abs (.ite (leOp (.const "0" .int) (.bvar 0 .int)) (Lambda.LExpr.const "1" none) (Lambda.LExpr.const "2" none))-/ #guard_msgs in -#eval (eraseTy (translateBounded' test2)) +#eval (eraseTy (translateBoundedExpr test2)) -- 3. λ(x: int), if foo x >= 0 then 1 else 0 (for foo: int -> Nat) (propagate op/application information) -- Expected: λ (x: int), if 0 <= foo x -> foo x >= 0 then 1 else 0 @@ -96,7 +95,7 @@ def test3 : LExprT String := .abs (.ite (geOp (.app (.op "foo" (.arrow .int natT (Lambda.LExpr.const "1" none) (Lambda.LExpr.const "0" none))-/ #guard_msgs in -#eval (eraseTy (translateBounded' test3)) +#eval (eraseTy (translateBoundedExpr test3)) -- 4. (x: Nat) + (y: Nat) >= 0 (free variable bounds) -- Expected: 0 <= (x: int) -> 0 <= (y : int) -> x + y >= 0 @@ -123,7 +122,7 @@ def test4 : LExprT String := geOp (addOp (.fvar "x" natTy) (.fvar "y" natTy)) (. (Lambda.LExpr.fvar "y" none))) (Lambda.LExpr.const "0" none)))-/ #guard_msgs in -#eval eraseTy (translateBounded' test4) +#eval eraseTy (translateBoundedExpr test4) -- 5. λ (x: Nat). λ (y: Nat). x + y >= 0 (multiple bound vars) -- Expected: λ (x: int). λ (y: int). 0 <= y -> 0 <= x -> x + y >= 0 @@ -154,7 +153,7 @@ def test5 : LExprT String := .abs (.abs (geOp (addOp (.bvar 1 .int) (.bvar 0 .in (Lambda.LExpr.bvar 0))) (Lambda.LExpr.const "0" none)))))-/ #guard_msgs in -#eval eraseTy (translateBounded' test5) +#eval eraseTy (translateBoundedExpr test5) -- 6. λ (x: Nat), if foo then λ (y: Nat). not (x = -1) else λ (y: Nat). x + y >= 0 (propagate inside branches of if-then-else) --Expected: λ (x: int), if 0 <= x -> foo then λ (y: int), 0 <= y -> 0 <= x -> not (x = -1) else λ (y: int). 0 <= y -> 0 <= x -> x + y >= 0 @@ -210,7 +209,7 @@ def test6 : LExprT String := .abs (.ite (.op "foo" .bool) (.abs (notOp (.eq (.bv (Lambda.LExpr.bvar 0))) (Lambda.LExpr.const "0" none))))))-/ #guard_msgs in -#eval eraseTy (translateBounded' test6) +#eval eraseTy (translateBoundedExpr test6) end TranslateTest @@ -419,7 +418,7 @@ def testNestedBoundedApps : LExprT String := (Lambda.LExpr.app (Lambda.LExpr.op "add" none) (Lambda.LExpr.const "3" none)) (Lambda.LExpr.const "2" none)-/ #guard_msgs in -#eval eraseTy (translateBounded' testNestedBoundedApps) +#eval eraseTy (translateBoundedExpr testNestedBoundedApps) -- Test 2: Bounded variable in quantifier with complex bound expression -- Input: ∀ (x : {x < 100 ∧ 0 ≤ x}). x = 42 @@ -457,7 +456,7 @@ def testComplexBoundInQuantifier : LExprT String := (Lambda.LExpr.eq (Lambda.LExpr.bvar 0) (Lambda.LExpr.const "42" none))) -/ #guard_msgs in -#eval eraseTy (translateBounded' testComplexBoundInQuantifier) +#eval eraseTy (translateBoundedExpr testComplexBoundInQuantifier) -- Test 3: If-then-else with bounded branches and boolean condition -- Input: if (0 < (getValue 5 : {0 ≤ x})) then (1 : {x < 10}) else (0 : {x < 10}) : {x < 10} @@ -510,7 +509,7 @@ def testBoundedIte : LExprT String := (Lambda.LExpr.const "1" none) (Lambda.LExpr.const "0" none)-/ #guard_msgs in -#eval eraseTy (translateBounded' testBoundedIte) +#eval eraseTy (translateBoundedExpr testBoundedIte) -- Test 4: Lambda with bounded parameter and bounded return type -- Input: λ (x : {0 ≤ x}). increment x : {y < 100} @@ -549,7 +548,7 @@ def testBoundedLambda : LExprT String := (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) (Lambda.LExpr.app (Lambda.LExpr.op "increment" none) (Lambda.LExpr.bvar 0))-/ #guard_msgs in -#eval eraseTy (translateBounded' testBoundedLambda) +#eval eraseTy (translateBoundedExpr testBoundedLambda) -- Test 5: Equality between bounded expressions -- Input: (square (3 : {-10 ≤ x ≤ 10}) : {0 ≤ y}) = (9 : {0 ≤ z}) @@ -588,7 +587,7 @@ def testBoundedEquality : LExprT String := (Lambda.LExpr.app (Lambda.LExpr.op "square" none) (Lambda.LExpr.const "3" none)) (Lambda.LExpr.const "9" none))-/ #guard_msgs in -#eval eraseTy (translateBounded' testBoundedEquality) +#eval eraseTy (translateBoundedExpr testBoundedEquality) -- Test 6: Free variable with bounded type in assumptions -- Input: compare (x : {x < 50}) 25, with assumption [x < 50] @@ -614,7 +613,7 @@ def testFreeVarWithAssumptions : LExprT String := (Lambda.LExpr.app (Lambda.LExpr.op "compare" none) (Lambda.LExpr.fvar "x" none)) (Lambda.LExpr.const "25" none))-/ #guard_msgs in -#eval eraseTy (translateBounded' testFreeVarWithAssumptions) +#eval eraseTy (translateBoundedExpr testFreeVarWithAssumptions) -- Test 7: Metadata preservation with bounded types -- Input: @metadata(42 : {0 ≤ x < 100}) @@ -638,7 +637,7 @@ def testMetadataWithBounds : LExprT String := /-- info: Lambda.LExpr.mdata { value := "test_info" } (Lambda.LExpr.const "42" none)-/ #guard_msgs in -#eval eraseTy (translateBounded' testMetadataWithBounds) +#eval eraseTy (translateBoundedExpr testMetadataWithBounds) -- Test 8: Chain of bounded function applications -- Input: f3 (f2 (f1 5 : {x < 10}) : {x < 20}) : {x < 30} @@ -700,7 +699,7 @@ def testBoundedChain : LExprT String := (Lambda.LExpr.op "f2" none) (Lambda.LExpr.app (Lambda.LExpr.op "f1" none) (Lambda.LExpr.const "5" none)))-/ #guard_msgs in -#eval eraseTy (translateBounded' testBoundedChain) +#eval eraseTy (translateBoundedExpr testBoundedChain) end OtherTest @@ -744,7 +743,7 @@ end OtherTest -- #guard_msgs in -- #eval runWFTest testBoundedInBoolContext -- #eval (translateBounded testBoundedInBoolContext [] true).assume --- #eval eraseTy (translateBounded' testBoundedInBoolContext) +-- #eval eraseTy (translateBoundedExpr testBoundedInBoolContext) -- test: (foo x < 0 -> False) should be same as above @@ -781,7 +780,7 @@ def test1 : LExprT String := (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.const "1" none))) (Lambda.LExpr.const "0" none))) -/ #guard_msgs in -#eval eraseTy (translateBounded' test1) +#eval eraseTy (translateBoundedExpr test1) -- 1a. (foo 1 < 0) where foo: int -> nat -- Expected: 0 <= foo 1 -> foo 1 < 0 @@ -801,7 +800,7 @@ info: Lambda.LExpr.app (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.const "1" none))) (Lambda.LExpr.const "0" none))-/ #guard_msgs in -#eval eraseTy (translateBounded' test1a) +#eval eraseTy (translateBoundedExpr test1a) -- 2. not (forall (x: int), foo x < 0) for foo of same type @@ -833,7 +832,7 @@ def test2 : LExprT String := (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.bvar 0))) (Lambda.LExpr.const "0" none)))) -/ #guard_msgs in -#eval eraseTy (translateBounded' test2) +#eval eraseTy (translateBoundedExpr test2) -- 2a. (forall (x: int), foo x < 0) for foo: int -> nat -- Expected: forall (x: int), 0 <= foo x -> foo x < 0 @@ -857,7 +856,7 @@ def test2a : LExprT String := (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.bvar 0))) (Lambda.LExpr.const "0" none)))-/ #guard_msgs in -#eval eraseTy (translateBounded' test2a) +#eval eraseTy (translateBoundedExpr test2a) -- 3. (foo 1 < 0 -> false) -- Expected: 0 <= foo 1 -> (foo 1 < 0 -> false) @@ -881,7 +880,7 @@ def test3 : LExprT String := (Lambda.LExpr.const "0" none))) (Lambda.LExpr.const "false" none))-/ #guard_msgs in -#eval eraseTy (translateBounded' test3) +#eval eraseTy (translateBoundedExpr test3) -- 3a. (b -> foo 1 < 0) for boolean constant b where foo: int -> nat -- Expected: b -> 0 <= foo 1 -> foo 1 < 0 @@ -903,7 +902,7 @@ def test3a : LExprT String := (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.const "1" none))) (Lambda.LExpr.const "0" none)))-/ #guard_msgs in -#eval eraseTy (translateBounded' test3a) +#eval eraseTy (translateBoundedExpr test3a) -- 4. (\x y. x -> y) (foo x < 0) false @@ -937,7 +936,7 @@ def test4 : LExprT String := (Lambda.LExpr.const "0" none))) (Lambda.LExpr.const "false" none))-/ #guard_msgs in -#eval eraseTy (translateBounded' test4) +#eval eraseTy (translateBoundedExpr test4) -- 5. (\x. foo x < 0) 1 -- Expected: (forall x, 0 <= foo x) -> (\x. foo x < 0) 1 @@ -966,7 +965,7 @@ def test5 : LExprT String := (Lambda.LExpr.const "0" none))) (Lambda.LExpr.const "1" none)) -/ #guard_msgs in -#eval eraseTy (translateBounded' test5) +#eval eraseTy (translateBoundedExpr test5) -- 5a. (\x. foo y < 0) 1 where foo: int -> nat and y is a free variable -- Expected: 0 <= foo y -> (\x. foo y < 0) 1 @@ -991,7 +990,7 @@ def test5a : LExprT String := (Lambda.LExpr.const "0" none))) (Lambda.LExpr.const "1" none))-/ #guard_msgs in -#eval eraseTy (translateBounded' test5a) +#eval eraseTy (translateBoundedExpr test5a) -- 6. (\x. foo x) 1 >= 0 @@ -1032,7 +1031,7 @@ def test6 : LExprT String := (Lambda.LExpr.const "1" none))) (Lambda.LExpr.const "0" none)))-/ #guard_msgs in -#eval eraseTy (translateBounded' test6) +#eval eraseTy (translateBoundedExpr test6) -- 7. (if foo x < 0 then bar == 1 else baz > 0) where bar and baz have type nat -- Expected: 0 <= foo x -> if foo x < 0 then 0 <= bar -> bar == 1 else 0 <= baz -> baz > 0 @@ -1071,7 +1070,7 @@ def test7 : LExprT String := (Lambda.LExpr.app (Lambda.LExpr.op "Int.Gt" none) (Lambda.LExpr.fvar "baz" none)) (Lambda.LExpr.const "0" none))))-/ #guard_msgs in -#eval eraseTy (translateBounded' test7) +#eval eraseTy (translateBoundedExpr test7) --TODO: test and/or From 2425d9144940078fede995bda839290a9314a8cb Mon Sep 17 00:00:00 2001 From: Josh Cohen Date: Mon, 6 Oct 2025 15:45:26 -0400 Subject: [PATCH 13/19] Add back bitvector special case --- Strata/DL/Lambda/LTy.lean | 16 ++++++++++------ Strata/DL/Lambda/LTyUnify.lean | 7 +++++++ 2 files changed, 17 insertions(+), 6 deletions(-) diff --git a/Strata/DL/Lambda/LTy.lean b/Strata/DL/Lambda/LTy.lean index 89e651f9..2fd09504 100644 --- a/Strata/DL/Lambda/LTy.lean +++ b/Strata/DL/Lambda/LTy.lean @@ -52,11 +52,8 @@ deriving Inhabited, Repr, DecidableEq /-- An additional parameter that restricts the values of a type or attaches additional metadata to a type. -For now, only bitvectors are supported. -/ inductive LTyRestrict : Type where - /-- Special support for bitvector types of every size. -/ - | bitvecdata (size: Nat) /-- Bounded integers -/ | bounddata (b: BoundExpr) | nodata @@ -71,13 +68,12 @@ inductive LMonoTy (TypeRestrictions: Type := LTyRestrict) : Type where | ftvar (name : TyIdentifier) /-- Type constructor. -/ | tcons (name : String) (args : List (LMonoTy TypeRestrictions)) (r: TypeRestrictions := by exact .nodata) + /-- Special support for bitvector types of every size. -/ + | bitvec (size : Nat) deriving Inhabited, Repr abbrev LMonoTys := List LMonoTy -@[match_pattern] -def LMonoTy.bitvec (n: Nat) := LMonoTy.tcons "bitvec" [] (LTyRestrict.bitvecdata n) - @[match_pattern] def LMonoTy.bool : LMonoTy := .tcons "bool" [] @@ -168,6 +164,7 @@ types. So we define our own induction principle below. @[induction_eliminator] theorem LMonoTy.induct {Meta: Type} {P : LMonoTy Meta → Prop} (ftvar : ∀f, P (.ftvar f)) + (bitvec : ∀n, P (.bitvec n)) (tcons : ∀name args r, (∀ ty ∈ args, P ty) → P (.tcons name args r)) : ∀ ty, P ty := by intro n @@ -186,6 +183,7 @@ def LMonoTy.size (ty : LMonoTy) : Nat := match ty with | .ftvar _ => 1 | .tcons _ args _ => 1 + LMonoTys.size args + | .bitvec _ => 1 def LMonoTys.size (args : LMonoTys) : Nat := match args with @@ -205,6 +203,7 @@ Boolean equality for `LMonoTy`. def LMonoTy.BEq (x y : LMonoTy) : Bool := match x, y with | .ftvar i, .ftvar j => i == j + | .bitvec i, .bitvec j => i == j | .tcons i1 j1 r1, .tcons i2 j2 r2 => i1 == i2 && j1.length == j2.length && r1 == r2 && go j1 j2 | _, _ => false @@ -233,6 +232,8 @@ instance : DecidableEq LMonoTy := induction x generalizing y case ftvar => unfold LMonoTy.BEq at h <;> split at h <;> try simp_all + case bitvec => + unfold LMonoTy.BEq at h <;> split at h <;> try simp_all case tcons => rename_i name args r ih cases y <;> try simp_all [LMonoTy.BEq] @@ -250,6 +251,8 @@ instance : DecidableEq LMonoTy := isFalse (by induction x generalizing y case ftvar => cases y <;> try simp_all [LMonoTy.BEq] + case bitvec n => + cases y <;> try simp_all [LMonoTy.BEq] case tcons name args r ih => cases y <;> try simp_all [LMonoTy.BEq] rename_i name' args' r' @@ -279,6 +282,7 @@ it. def LMonoTy.freeVars (mty : LMonoTy) : List TyIdentifier := match mty with | .ftvar x => [x] + | .bitvec _ => [] | .tcons _ ltys _ => LMonoTys.freeVars ltys /-- diff --git a/Strata/DL/Lambda/LTyUnify.lean b/Strata/DL/Lambda/LTyUnify.lean index 79987de8..70e3dbda 100644 --- a/Strata/DL/Lambda/LTyUnify.lean +++ b/Strata/DL/Lambda/LTyUnify.lean @@ -144,6 +144,7 @@ def LMonoTy.subst (S : Subst) (mty : LMonoTy) : LMonoTy := match mty with | .ftvar x => match S.find? x with | some sty => sty | none => mty + | .bitvec _ => mty | .tcons name ltys r => .tcons name (LMonoTys.subst S ltys) r /-- @@ -218,6 +219,10 @@ theorem LMonoTy.subst_keys_not_in_substituted_type (h : SubstWF S) : have := @Map.find?_of_not_mem_values _ _ i _ S simp_all exact ne_of_mem_of_not_mem hid this + case bitvec n => + simp_all [LMonoTy.subst] + unfold LMonoTy.freeVars + simp case tcons name args r h1 => simp_all simp [subst] @@ -259,6 +264,8 @@ theorem LMonoTy.freeVars_of_subst_subset (S : Subst) (mty : LMonoTy) : apply @Map.find?_mem_values _ _ x sty _ S h_find · -- Case: S.find? x = none simp [freeVars] + case bitvec n => + simp [subst] case tcons name args r ih => simp [LMonoTy.subst, LMonoTy.freeVars] induction args From 213d91be7fade56de8027a94637081eedc5630f0 Mon Sep 17 00:00:00 2001 From: Josh Cohen Date: Mon, 6 Oct 2025 16:03:53 -0400 Subject: [PATCH 14/19] Add test for "and", move location --- .../DL/Bounded/BExprTest.lean | 207 +++++++++--------- 1 file changed, 98 insertions(+), 109 deletions(-) rename {Strata => StrataTest}/DL/Bounded/BExprTest.lean (86%) diff --git a/Strata/DL/Bounded/BExprTest.lean b/StrataTest/DL/Bounded/BExprTest.lean similarity index 86% rename from Strata/DL/Bounded/BExprTest.lean rename to StrataTest/DL/Bounded/BExprTest.lean index d1eb0d02..f6f77afa 100644 --- a/Strata/DL/Bounded/BExprTest.lean +++ b/StrataTest/DL/Bounded/BExprTest.lean @@ -21,7 +21,7 @@ def mulOp (e1 e2: LExprT String) : LExprT String := .app (.app (LFunc.opExprT in def notOp (e: LExprT String) : LExprT String := .app (LFunc.opExprT boolNotFunc) e .bool --- easier to read +-- easier to read without huge type ASTs def eraseTy (x: LExprT String) := LExpr.eraseTypes (LExprT.toLExpr x) @@ -32,7 +32,7 @@ namespace TranslateTest -- 1. ∀ (x: Nat), 0 <= x (quantified assumption) -- Expected: ∀ (x: int), 0 <= x -> 0 <= x -def test1 := (@LExprT.quant String .all natTy (.bvar 0 natTy) (leOp (.const "0" .int) (.bvar 0 .int))) +def testQuantAssume := (@LExprT.quant String .all natTy (.bvar 0 natTy) (leOp (.const "0" .int) (.bvar 0 .int))) /-- info: Lambda.LExpr.quant (Lambda.QuantifierKind.all) @@ -49,12 +49,12 @@ def test1 := (@LExprT.quant String .all natTy (.bvar 0 natTy) (leOp (.const "0" (Lambda.LExpr.bvar 0))) -/ #guard_msgs in -#eval (eraseTy (translateBoundedExpr test1) ) +#eval (eraseTy (translateBoundedExpr testQuantAssume) ) -- 2. λ(x: Nat), if 0 <= x then 1 else 2 (assumption inside ite) -- Expected: λ (x: int), if 0 <= x -> 0 <= x then 1 else 2 -def test2 : LExprT String := .abs (.ite (leOp (.const "0" .int) (.bvar 0 .int)) (.const "1" .int) (.const "2" .int) .int) (.arrow natTy .int) +def testAssumeIte : LExprT String := .abs (.ite (leOp (.const "0" .int) (.bvar 0 .int)) (.const "1" .int) (.const "2" .int) .int) (.arrow natTy .int) /-- info: Lambda.LExpr.abs (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) @@ -71,12 +71,12 @@ def test2 : LExprT String := .abs (.ite (leOp (.const "0" .int) (.bvar 0 .int)) (Lambda.LExpr.const "1" none) (Lambda.LExpr.const "2" none))-/ #guard_msgs in -#eval (eraseTy (translateBoundedExpr test2)) +#eval (eraseTy (translateBoundedExpr testAssumeIte)) -- 3. λ(x: int), if foo x >= 0 then 1 else 0 (for foo: int -> Nat) (propagate op/application information) -- Expected: λ (x: int), if 0 <= foo x -> foo x >= 0 then 1 else 0 -def test3 : LExprT String := .abs (.ite (geOp (.app (.op "foo" (.arrow .int natTy)) (.bvar 0 .int) natTy) (.const "0" .int)) (.const "1" .int) (.const "0" .int) .int) (.arrow .int .int) +def testAppAssume : LExprT String := .abs (.ite (geOp (.app (.op "foo" (.arrow .int natTy)) (.bvar 0 .int) natTy) (.const "0" .int)) (.const "1" .int) (.const "0" .int) .int) (.arrow .int .int) /-- info: Lambda.LExpr.abs (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) @@ -95,12 +95,12 @@ def test3 : LExprT String := .abs (.ite (geOp (.app (.op "foo" (.arrow .int natT (Lambda.LExpr.const "1" none) (Lambda.LExpr.const "0" none))-/ #guard_msgs in -#eval (eraseTy (translateBoundedExpr test3)) +#eval (eraseTy (translateBoundedExpr testAppAssume)) -- 4. (x: Nat) + (y: Nat) >= 0 (free variable bounds) -- Expected: 0 <= (x: int) -> 0 <= (y : int) -> x + y >= 0 -def test4 : LExprT String := geOp (addOp (.fvar "x" natTy) (.fvar "y" natTy)) (.const "0" .int) +def testFreeVar : LExprT String := geOp (addOp (.fvar "x" natTy) (.fvar "y" natTy)) (.const "0" .int) /-- info: Lambda.LExpr.app (Lambda.LExpr.app @@ -122,12 +122,12 @@ def test4 : LExprT String := geOp (addOp (.fvar "x" natTy) (.fvar "y" natTy)) (. (Lambda.LExpr.fvar "y" none))) (Lambda.LExpr.const "0" none)))-/ #guard_msgs in -#eval eraseTy (translateBoundedExpr test4) +#eval eraseTy (translateBoundedExpr testFreeVar) -- 5. λ (x: Nat). λ (y: Nat). x + y >= 0 (multiple bound vars) -- Expected: λ (x: int). λ (y: int). 0 <= y -> 0 <= x -> x + y >= 0 -def test5 : LExprT String := .abs (.abs (geOp (addOp (.bvar 1 .int) (.bvar 0 .int)) (.const "0" .int)) (.arrow natTy .bool)) (.arrow natTy (.arrow natTy .bool)) +def testMultiBvar : LExprT String := .abs (.abs (geOp (addOp (.bvar 1 .int) (.bvar 0 .int)) (.const "0" .int)) (.arrow natTy .bool)) (.arrow natTy (.arrow natTy .bool)) /-- info: Lambda.LExpr.abs (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) @@ -153,12 +153,12 @@ def test5 : LExprT String := .abs (.abs (geOp (addOp (.bvar 1 .int) (.bvar 0 .in (Lambda.LExpr.bvar 0))) (Lambda.LExpr.const "0" none)))))-/ #guard_msgs in -#eval eraseTy (translateBoundedExpr test5) +#eval eraseTy (translateBoundedExpr testMultiBvar) -- 6. λ (x: Nat), if foo then λ (y: Nat). not (x = -1) else λ (y: Nat). x + y >= 0 (propagate inside branches of if-then-else) --Expected: λ (x: int), if 0 <= x -> foo then λ (y: int), 0 <= y -> 0 <= x -> not (x = -1) else λ (y: int). 0 <= y -> 0 <= x -> x + y >= 0 -def test6 : LExprT String := .abs (.ite (.op "foo" .bool) (.abs (notOp (.eq (.bvar 1 .int) (.const "-1" .int) .bool)) (.arrow natTy .bool)) (.abs (geOp (addOp (.bvar 1 .int) (.bvar 0 .int)) (.const "0" .int)) (.arrow natTy .bool)) (.arrow natTy .bool)) (.arrow natTy (.arrow natTy .bool)) +def testBranchIte : LExprT String := .abs (.ite (.op "foo" .bool) (.abs (notOp (.eq (.bvar 1 .int) (.const "-1" .int) .bool)) (.arrow natTy .bool)) (.abs (geOp (addOp (.bvar 1 .int) (.bvar 0 .int)) (.const "0" .int)) (.arrow natTy .bool)) (.arrow natTy .bool)) (.arrow natTy (.arrow natTy .bool)) /-- info: Lambda.LExpr.abs (some (Lambda.LMonoTy.tcons "int" [] (Lambda.LTyRestrict.nodata))) @@ -209,7 +209,7 @@ def test6 : LExprT String := .abs (.ite (.op "foo" .bool) (.abs (notOp (.eq (.bv (Lambda.LExpr.bvar 0))) (Lambda.LExpr.const "0" none))))))-/ #guard_msgs in -#eval eraseTy (translateBoundedExpr test6) +#eval eraseTy (translateBoundedExpr testBranchIte) end TranslateTest @@ -225,30 +225,30 @@ namespace WFTest -- 1. constant: (1: Nat) -- Expected: 0 <= 1 -def test1 : LExprT String := .const "1" natTy +def testWfConst : LExprT String := .const "1" natTy /--info: ok: [Lambda.LExpr.app (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) (Lambda.LExpr.const "1" none)] -/ #guard_msgs in -#eval runWFTest test1 +#eval runWFTest testWfConst -- 2. application: (λ x: Nat. x) 1 -- Expected: 0 <= 1 -def test2 : LExprT String := .app (.abs (.bvar 0 .int) (.arrow natTy .int)) (.const "1" .int) .int +def testWfApp : LExprT String := .app (.abs (.bvar 0 .int) (.arrow natTy .int)) (.const "1" .int) .int /-- info: ok: [Lambda.LExpr.app (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) (Lambda.LExpr.const "1" none)]-/ #guard_msgs in -#eval runWFTest test2 +#eval runWFTest testWfApp -- 3. application with assumption (bottom up): (λ x: Nat. x) (foo : Nat) -- Expected: 0 <= foo -> 0 <= foo -def test3 : LExprT String := .app (.abs (.bvar 0 .int) (.arrow natTy .int)) (.op "foo" natTy) .int +def testWfAppAssume1 : LExprT String := .app (.abs (.bvar 0 .int) (.arrow natTy .int)) (.op "foo" natTy) .int /-- info: ok: [Lambda.LExpr.app (Lambda.LExpr.app @@ -260,12 +260,12 @@ def test3 : LExprT String := .app (.abs (.bvar 0 .int) (.arrow natTy .int)) (.op (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) (Lambda.LExpr.op "foo" none))]-/ #guard_msgs in -#eval runWFTest test3 +#eval runWFTest testWfAppAssume1 -- 4. application with assumption (top down): (λ x: Nat. (λ y: Nat. y) x) -- Expected: 0 <= x -> 0 <= x (no quantifiers) -def test4 : LExprT String := .abs (.app (.abs (.bvar 0 .int) (.arrow natTy .int)) (.bvar 0 .int) .int) (.arrow natTy .int) +def testWfAppAssume2 : LExprT String := .abs (.app (.abs (.bvar 0 .int) (.arrow natTy .int)) (.bvar 0 .int) .int) (.arrow natTy .int) /-- info: ok: [Lambda.LExpr.app (Lambda.LExpr.app @@ -277,12 +277,12 @@ def test4 : LExprT String := .abs (.app (.abs (.bvar 0 .int) (.arrow natTy .int) (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) (Lambda.LExpr.fvar "$__var0" none))]-/ #guard_msgs in -#eval runWFTest test4 +#eval runWFTest testWfAppAssume2 -- 5. abstraction with assumption: (λ x: Nat. foo (x + 1)) (foo: Nat -> int) -- Expected: 0 <= x -> 0 <= x + 1 -def test5 : LExprT String := .abs (.app (.op "foo" (.arrow natTy .int)) (addOp (.bvar 0 .int) (.const "1" .int)) .int) (.arrow natTy .int) +def testWfAbs : LExprT String := .abs (.app (.op "foo" (.arrow natTy .int)) (addOp (.bvar 0 .int) (.const "1" .int)) .int) (.arrow natTy .int) /--info: ok: [Lambda.LExpr.app (Lambda.LExpr.app @@ -296,12 +296,12 @@ def test5 : LExprT String := .abs (.app (.op "foo" (.arrow natTy .int)) (addOp ( (Lambda.LExpr.app (Lambda.LExpr.op "Int.Add" none) (Lambda.LExpr.fvar "$__var0" none)) (Lambda.LExpr.const "1" none)))]-/ #guard_msgs in -#eval runWFTest test5 +#eval runWFTest testWfAbs -- 6. quantified assumption: (∃ x: Nat. foo x) where (foo: Nat -> int) -- Expected: 0 <= x -> 0 <= x -def test6 : LExprT String := .quant .exist natTy (.bvar 0 .int) (.app (.op "foo" (.arrow natTy .int)) (.bvar 0 .int) .int) +def testWfQuantAssume : LExprT String := .quant .exist natTy (.bvar 0 .int) (.app (.op "foo" (.arrow natTy .int)) (.bvar 0 .int) .int) /-- info: ok: [Lambda.LExpr.app (Lambda.LExpr.app @@ -313,12 +313,12 @@ def test6 : LExprT String := .quant .exist natTy (.bvar 0 .int) (.app (.op "foo" (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) (Lambda.LExpr.fvar "$__var0" none))]-/ #guard_msgs in -#eval runWFTest test6 +#eval runWFTest testWfQuantAssume -- 7. Lambda with bounded body (λ x: Nat. (x + 1: Nat)) -- Expected: 0 <= x -> 0 <= x + 1 -def test7 : LExprT String := .abs (addOp (.bvar 0 .int) (.const "1" .int)) (.arrow natTy natTy) +def testWfAbsBody : LExprT String := .abs (addOp (.bvar 0 .int) (.const "1" .int)) (.arrow natTy natTy) /-- info: ok: [Lambda.LExpr.app (Lambda.LExpr.app @@ -332,22 +332,22 @@ def test7 : LExprT String := .abs (addOp (.bvar 0 .int) (.const "1" .int)) (.arr (Lambda.LExpr.app (Lambda.LExpr.op "Int.Add" none) (Lambda.LExpr.fvar "$__var0" none)) (Lambda.LExpr.const "1" none)))]-/ #guard_msgs in -#eval runWFTest test7 +#eval runWFTest testWfAbsBody -- 8. Application with bounded body: (foo x) : Nat where foo: int -> Nat -- Expected: [] (foo assumed) -def test8 : LExprT String := .app (.op "foo" (.arrow .int natTy)) (.fvar "x" .int) natTy +def testWfBoundBody : LExprT String := .app (.op "foo" (.arrow .int natTy)) (.fvar "x" .int) natTy /-- info: ok: []-/ #guard_msgs in -#eval runWFTest test8 +#eval runWFTest testWfBoundBody -- 9. Application with bounded body, no assumption: (λ (x: int) . x * x) 1 : Nat -- Expected: 0 <= (λ (x: int) . x * x) 1 -def test9 : LExprT String := .app (.abs (mulOp (.bvar 0 .int) (.bvar 0 .int)) (.arrow .int .int)) (.const "1" .int) natTy +def testWfBoundBodyNoAssume : LExprT String := .app (.abs (mulOp (.bvar 0 .int) (.bvar 0 .int)) (.arrow .int .int)) (.const "1" .int) natTy /-- info: ok: [Lambda.LExpr.app (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) @@ -359,23 +359,23 @@ def test9 : LExprT String := .app (.abs (mulOp (.bvar 0 .int) (.bvar 0 .int)) (. (Lambda.LExpr.bvar 0))) (Lambda.LExpr.const "1" none))]-/ #guard_msgs in -#eval runWFTest test9 +#eval runWFTest testWfBoundBodyNoAssume -- 10. If-then-else with bounded body: if b then 1 else 0 : Nat -- Expected: 0 <= if b then 1 else 0 -def test10 : LExprT String := .ite (.const "b" .bool) (.const "1" .int) (.const "0" .int) natTy +def testWfIteBoundBody : LExprT String := .ite (.const "b" .bool) (.const "1" .int) (.const "0" .int) natTy /-- info: ok: [Lambda.LExpr.app (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) (Lambda.LExpr.ite (Lambda.LExpr.const "b" none) (Lambda.LExpr.const "1" none) (Lambda.LExpr.const "0" none))]-/ #guard_msgs in -#eval runWFTest test10 +#eval runWFTest testWfIteBoundBody -- 11. If-then-else with bound functions: if b then λ (x: int). x * x : Nat else λ (y: int). 0 : Nat (whole type int -> Nat) -- Expected: 0 <= x * x; 0 <= 0 -def test11 : LExprT String := .ite (.const "b" .bool) (.abs (mulOp (.bvar 0 .int) (.bvar 0 .int)) (.arrow .int natTy)) (.abs (.const "0" .int) (.arrow .int natTy)) (.arrow .int natTy) +def testWfIteBoundAbs : LExprT String := .ite (.const "b" .bool) (.abs (mulOp (.bvar 0 .int) (.bvar 0 .int)) (.arrow .int natTy)) (.abs (.const "0" .int) (.arrow .int natTy)) (.arrow .int natTy) /-- info: ok: [Lambda.LExpr.app (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) @@ -385,13 +385,13 @@ def test11 : LExprT String := .ite (.const "b" .bool) (.abs (mulOp (.bvar 0 .int (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) (Lambda.LExpr.const "0" none)]-/ #guard_msgs in -#eval runWFTest test11 +#eval runWFTest testWfIteBoundAbs end WFTest ---tests with more sophisticated bounded types (mostly AI generated) +--Tests with more sophisticated bounded types (mostly AI generated) -namespace OtherTest +namespace MoreBoundTests -- Test 1: Nested bounded function applications -- Input: add : {x < 10} → {y < 5} → {z < 15}, applied to (3 : {x < 10}) and (2 : {x < 5}) @@ -701,59 +701,12 @@ def testBoundedChain : LExprT String := #guard_msgs in #eval eraseTy (translateBoundedExpr testBoundedChain) -end OtherTest +end MoreBoundTests -- Tests relating to positivity --- section PositiveTest - --- -- Test 9: Bounded type in boolean context with negation --- -- Input: ¬((getValue 10 : {0 ≤ x}) < 5) --- -- Expected: translate = (0 ≤ getValue 10) → ¬(getValue 10 < 5), wfCond = [] --- def testBoundedInBoolContext : LExprT String := --- .app (LFunc.opExprT boolNotFunc) --- (.app (.app (LFunc.opExprT intLtFunc) --- (.app (.op "getValue" (.arrow .int (.bounded (.ble (.bconst 0) (.bvar))))) --- (.const "10" .int) --- (.bounded (.ble (.bconst 0) (.bvar)))) --- (.arrow .int .bool)) --- (.const "5" .int) --- .bool) --- .bool - --- def foo : LExprT String := (.app (.app (LFunc.opExprT intLtFunc) --- (.app (.op "getValue" (.arrow .int (.bounded (.ble (.bconst 0) (.bvar))))) --- (.const "10" .int) --- (.bounded (.ble (.bconst 0) (.bvar)))) --- (.arrow .int .bool)) --- (.const "5" .int) --- .bool) - --- #eval (eraseTys (translateBounded foo [] true).assume) - - --- -- NOTE: all bottom-up assumptions need to be kept, even in boolean-valued terms, since they have to go to the top level - --- --ex: not (foo x < 0) for foo: int -> nat --- -- want (I think): 0 <= foo x -> not (foo x < 0) - TRUE --- -- we do NOT want: not (0 <= foo x -> foo x < 0) - this is FALSE --- -- think we need to collect bottom-up assumptions differently - --- /-- info: ok: []-/ --- #guard_msgs in --- #eval runWFTest testBoundedInBoolContext --- #eval (translateBounded testBoundedInBoolContext [] true).assume --- #eval eraseTy (translateBoundedExpr testBoundedInBoolContext) - --- test: (foo x < 0 -> False) should be same as above - --- not (forall (x: int), foo x < 0) - --- (λ x y. x -> y) (foo x < 0) false --- should be (forall x, foo x >= 0) -> (λ x y. x -> y) (foo x < 0) false - /- -Here we have several tests for positivity. For many, we have the main test (e.g. Test 1) as well as an auxilliary test (Test 1a) showing the ordinary (i.e. no positivity complications) version. This demonstrates that we can create better translations when we know that there are no positivity restrictions. +Here we have several tests for positivity. For many, we have the main test as well as an auxilliary test showing the ordinary (i.e. no positivity complications) version. This demonstrates that we can create better translations when we know that there are no positivity restrictions. -/ namespace PositiveTest @@ -763,7 +716,7 @@ def impliesOp (e1 e2: LExprT String) : LExprT String := .app (.app (LFunc.opExpr -- 1. not (foo 1 < 0) where foo: int -> nat -- Expected: 0 <= foo 1 -> not (foo 1 < 0) -def test1 : LExprT String := +def testPosNot : LExprT String := notOp (ltOp (.app (.op "foo" (.arrow .int natTy)) (.const "1" .int) natTy) (.const "0" .int)) /-- info: Lambda.LExpr.app @@ -780,11 +733,11 @@ def test1 : LExprT String := (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.const "1" none))) (Lambda.LExpr.const "0" none))) -/ #guard_msgs in -#eval eraseTy (translateBoundedExpr test1) +#eval eraseTy (translateBoundedExpr testPosNot) -- 1a. (foo 1 < 0) where foo: int -> nat -- Expected: 0 <= foo 1 -> foo 1 < 0 -def test1a : LExprT String := +def testPosNotAux : LExprT String := ltOp (.app (.op "foo" (.arrow .int natTy)) (.const "1" .int) natTy) (.const "0" .int) /-- @@ -800,12 +753,12 @@ info: Lambda.LExpr.app (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.const "1" none))) (Lambda.LExpr.const "0" none))-/ #guard_msgs in -#eval eraseTy (translateBoundedExpr test1a) +#eval eraseTy (translateBoundedExpr testPosNotAux) -- 2. not (forall (x: int), foo x < 0) for foo of same type -- Expected: (forall (x: int), 0 <= foo x) -> not (forall (x: int), foo x < 0) -def test2 : LExprT String := +def testPosQuant : LExprT String := notOp (.quant .all .int (.bvar 0 .int) (ltOp (.app (.op "foo" (.arrow .int natTy)) (.bvar 0 .int) natTy) (.const "0" .int))) @@ -832,11 +785,11 @@ def test2 : LExprT String := (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.bvar 0))) (Lambda.LExpr.const "0" none)))) -/ #guard_msgs in -#eval eraseTy (translateBoundedExpr test2) +#eval eraseTy (translateBoundedExpr testPosQuant) -- 2a. (forall (x: int), foo x < 0) for foo: int -> nat -- Expected: forall (x: int), 0 <= foo x -> foo x < 0 -def test2a : LExprT String := +def testPosQuantAux : LExprT String := .quant .all .int (.bvar 0 .int) (ltOp (.app (.op "foo" (.arrow .int natTy)) (.bvar 0 .int) natTy) (.const "0" .int)) @@ -856,11 +809,11 @@ def test2a : LExprT String := (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.bvar 0))) (Lambda.LExpr.const "0" none)))-/ #guard_msgs in -#eval eraseTy (translateBoundedExpr test2a) +#eval eraseTy (translateBoundedExpr testPosQuantAux) -- 3. (foo 1 < 0 -> false) -- Expected: 0 <= foo 1 -> (foo 1 < 0 -> false) -def test3 : LExprT String := +def testPosImpl : LExprT String := impliesOp (ltOp (.app (.op "foo" (.arrow .int natTy)) (.const "1" .int) natTy) (.const "0" .int)) (.const "false" .bool) @@ -880,11 +833,11 @@ def test3 : LExprT String := (Lambda.LExpr.const "0" none))) (Lambda.LExpr.const "false" none))-/ #guard_msgs in -#eval eraseTy (translateBoundedExpr test3) +#eval eraseTy (translateBoundedExpr testPosImpl) -- 3a. (b -> foo 1 < 0) for boolean constant b where foo: int -> nat -- Expected: b -> 0 <= foo 1 -> foo 1 < 0 -def test3a : LExprT String := +def testPosImplAux : LExprT String := impliesOp (.const "b" .bool) (ltOp (.app (.op "foo" (.arrow .int natTy)) (.const "1" .int) natTy) (.const "0" .int)) @@ -902,12 +855,12 @@ def test3a : LExprT String := (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.const "1" none))) (Lambda.LExpr.const "0" none)))-/ #guard_msgs in -#eval eraseTy (translateBoundedExpr test3a) +#eval eraseTy (translateBoundedExpr testPosImplAux) -- 4. (\x y. x -> y) (foo x < 0) false -- Expected: (forall x, 0 <= foo x) -> (\x y: bool. x -> y) (foo x < 0) false -def test4 : LExprT String := +def testPosImplApp : LExprT String := .app (.app (.abs (.abs (impliesOp (.bvar 1 .bool) (.bvar 0 .bool)) (.arrow .bool .bool)) (.arrow .bool (.arrow .bool .bool))) (ltOp (.app (.op "foo" (.arrow .int natTy)) (.fvar "x" .int) natTy) (.const "0" .int)) (.arrow .bool .bool)) @@ -936,11 +889,11 @@ def test4 : LExprT String := (Lambda.LExpr.const "0" none))) (Lambda.LExpr.const "false" none))-/ #guard_msgs in -#eval eraseTy (translateBoundedExpr test4) +#eval eraseTy (translateBoundedExpr testPosImplApp) -- 5. (\x. foo x < 0) 1 -- Expected: (forall x, 0 <= foo x) -> (\x. foo x < 0) 1 -def test5 : LExprT String := +def testPosApp : LExprT String := .app (.abs (ltOp (.app (.op "foo" (.arrow .int natTy)) (.bvar 0 .int) natTy) (.const "0" .int)) (.arrow .int .bool)) (.const "1" .int) .bool @@ -965,11 +918,11 @@ def test5 : LExprT String := (Lambda.LExpr.const "0" none))) (Lambda.LExpr.const "1" none)) -/ #guard_msgs in -#eval eraseTy (translateBoundedExpr test5) +#eval eraseTy (translateBoundedExpr testPosApp) -- 5a. (\x. foo y < 0) 1 where foo: int -> nat and y is a free variable -- Expected: 0 <= foo y -> (\x. foo y < 0) 1 -def test5a : LExprT String := +def testPosAppAux : LExprT String := .app (.abs (ltOp (.app (.op "foo" (.arrow .int natTy)) (.fvar "y" .int) natTy) (.const "0" .int)) (.arrow .int .bool)) (.const "1" .int) .bool @@ -990,12 +943,12 @@ def test5a : LExprT String := (Lambda.LExpr.const "0" none))) (Lambda.LExpr.const "1" none))-/ #guard_msgs in -#eval eraseTy (translateBoundedExpr test5a) +#eval eraseTy (translateBoundedExpr testPosAppAux) -- 6. (\x. foo x) 1 >= 0 -- Expected: 0 <= (\x. foo x) 1 -> (forall x, 0 <= foo x) -> (\x. foo x) 1 >= 0 -def test6 : LExprT String := +def testAbsQuant : LExprT String := geOp (.app (.abs (.app (.op "foo" (.arrow .int natTy)) (.bvar 0 .int) natTy) (.arrow .int natTy)) (.const "1" .int) natTy) @@ -1031,11 +984,11 @@ def test6 : LExprT String := (Lambda.LExpr.const "1" none))) (Lambda.LExpr.const "0" none)))-/ #guard_msgs in -#eval eraseTy (translateBoundedExpr test6) +#eval eraseTy (translateBoundedExpr testAbsQuant) -- 7. (if foo x < 0 then bar == 1 else baz > 0) where bar and baz have type nat -- Expected: 0 <= foo x -> if foo x < 0 then 0 <= bar -> bar == 1 else 0 <= baz -> baz > 0 -def test7 : LExprT String := +def testPosIte : LExprT String := .ite (ltOp (.app (.op "foo" (.arrow .int natTy)) (.fvar "x" .int) natTy) (.const "0" .int)) (.eq (.fvar "bar" natTy) (.const "1" .int) .bool) (.app (.app (LFunc.opExprT intGtFunc) (.fvar "baz" natTy) (.arrow .int .bool)) (.const "0" .int) .bool) @@ -1070,9 +1023,45 @@ def test7 : LExprT String := (Lambda.LExpr.app (Lambda.LExpr.op "Int.Gt" none) (Lambda.LExpr.fvar "baz" none)) (Lambda.LExpr.const "0" none))))-/ #guard_msgs in -#eval eraseTy (translateBoundedExpr test7) +#eval eraseTy (translateBoundedExpr testPosIte) + +-- 8. (foo x >= 0) && (bar y > 0) where foo: int -> nat and bar: int -> nat +-- Expected: (0 <= foo x -> foo x >= 0) && (0 <= bar y -> bar y > 0) +def andOp (e1 e2: LExprT String) : LExprT String := .app (.app (LFunc.opExprT boolAndFunc) e1 (.arrow .bool .bool)) e2 .bool +def gtOp (e1 e2: LExprT String) : LExprT String := .app (.app (LFunc.opExprT intGtFunc) e1 (.arrow .int .bool)) e2 .bool + +def testPosAnd : LExprT String := + andOp (geOp (.app (.op "foo" (.arrow .int natTy)) (.fvar "x" .int) natTy) (.const "0" .int)) + (gtOp (.app (.op "bar" (.arrow .int natTy)) (.fvar "y" .int) natTy) (.const "0" .int)) + +/-- info: Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.And" none) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.fvar "x" none)))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Int.Ge" none) + (Lambda.LExpr.app (Lambda.LExpr.op "foo" none) (Lambda.LExpr.fvar "x" none))) + (Lambda.LExpr.const "0" none)))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Bool.Implies" none) + (Lambda.LExpr.app + (Lambda.LExpr.app (Lambda.LExpr.op "Int.Le" none) (Lambda.LExpr.const "0" none)) + (Lambda.LExpr.app (Lambda.LExpr.op "bar" none) (Lambda.LExpr.fvar "y" none)))) + (Lambda.LExpr.app + (Lambda.LExpr.app + (Lambda.LExpr.op "Int.Gt" none) + (Lambda.LExpr.app (Lambda.LExpr.op "bar" none) (Lambda.LExpr.fvar "y" none))) + (Lambda.LExpr.const "0" none)))-/ +#guard_msgs in +#eval eraseTy (translateBoundedExpr testPosAnd) ---TODO: test and/or end PositiveTest From 8a79559df2039b52d8f02f47d1b1c6c0ea447610 Mon Sep 17 00:00:00 2001 From: Josh Cohen Date: Mon, 6 Oct 2025 17:28:02 -0400 Subject: [PATCH 15/19] Clean up BExpr, add clean translation for triggers --- Strata/DL/Bounded/BExpr.lean | 85 +++++++++++++++++++++--------------- 1 file changed, 50 insertions(+), 35 deletions(-) diff --git a/Strata/DL/Bounded/BExpr.lean b/Strata/DL/Bounded/BExpr.lean index 65e54c32..4ee390bf 100644 --- a/Strata/DL/Bounded/BExpr.lean +++ b/Strata/DL/Bounded/BExpr.lean @@ -10,30 +10,28 @@ import Strata.DL.Lambda.IntBoolFactory /-! ## Lambda Expressions with Bounded Ints -Here, we parameterize LExprTs with bounded int types. These expressions can be desugared into regular Lambda expressions by performing two steps: -1. Translating the term, keeping track of assumptions implicit in the bounded types -2. Generating well-formedness conditions e.g. for function calls +Here, we desugar LExprTs with bounded int types into regular Lambda expressions by performing two steps: +1. Translating the term, keeping track of assumptions implicit in the bounded types. +2. Generating well-formedness conditions e.g. for function calls. Most bounds must be explicitly stated; the only inferred/checked ones occur when calling a function requiring a bounded argument on a int-typed argument. -This process - -/ /- -The process of keeping track of assumptions and constraints is complicated. Here we give the key ideas (using nat = {x | 0 ≤ x} as an example): +The process of keeping track of assumptions and constraints is complicated. Here we give the key ideas (using the type nat = {x | 0 ≤ x} as an example): 1. There are two kinds of assumptions: context-dependent (e.g. for ∀ (x: nat). e, we assume 0 ≤ x when translating e) and context-independent (e.g. for external operator foo: int → nat, we are always free to assume 0 ≤ foo 1). 2. The context-dependent assumptions must be propagated top-down (when recursively translating the body of a quantifier or lambda), while the context-independent assumptions must be propagated bottom-up (as we do not know which external symbols appear in a subterm without inspecting it). -3. Assumptions can only be added to boolean-valued terms. Therefore, the basic algorithm is as follows: mantain a set of the top-down assumptions. Recursively traverse the term, removing bounded types. Whenever a bool-valued expression is reached, recursively translate the body and collect the bottom-up assumptions. Then add all assumptions to the translated body. +3. Assumptions can only be added to boolean-valued terms. Therefore, the basic algorithm is as follows: maintain a set of the top-down assumptions and recursively traverse the term, removing bounded types. Whenever a bool-valued expression is reached, recursively translate the body and collect the bottom-up assumptions. Then add all assumptions to the translated body. 4. However, we cannot freely add assumptions in non-strictly-positive positions without changing the semantics. Therefore, if in such a position, we do not add the bottom-up assumptions but rather keep propagating upward until we reach a strictly positive position (which must exist at least at the outer layer). Since it is not possible to always determine whether a position is positive without β-reduction, we are conservative and assume all unknown positions are non-positive. 5. We must be careful about binders and quantify bottom-up assumptions that escape the scope of a quantified variable due to the above positivity restriction. For example, (∃ (x: int). foo x < 0) → False is translated to (∀ (x: int). 0 ≤ foo' x) → ((∃ (x: int). foo' x < 0) → False) where foo' has type int → int -Generating well-formedness conditions is relatively simpler. The main rules are: +Generating well-formedness conditions is somewhat simpler. The main rules are: 1. Constants that claim to have a bounded type are checked (e.g. 3: nat) -2. For application e₁e₂ where e₁ has type b → T, we check that b(tr(e₂)) -3. For application e₁e₂ : b where e₁ has type T -> int, we check that the translated application satisfies b. (Note: we cannot do this in general because some bounds e.g. for external operators are assumed) -4. For λ (x : T). e : T -> b, check that e satisfies bound b (possibly assuming any bounds for x) +2. For application e₁e₂ where e₁ has type b → T, we check that b holds of the translated e₂ +3. For application e₁e₂ : b where e₁ has type T -> int, we check that the translated application satisfies b. (Note: we cannot just check bounded types in general because some bounds are assumed e.g. for external operators) +4. For λ (x : T). e : T -> b, we check that e satisfies bound b (possibly assuming any bounds for x) Proving these well-formedness conditions may require both the top-down and bottom-up assumptions, so we compute all at the same time. -See "BExprTest.lean" for test cases that further illustrate the expected translations and well-formedness conditions. +See StrataTest/DL/Bounded/BExprTest.lean for test cases that further illustrate the expected translations and well-formedness conditions. -/ @@ -43,9 +41,10 @@ open Lambda variable {Identifier : Type} [ToString Identifier] [DecidableEq Identifier] [ToFormat Identifier] [HasGen Identifier] --- Translate bounded integer expressions to LExprT - -def BoundValToLExprT (b: BoundVal) (e: LExprT Identifier) : LExprT Identifier := +/-- +Translate bounded integer expression b to LExprT with holes filled by e +-/ +def boundValToLExprT (b: BoundVal) (e: LExprT Identifier) : LExprT Identifier := match b with | .bvar => e | .bconst val => .const (val.repr) LMonoTy.int @@ -93,15 +92,15 @@ def intLeExprT [Coe String Identifier] (e1 e2: LExprT Identifier) : LExprT Ident /-- Translate BoundExpr b to corresponding Lambda expression -/ -def BoundExprToLExprT [Coe String Identifier] (b: BoundExpr) (e: LExprT Identifier) : +def boundExprToLExprT [Coe String Identifier] (b: BoundExpr) (e: LExprT Identifier) : LExprT Identifier := match b with | .beq b1 b2 => - .eq (BoundValToLExprT b1 e) (BoundValToLExprT b2 e) LMonoTy.bool - | .blt b1 b2 => intLtExprT (BoundValToLExprT b1 e) (BoundValToLExprT b2 e) - | .ble b1 b2 => intLeExprT (BoundValToLExprT b1 e) (BoundValToLExprT b2 e) - | .band e1 e2 => andExprT (BoundExprToLExprT e1 e) (BoundExprToLExprT e2 e) - | .bnot e1 => notExprT (BoundExprToLExprT e1 e) + .eq (boundValToLExprT b1 e) (boundValToLExprT b2 e) LMonoTy.bool + | .blt b1 b2 => intLtExprT (boundValToLExprT b1 e) (boundValToLExprT b2 e) + | .ble b1 b2 => intLeExprT (boundValToLExprT b1 e) (boundValToLExprT b2 e) + | .band e1 e2 => andExprT (boundExprToLExprT e1 e) (boundExprToLExprT e2 e) + | .bnot e1 => notExprT (boundExprToLExprT e1 e) -- Auxilliary functions for translation @@ -118,7 +117,7 @@ private def addAndExprs [Coe String Identifier] (l: List (LExprT Identifier)) (e List.foldr andExprT e l /- -An inefficient method of mantaining a list with no duplicates. Should be replaced with a HashSet or similar +An inefficient method of maintaining a list with no duplicates. Should be replaced with a HashSet or similar -/ def ListSet α := List α @@ -137,7 +136,7 @@ def ListSet.union {α} [DecidableEq α] (l: List (ListSet α)) : ListSet α := -- Auxiliary functions for producing bounds/assumptions /- -All top-down assumptions initially start as expressions over (.bvar 0 .int). As we pass through binders, these bvars must be increased. We only care about expressions consisting of bvar, eq, app (output of BoundExprToLExprT) +All top-down assumptions initially start as expressions over (.bvar 0 .int). As we pass through binders, these bvars must be increased. We only care about expressions consisting of bvar, eq, app (output of boundExprToLExprT) -/ private def increaseBVar (e: LExprT Identifier) : LExprT Identifier := match e with @@ -164,7 +163,7 @@ private def removeTyBound (t: LMonoTy) : LMonoTy := If ty is .bounded b, produces expression b(e) -/ private def boundExprIfType [Coe String Identifier] (ty: LMonoTy) (e: LExprT Identifier) : List (LExprT Identifier) := - ((isBounded ty).map (fun b => BoundExprToLExprT b e)).toList + ((isBounded ty).map (fun b => boundExprToLExprT b e)).toList /-- If ty is .bounded b, produces expression b(.bvar 0 .int) @@ -178,7 +177,7 @@ def wfCallCondition [Coe String Identifier] (assume : List (LExprT Identifier)) match LExprT.toLMonoTy e1 with | .arrow (.bounded b) _ => -- check that translated e2 satisfies precondition under assumptions - [addAssumptions assume (BoundExprToLExprT b e2)] + [addAssumptions assume (boundExprToLExprT b e2)] | _ => [] /-- @@ -223,6 +222,22 @@ def conditionalAdd (b : Bool) (base extra : List α) : List α := def conditionalReturn (pos : Bool) (value : List α) : List α := if pos then [] else value +/-- +A "clean" translation that only removes types; used for triggers +-/ +private def translateClean [Coe String Identifier] (e: LExprT Identifier) : LExprT Identifier := + match e with + | .const c ty => .const c (removeTyBound ty) + | .op o ty => .op o (removeTyBound ty) + | .bvar b ty => .bvar b (removeTyBound ty) + | .fvar f ty => .fvar f (removeTyBound ty) + | .app e1 e2 ty => .app (translateClean e1) (translateClean e2) (removeTyBound ty) + | .abs e ty => .abs (translateClean e) (removeTyBound ty) + | .quant qk ty tr e' => .quant qk (removeTyBound ty) (translateClean tr) (translateClean e') + | .ite c t f ty => .ite (translateClean c) (translateClean t) (translateClean f) (removeTyBound ty) + | .eq e1 e2 ty => .eq (translateClean e1) (translateClean e2) (removeTyBound ty) + | .mdata m e => .mdata m (translateClean e) + structure translationRes Identifier where (translate : LExprT Identifier) (wfCond : ListSet (LExprT Identifier)) @@ -280,14 +295,14 @@ def translateBounded [Coe String Identifier] [DecidableEq Identifier] (e: LExprT | .app e1 e2 ty => let e1' := translateBounded e1 assume pos; let e2' := translateBounded e2 assume pos; - let e' := .app e1'.translate e2'.translate (removeTyBound ty); + let res := .app e1'.translate e2'.translate (removeTyBound ty); -- If we have application f x where f : _ -> int and f x has bounded type, need wf condition that application result is bounded let extraWf := match LExprT.toLMonoTy e1, ty with | .arrow _ .int, .bounded _ => boundExprIfType ty e' | _, _ => []; - ⟨e', ListSet.union [wfCallCondition (assume ++ e2'.assume) e1 e2'.translate, extraWf, e1'.wfCond, e2'.wfCond], boundExprIfType ty e' ++ e1'.assume ++ e2'.assume⟩ + ⟨res, ListSet.union [wfCallCondition (assume ++ e2'.assume) e1 e2'.translate, extraWf, e1'.wfCond, e2'.wfCond], boundExprIfType ty e' ++ e1'.assume ++ e2'.assume⟩ /- Lambda abstraction: 1. If the argument is bounded, add as top-down assumption @@ -302,24 +317,24 @@ def translateBounded [Coe String Identifier] [DecidableEq Identifier] (e: LExprT | .abs e (.arrow ty1 ty2) => let all_assume := makeBoundAssumption ty1 ++ (increaseBVars assume); let e' := translateBounded e all_assume pos; - let e'' := .abs e'.translate (removeTyBound (.arrow ty1 ty2)); + let res := .abs e'.translate (removeTyBound (.arrow ty1 ty2)); -- Note: don't add assumptions to e'.wfCond, already included - ⟨e'', addBoundedWf all_assume (boundExprIfType ty2 e'.translate) e'.wfCond, List.map (quantifyAssumptions ty1) e'.assume⟩ + ⟨res, addBoundedWf all_assume (boundExprIfType ty2 e'.translate) e'.wfCond, List.map (quantifyAssumptions ty1) e'.assume⟩ | .abs _ _ => ⟨.const "0" .int, [], []⟩ -- error case /- Quantifiers are simpler because they are boolean-valued. ∀ (x : nat). e adds an assumption x >= 0 -> ..., while ∃ (x: nat). e results in x >= 0 ∧ .. -/ | .quant .all ty tr e => let e' := translateBounded e [] pos; - let tr' := translateBounded tr [] pos; --TODO: need "clean" one here + let tr' := translateClean tr; let all_assume := conditionalAdd pos (makeBoundAssumption ty ++ (increaseBVars assume)) e'.assume; - ⟨.quant .all (removeTyBound ty) tr'.translate (addAssumptions all_assume e'.translate), addBoundedWfAssume all_assume e'.wfCond, conditionalReturn pos (List.map (quantifyAssumptions ty) e'.assume)⟩ + ⟨.quant .all (removeTyBound ty) tr' (addAssumptions all_assume e'.translate), addBoundedWfAssume all_assume e'.wfCond, conditionalReturn pos (List.map (quantifyAssumptions ty) e'.assume)⟩ | .quant .exist ty tr e => let newAssumption := makeBoundAssumption ty; let e' := translateBounded e [] pos; - let tr' := translateBounded tr [] pos; --TODO: need "clean" one here + let tr' := translateClean tr; let all_assume := conditionalAdd pos (increaseBVars assume) e'.assume; - ⟨.quant .exist (removeTyBound ty) tr'.translate (addAssumptions all_assume (addAndExprs newAssumption e'.translate)), addBoundedWfAssume (newAssumption ++ all_assume) e'.wfCond, conditionalReturn pos (List.map (quantifyAssumptions ty) e'.assume)⟩ + ⟨.quant .exist (removeTyBound ty) tr' (addAssumptions all_assume (addAndExprs newAssumption e'.translate)), addBoundedWfAssume (newAssumption ++ all_assume) e'.wfCond, conditionalReturn pos (List.map (quantifyAssumptions ty) e'.assume)⟩ /- For if-then-else, add assumptions to the condition, which is always bool-valued. For a bool-valued result, can add the conditions freely. For a bounded-valued term, produce a wf condition proving this. -/ @@ -335,8 +350,8 @@ def translateBounded [Coe String Identifier] [DecidableEq Identifier] (e: LExprT let t' := translateBounded t assume pos; let f' := translateBounded f assume pos; -- here can add inside if positive, never add outside - let e' := .ite (addAssumptions (conditionalAdd pos assume c'.assume) c'.translate) t'.translate f'.translate ty; - ⟨e', ListSet.union [boundExprIfType ty e', c'.wfCond, t'.wfCond, f'.wfCond], conditionalReturn pos (c'.assume ++ t'.assume ++ f'.assume)⟩ + let res := .ite (addAssumptions (conditionalAdd pos assume c'.assume) c'.translate) t'.translate f'.translate ty; + ⟨res, ListSet.union [boundExprIfType ty res, c'.wfCond, t'.wfCond, f'.wfCond], conditionalReturn pos (c'.assume ++ t'.assume ++ f'.assume)⟩ /- Equality is always bool-valued, so can add assumptions Equality/iff are equivalent, NOT positive From da61c4c179fdb6dbeaaa4b08c882e4178e8aee06 Mon Sep 17 00:00:00 2001 From: Josh Cohen Date: Mon, 6 Oct 2025 17:34:19 -0400 Subject: [PATCH 16/19] Fix changes from Lambda that had been introduced earlier --- Strata/DL/Bounded/BExpr.lean | 4 ++-- Strata/DL/Lambda/LExprT.lean | 4 ++-- Strata/DL/Lambda/LExprTypeEnv.lean | 4 ++-- Strata/DL/Lambda/LTy.lean | 14 +++++++------- 4 files changed, 13 insertions(+), 13 deletions(-) diff --git a/Strata/DL/Bounded/BExpr.lean b/Strata/DL/Bounded/BExpr.lean index 4ee390bf..89f76dd2 100644 --- a/Strata/DL/Bounded/BExpr.lean +++ b/Strata/DL/Bounded/BExpr.lean @@ -300,9 +300,9 @@ def translateBounded [Coe String Identifier] [DecidableEq Identifier] (e: LExprT let extraWf := match LExprT.toLMonoTy e1, ty with | .arrow _ .int, .bounded _ => - boundExprIfType ty e' + boundExprIfType ty res | _, _ => []; - ⟨res, ListSet.union [wfCallCondition (assume ++ e2'.assume) e1 e2'.translate, extraWf, e1'.wfCond, e2'.wfCond], boundExprIfType ty e' ++ e1'.assume ++ e2'.assume⟩ + ⟨res, ListSet.union [wfCallCondition (assume ++ e2'.assume) e1 e2'.translate, extraWf, e1'.wfCond, e2'.wfCond], boundExprIfType ty res ++ e1'.assume ++ e2'.assume⟩ /- Lambda abstraction: 1. If the argument is bounded, add as top-down assumption diff --git a/Strata/DL/Lambda/LExprT.lean b/Strata/DL/Lambda/LExprT.lean index 1148c6f9..c9e8f00d 100644 --- a/Strata/DL/Lambda/LExprT.lean +++ b/Strata/DL/Lambda/LExprT.lean @@ -278,8 +278,8 @@ def inferConst (T : (TEnv Identifier)) (c : String) (cty : Option LMonoTy) : {@LExpr.const LMonoTy Identifier c cty}\n\ Known Types: {T.knownTypes}" -- Annotated BitVecs - | c, some (.bitvec n) => - let ty := .bitvec n + | c, some (LMonoTy.bitvec n) => + let ty := LMonoTy.bitvec n if { name := "bitvec", arity := 1 } ∈ T.knownTypes then (.ok (ty, T)) else diff --git a/Strata/DL/Lambda/LExprTypeEnv.lean b/Strata/DL/Lambda/LExprTypeEnv.lean index ee2dd447..75724987 100644 --- a/Strata/DL/Lambda/LExprTypeEnv.lean +++ b/Strata/DL/Lambda/LExprTypeEnv.lean @@ -192,8 +192,8 @@ def KnownType.toLTy (k : KnownType) : LTy := def LTy.toKnownType! (lty : LTy) : KnownType := match lty with - | .forAll [] (.bitvec _) => { name := "bitvec", arity := 1 } | .forAll _ (.tcons name args) => { name, arity := args.length } + | .forAll [] (.bitvec _) => { name := "bitvec", arity := 1 } | _ => panic! s!"Unsupported known type: {lty}" instance : ToFormat KnownType where @@ -451,7 +451,7 @@ def LMonoTy.aliasDef? (mty : LMonoTy) (T : (TEnv Identifier)) : (Option LMonoTy -- We can't have a free variable be the LHS of an alias definition because -- then it will unify with every type. (none, T) - | bitvec _ => + | .bitvec _ => -- A bitvector cannot be a type alias. (none, T) | .tcons name args r => diff --git a/Strata/DL/Lambda/LTy.lean b/Strata/DL/Lambda/LTy.lean index 2fd09504..6bc6a2b1 100644 --- a/Strata/DL/Lambda/LTy.lean +++ b/Strata/DL/Lambda/LTy.lean @@ -92,23 +92,23 @@ def LMonoTy.bounded (b: BoundExpr) : LMonoTy := @[match_pattern] def LMonoTy.bv1 : LMonoTy := - bitvec 1 + .bitvec 1 @[match_pattern] def LMonoTy.bv8 : LMonoTy := - bitvec 8 + .bitvec 8 @[match_pattern] def LMonoTy.bv16 : LMonoTy := - bitvec 16 + .bitvec 16 @[match_pattern] def LMonoTy.bv32 : LMonoTy := - bitvec 32 + .bitvec 32 @[match_pattern] def LMonoTy.bv64 : LMonoTy := - bitvec 64 + .bitvec 64 @[match_pattern] def LMonoTy.string : LMonoTy := @@ -162,7 +162,7 @@ Lean's default `induction` tactic does not support nested or mutual inductive types. So we define our own induction principle below. -/ @[induction_eliminator] -theorem LMonoTy.induct {Meta: Type} {P : LMonoTy Meta → Prop} +theorem LMonoTy.induct {R: Type} {P : LMonoTy R → Prop} (ftvar : ∀f, P (.ftvar f)) (bitvec : ∀n, P (.bitvec n)) (tcons : ∀name args r, (∀ ty ∈ args, P ty) → P (.tcons name args r)) : @@ -305,7 +305,7 @@ Get all type constructors in monotype `mty`. def LMonoTy.getTyConstructors (mty : LMonoTy) : List LMonoTy := match mty with | .ftvar _ => [] - | bitvec _ => [] + | .bitvec _ => [] | .tcons name mtys _ => let typeargs := List.replicate mtys.length "_dummy" let args := typeargs.mapIdx (fun i elem => LMonoTy.ftvar (elem ++ toString i)) From f01ba99839fd95a309909baf5dadf16fc35b7cc1 Mon Sep 17 00:00:00 2001 From: Josh Cohen Date: Tue, 7 Oct 2025 11:33:47 -0400 Subject: [PATCH 17/19] Add bitvector to LTyRestrict Still special case but need this way so that we can have polymorphic non-restricted type applications --- Strata/DL/Lambda/LExpr.lean | 178 ++++++++++++----------- Strata/DL/Lambda/LTy.lean | 84 ++++++----- Strata/DL/Lambda/LTyUnify.lean | 7 - Strata/Languages/Boogie/Identifiers.lean | 2 +- 4 files changed, 136 insertions(+), 135 deletions(-) diff --git a/Strata/DL/Lambda/LExpr.lean b/Strata/DL/Lambda/LExpr.lean index 5e4e9ff9..ff5e61bb 100644 --- a/Strata/DL/Lambda/LExpr.lean +++ b/Strata/DL/Lambda/LExpr.lean @@ -322,51 +322,51 @@ scoped syntax lconstmono : lexprmono def elabLConstMono (Identifier : Type) [MkIdent Identifier] : Lean.Syntax → MetaM Expr | `(lconstmono| #$n:num) => do - let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit (toString n.getNat), none] | `(lconstmono| (#$n:num : $ty:lmonoty)) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy ty - let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) lmonoty - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) lmonoty + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit (toString n.getNat), lmonoty] | `(lconstmono| #-$n:num) => do - let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit ("-" ++ (toString n.getNat)), none] | `(lconstmono| (#-$n:num : $ty:lmonoty)) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy ty - let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) lmonoty - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) lmonoty + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit ("-" ++ (toString n.getNat)), lmonoty] | `(lconstmono| #true) => do - let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit "true", none] | `(lconstmono| (#true : $ty:lmonoty)) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy ty - let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) lmonoty - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) lmonoty + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit "true", lmonoty] | `(lconstmono| #false) => do - let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit "false", none] | `(lconstmono| (#false : $ty:lmonoty)) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy ty - let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) lmonoty - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) lmonoty + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit "false", lmonoty] | `(lconstmono| #$s:ident) => do - let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) let s := toString s.getId - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit s, none] | `(lconstmono| (#$s:ident : $ty:lmonoty)) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy ty - let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) lmonoty + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) lmonoty let s := toString s.getId - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) return mkAppN (.const ``LExpr.const []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkStrLit s, lmonoty] | _ => throwUnsupportedSyntax @@ -377,13 +377,13 @@ scoped syntax lopmono : lexprmono def elabLOpMono (Identifier : Type) [MkIdent Identifier] : Lean.Syntax → MetaM Expr | `(lopmono| ~$s:lidentmono) => do - let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) let ident ← MkIdent.elabIdent Identifier s - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) return mkAppN (.const ``LExpr.op []) #[typeTypeExpr, MkIdent.toExpr Identifier, ident, none] | `(lopmono| (~$s:lidentmono : $ty:lmonoty)) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy ty - let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) lmonoty + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) lmonoty mkAppM ``LExpr.op #[← MkIdent.elabIdent Identifier s, lmonoty] | _ => throwUnsupportedSyntax @@ -391,7 +391,7 @@ declare_syntax_cat lbvarmono scoped syntax "%" noWs num : lbvarmono def elabLBVarMono (Identifier : Type) [MkIdent Identifier] : Lean.Syntax → MetaM Expr | `(lbvarmono| %$n:num) => - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) return mkAppN (.const ``LExpr.bvar []) #[typeTypeExpr, MkIdent.toExpr Identifier, mkNatLit n.getNat] | _ => throwUnsupportedSyntax scoped syntax lbvarmono : lexprmono @@ -402,11 +402,11 @@ scoped syntax "(" lidentmono ":" lmonoty ")" : lfvarmono def elabLFVarMono (Identifier : Type) [MkIdent Identifier] : Lean.Syntax → MetaM Expr | `(lfvarmono| $i:lidentmono) => do - let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let none ← mkNone (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) mkAppM ``LExpr.fvar #[← MkIdent.elabIdent Identifier i, none] | `(lfvarmono| ($i:lidentmono : $ty:lmonoty)) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy ty - let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) lmonoty + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) lmonoty mkAppM ``LExpr.fvar #[← MkIdent.elabIdent Identifier i, lmonoty] | _ => throwUnsupportedSyntax scoped syntax lfvarmono : lexprmono @@ -455,48 +455,48 @@ partial def elabLExprMono (Identifier : Type) [MkIdent Identifier] : Lean.Syntax | `(lexprmono| $f:lfvarmono) => elabLFVarMono Identifier f | `(lexprmono| λ $e:lexprmono) => do let e' ← elabLExprMono Identifier e - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) return mkAppN (.const ``LExpr.absUntyped []) #[typeTypeExpr, MkIdent.toExpr Identifier, e'] | `(lexprmono| λ ($mty:lmonoty): $e:lexprmono) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy mty - let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) lmonoty + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) lmonoty let e' ← elabLExprMono Identifier e - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) return mkAppN (.const ``LExpr.abs []) #[typeTypeExpr, MkIdent.toExpr Identifier, lmonoty, e'] | `(lexprmono| ∀ $e:lexprmono) => do let e' ← elabLExprMono Identifier e - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) return mkAppN (.const ``LExpr.allUntyped []) #[typeTypeExpr, MkIdent.toExpr Identifier, e'] | `(lexprmono| ∀ {$tr}$e:lexprmono) => do let e' ← elabLExprMono Identifier e let tr' ← elabLExprMono Identifier tr - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) return mkAppN (.const ``LExpr.allUntypedTr []) #[typeTypeExpr, MkIdent.toExpr Identifier, tr', e'] | `(lexprmono| ∀ ($mty:lmonoty): $e:lexprmono) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy mty - let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) lmonoty + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) lmonoty let e' ← elabLExprMono Identifier e - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) return mkAppN (.const ``LExpr.all []) #[typeTypeExpr, MkIdent.toExpr Identifier, lmonoty, e'] | `(lexprmono| ∀ ($mty:lmonoty):{$tr} $e:lexprmono) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy mty - let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) lmonoty + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) lmonoty let e' ← elabLExprMono Identifier e let tr' ← elabLExprMono Identifier tr - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) return mkAppN (.const ``LExpr.allTr []) #[typeTypeExpr, MkIdent.toExpr Identifier, lmonoty, tr', e'] | `(lexprmono| ∃ ($mty:lmonoty): $e:lexprmono) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy mty - let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) lmonoty + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) lmonoty let e' ← elabLExprMono Identifier e - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) return mkAppN (.const ``LExpr.exist []) #[typeTypeExpr, MkIdent.toExpr Identifier, lmonoty, e'] | `(lexprmono| ∃ ($mty:lmonoty):{$tr} $e:lexprmono) => do let lmonoty ← Lambda.LTy.Syntax.elabLMonoTy mty - let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) lmonoty + let lmonoty ← mkSome (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) lmonoty let e' ← elabLExprMono Identifier e let tr' ← elabLExprMono Identifier tr - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) return mkAppN (.const ``LExpr.existTr []) #[typeTypeExpr, MkIdent.toExpr Identifier, lmonoty, tr', e'] | `(lexprmono| ∃ $e:lexprmono) => do let e' ← elabLExprMono Identifier e @@ -508,18 +508,18 @@ partial def elabLExprMono (Identifier : Type) [MkIdent Identifier] : Lean.Syntax | `(lexprmono| ($e1:lexprmono $e2:lexprmono)) => do let e1' ← elabLExprMono Identifier e1 let e2' ← elabLExprMono Identifier e2 - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) return mkAppN (.const ``LExpr.app []) #[typeTypeExpr, MkIdent.toExpr Identifier, e1', e2'] | `(lexprmono| $e1:lexprmono == $e2:lexprmono) => do let e1' ← elabLExprMono Identifier e1 let e2' ← elabLExprMono Identifier e2 - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) return mkAppN (.const ``LExpr.eq []) #[typeTypeExpr, MkIdent.toExpr Identifier, e1', e2'] | `(lexprmono| if $e1:lexprmono then $e2:lexprmono else $e3:lexprmono) => do let e1' ← elabLExprMono Identifier e1 let e2' ← elabLExprMono Identifier e2 let e3' ← elabLExprMono Identifier e3 - let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) + let typeTypeExpr := (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) return mkAppN (.const ``LExpr.ite []) #[typeTypeExpr, MkIdent.toExpr Identifier, e1', e2', e3'] | `(lexprmono| ($e:lexprmono)) => elabLExprMono Identifier e | _ => throwUnsupportedSyntax @@ -544,7 +544,7 @@ open LTy.Syntax #guard_msgs in #check esM[((λ %0) #5)] -/-- info: (abs (some (LMonoTy.tcons LTyRestrict "bool" [] LTyRestrict.nodata)) (bvar 0)).app +/-- info: (abs (some (LMonoTy.tcons BoundTyRestrict "bool" [] LTyRestrict.nodata)) (bvar 0)).app (const "true" none) : LExpr LMonoTy String -/ #guard_msgs in #check esM[((λ (bool): %0) #true)] @@ -557,12 +557,12 @@ open LTy.Syntax #guard_msgs in #check esM[(∃ %0 == #5)] -/-- info: exist (some (LMonoTy.tcons LTyRestrict "int" [] LTyRestrict.nodata)) - ((bvar 0).eq (const "5" none)) : LExpr LMonoTy String -/ +/-- info: exist (some (LMonoTy.tcons BoundTyRestrict "int" [] LTyRestrict.nodata)) + ((bvar 0).eq (const "5" none)) : LExpr LMonoTy String-/ #guard_msgs in #check esM[(∃ (int): %0 == #5)] -/-- info: fvar "x" (some (LMonoTy.tcons LTyRestrict "bool" [] LTyRestrict.nodata)) : LExpr LMonoTy String -/ +/-- info: fvar "x" (some (LMonoTy.tcons BoundTyRestrict "bool" [] LTyRestrict.nodata)) : LExpr LMonoTy String -/ #guard_msgs in #check esM[(x : bool)] @@ -570,31 +570,33 @@ open LTy.Syntax /-- info: all (some - (LMonoTy.tcons LTyRestrict "Map" [LMonoTy.ftvar LTyRestrict "k", LMonoTy.ftvar LTyRestrict "v"] LTyRestrict.nodata)) - (all (some (LMonoTy.ftvar LTyRestrict "k")) - (all (some (LMonoTy.ftvar LTyRestrict "v")) + (LMonoTy.tcons BoundTyRestrict "Map" [LMonoTy.ftvar BoundTyRestrict "k", LMonoTy.ftvar BoundTyRestrict "v"] + LTyRestrict.nodata)) + (all (some (LMonoTy.ftvar BoundTyRestrict "k")) + (all (some (LMonoTy.ftvar BoundTyRestrict "v")) ((((op "select" (some - (LMonoTy.tcons LTyRestrict "Map" - [LMonoTy.ftvar LTyRestrict "k", - LMonoTy.tcons LTyRestrict "arrow" - [LMonoTy.ftvar LTyRestrict "v", - LMonoTy.tcons LTyRestrict "arrow" - [LMonoTy.ftvar LTyRestrict "k", LMonoTy.ftvar LTyRestrict "v"] LTyRestrict.nodata] + (LMonoTy.tcons BoundTyRestrict "Map" + [LMonoTy.ftvar BoundTyRestrict "k", + LMonoTy.tcons BoundTyRestrict "arrow" + [LMonoTy.ftvar BoundTyRestrict "v", + LMonoTy.tcons BoundTyRestrict "arrow" + [LMonoTy.ftvar BoundTyRestrict "k", LMonoTy.ftvar BoundTyRestrict "v"] + LTyRestrict.nodata] LTyRestrict.nodata] LTyRestrict.nodata))).app ((((op "update" (some - (LMonoTy.tcons LTyRestrict "Map" - [LMonoTy.ftvar LTyRestrict "k", - LMonoTy.tcons LTyRestrict "arrow" - [LMonoTy.ftvar LTyRestrict "v", - LMonoTy.tcons LTyRestrict "arrow" - [LMonoTy.ftvar LTyRestrict "k", - LMonoTy.tcons LTyRestrict "arrow" - [LMonoTy.ftvar LTyRestrict "v", - LMonoTy.tcons LTyRestrict "Map" - [LMonoTy.ftvar LTyRestrict "k", LMonoTy.ftvar LTyRestrict "v"] + (LMonoTy.tcons BoundTyRestrict "Map" + [LMonoTy.ftvar BoundTyRestrict "k", + LMonoTy.tcons BoundTyRestrict "arrow" + [LMonoTy.ftvar BoundTyRestrict "v", + LMonoTy.tcons BoundTyRestrict "arrow" + [LMonoTy.ftvar BoundTyRestrict "k", + LMonoTy.tcons BoundTyRestrict "arrow" + [LMonoTy.ftvar BoundTyRestrict "v", + LMonoTy.tcons BoundTyRestrict "Map" + [LMonoTy.ftvar BoundTyRestrict "k", LMonoTy.ftvar BoundTyRestrict "v"] LTyRestrict.nodata] LTyRestrict.nodata] LTyRestrict.nodata] @@ -872,7 +874,7 @@ open LTy.Syntax #guard_msgs in #check es[((λ %0) #5)] -/-- info: (abs (some (LTy.forAll [] (LMonoTy.tcons LTyRestrict "bool" [] LTyRestrict.nodata))) (bvar 0)).app +/-- info: (abs (some (LTy.forAll [] (LMonoTy.tcons BoundTyRestrict "bool" [] LTyRestrict.nodata))) (bvar 0)).app (const "true" none) : LExpr LTy String -/ #guard_msgs in #check es[((λ (bool): %0) #true)] @@ -885,12 +887,12 @@ open LTy.Syntax #guard_msgs in #check es[(∃ %0 == #5)] -/-- info: exist (some (LTy.forAll [] (LMonoTy.tcons LTyRestrict "int" [] LTyRestrict.nodata))) +/-- info: exist (some (LTy.forAll [] (LMonoTy.tcons BoundTyRestrict "int" [] LTyRestrict.nodata))) ((bvar 0).eq (const "5" none)) : LExpr LTy String -/ #guard_msgs in #check es[(∃ (int): %0 == #5)] -/-- info: fvar "x" (some (LTy.forAll [] (LMonoTy.tcons LTyRestrict "bool" [] LTyRestrict.nodata))) : LExpr LTy String -/ +/-- info: fvar "x" (some (LTy.forAll [] (LMonoTy.tcons BoundTyRestrict "bool" [] LTyRestrict.nodata))) : LExpr LTy String -/ #guard_msgs in #check es[(x : bool)] @@ -899,34 +901,36 @@ open LTy.Syntax info: all (some (LTy.forAll [] - (LMonoTy.tcons LTyRestrict "Map" [LMonoTy.ftvar LTyRestrict "k", LMonoTy.ftvar LTyRestrict "v"] + (LMonoTy.tcons BoundTyRestrict "Map" [LMonoTy.ftvar BoundTyRestrict "k", LMonoTy.ftvar BoundTyRestrict "v"] LTyRestrict.nodata))) - (all (some (LTy.forAll [] (LMonoTy.ftvar LTyRestrict "k"))) - (all (some (LTy.forAll [] (LMonoTy.ftvar LTyRestrict "v"))) + (all (some (LTy.forAll [] (LMonoTy.ftvar BoundTyRestrict "k"))) + (all (some (LTy.forAll [] (LMonoTy.ftvar BoundTyRestrict "v"))) ((((op "select" (some (LTy.forAll [] - (LMonoTy.tcons LTyRestrict "Map" - [LMonoTy.ftvar LTyRestrict "k", - LMonoTy.tcons LTyRestrict "arrow" - [LMonoTy.ftvar LTyRestrict "v", - LMonoTy.tcons LTyRestrict "arrow" - [LMonoTy.ftvar LTyRestrict "k", LMonoTy.ftvar LTyRestrict "v"] LTyRestrict.nodata] + (LMonoTy.tcons BoundTyRestrict "Map" + [LMonoTy.ftvar BoundTyRestrict "k", + LMonoTy.tcons BoundTyRestrict "arrow" + [LMonoTy.ftvar BoundTyRestrict "v", + LMonoTy.tcons BoundTyRestrict "arrow" + [LMonoTy.ftvar BoundTyRestrict "k", LMonoTy.ftvar BoundTyRestrict "v"] + LTyRestrict.nodata] LTyRestrict.nodata] LTyRestrict.nodata)))).app ((((op "update" (some (LTy.forAll [] - (LMonoTy.tcons LTyRestrict "Map" - [LMonoTy.ftvar LTyRestrict "k", - LMonoTy.tcons LTyRestrict "arrow" - [LMonoTy.ftvar LTyRestrict "v", - LMonoTy.tcons LTyRestrict "arrow" - [LMonoTy.ftvar LTyRestrict "k", - LMonoTy.tcons LTyRestrict "arrow" - [LMonoTy.ftvar LTyRestrict "v", - LMonoTy.tcons LTyRestrict "Map" - [LMonoTy.ftvar LTyRestrict "k", LMonoTy.ftvar LTyRestrict "v"] + (LMonoTy.tcons BoundTyRestrict "Map" + [LMonoTy.ftvar BoundTyRestrict "k", + LMonoTy.tcons BoundTyRestrict "arrow" + [LMonoTy.ftvar BoundTyRestrict "v", + LMonoTy.tcons BoundTyRestrict "arrow" + [LMonoTy.ftvar BoundTyRestrict "k", + LMonoTy.tcons BoundTyRestrict "arrow" + [LMonoTy.ftvar BoundTyRestrict "v", + LMonoTy.tcons BoundTyRestrict "Map" + [LMonoTy.ftvar BoundTyRestrict "k", + LMonoTy.ftvar BoundTyRestrict "v"] LTyRestrict.nodata] LTyRestrict.nodata] LTyRestrict.nodata] diff --git a/Strata/DL/Lambda/LTy.lean b/Strata/DL/Lambda/LTy.lean index 6bc6a2b1..90842bbe 100644 --- a/Strata/DL/Lambda/LTy.lean +++ b/Strata/DL/Lambda/LTy.lean @@ -49,13 +49,19 @@ inductive BoundExpr : Type where | bnot (e: BoundExpr) deriving Inhabited, Repr, DecidableEq -/-- +/- An additional parameter that restricts the values of a type or attaches additional metadata to a type. -/ -inductive LTyRestrict : Type where +inductive BoundTyRestrict : Type where /-- Bounded integers -/ | bounddata (b: BoundExpr) +deriving Inhabited, Repr, DecidableEq + +inductive LTyRestrict (ExtraRestrict : Type) : Type where + /-- Special support for bitvector types of every size. -/ + | bitvecdata (size: Nat) + | extradata (e: ExtraRestrict) | nodata deriving Inhabited, Repr, DecidableEq @@ -63,17 +69,18 @@ deriving Inhabited, Repr, DecidableEq Types in Lambda: these are mono-types. Note that all free type variables (`.ftvar`) are implicitly universally quantified. -/ -inductive LMonoTy (TypeRestrictions: Type := LTyRestrict) : Type where +inductive LMonoTy (ExtraRestrict: Type := BoundTyRestrict) : Type where /-- Type variable. -/ | ftvar (name : TyIdentifier) /-- Type constructor. -/ - | tcons (name : String) (args : List (LMonoTy TypeRestrictions)) (r: TypeRestrictions := by exact .nodata) - /-- Special support for bitvector types of every size. -/ - | bitvec (size : Nat) + | tcons (name : String) (args : List (LMonoTy ExtraRestrict)) (r: LTyRestrict ExtraRestrict := by exact .nodata) deriving Inhabited, Repr abbrev LMonoTys := List LMonoTy +@[match_pattern] +def LMonoTy.bitvec (n: Nat) : LMonoTy := LMonoTy.tcons "bitvec" [] (LTyRestrict.bitvecdata n) + @[match_pattern] def LMonoTy.bool : LMonoTy := .tcons "bool" [] @@ -88,7 +95,7 @@ def LMonoTy.real : LMonoTy := @[match_pattern] def LMonoTy.bounded (b: BoundExpr) : LMonoTy := - .tcons "int" [] (LTyRestrict.bounddata b) + .tcons "int" [] (.extradata (.bounddata b)) @[match_pattern] def LMonoTy.bv1 : LMonoTy := @@ -164,7 +171,6 @@ types. So we define our own induction principle below. @[induction_eliminator] theorem LMonoTy.induct {R: Type} {P : LMonoTy R → Prop} (ftvar : ∀f, P (.ftvar f)) - (bitvec : ∀n, P (.bitvec n)) (tcons : ∀name args r, (∀ ty ∈ args, P ty) → P (.tcons name args r)) : ∀ ty, P ty := by intro n @@ -183,7 +189,6 @@ def LMonoTy.size (ty : LMonoTy) : Nat := match ty with | .ftvar _ => 1 | .tcons _ args _ => 1 + LMonoTys.size args - | .bitvec _ => 1 def LMonoTys.size (args : LMonoTys) : Nat := match args with @@ -203,7 +208,6 @@ Boolean equality for `LMonoTy`. def LMonoTy.BEq (x y : LMonoTy) : Bool := match x, y with | .ftvar i, .ftvar j => i == j - | .bitvec i, .bitvec j => i == j | .tcons i1 j1 r1, .tcons i2 j2 r2 => i1 == i2 && j1.length == j2.length && r1 == r2 && go j1 j2 | _, _ => false @@ -232,8 +236,6 @@ instance : DecidableEq LMonoTy := induction x generalizing y case ftvar => unfold LMonoTy.BEq at h <;> split at h <;> try simp_all - case bitvec => - unfold LMonoTy.BEq at h <;> split at h <;> try simp_all case tcons => rename_i name args r ih cases y <;> try simp_all [LMonoTy.BEq] @@ -251,8 +253,6 @@ instance : DecidableEq LMonoTy := isFalse (by induction x generalizing y case ftvar => cases y <;> try simp_all [LMonoTy.BEq] - case bitvec n => - cases y <;> try simp_all [LMonoTy.BEq] case tcons name args r ih => cases y <;> try simp_all [LMonoTy.BEq] rename_i name' args' r' @@ -282,7 +282,6 @@ it. def LMonoTy.freeVars (mty : LMonoTy) : List TyIdentifier := match mty with | .ftvar x => [x] - | .bitvec _ => [] | .tcons _ ltys _ => LMonoTys.freeVars ltys /-- @@ -449,18 +448,18 @@ scoped syntax "(" lmonoty ")" : lmonoty open Lean Elab Meta in partial def elabLMonoTy : Lean.Syntax → MetaM Expr | `(lmonoty| %$f:ident) => do - mkAppOptM ``LMonoTy.ftvar #[mkConst `Lambda.LTyRestrict, mkStrLit (toString f.getId)] + mkAppOptM ``LMonoTy.ftvar #[mkConst `Lambda.BoundTyRestrict, mkStrLit (toString f.getId)] | `(lmonoty| $ty1:lmonoty → $ty2:lmonoty) => do let ty1' ← elabLMonoTy ty1 let ty2' ← elabLMonoTy ty2 - let tys ← mkListLit (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) [ty1', ty2'] - mkAppM ``LMonoTy.tcons #[(mkStrLit "arrow"), tys, (mkConst `Lambda.LTyRestrict.nodata)] + let tys ← mkListLit (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) [ty1', ty2'] + mkAppM ``LMonoTy.tcons #[(mkStrLit "arrow"), tys, (.app (mkConst `Lambda.LTyRestrict.nodata) (mkConst `Lambda.BoundTyRestrict))] | `(lmonoty| int) => do - let argslist ← mkListLit (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) [] - mkAppM ``LMonoTy.tcons #[(mkStrLit "int"), argslist, (mkConst `Lambda.LTyRestrict.nodata)] + let argslist ← mkListLit (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) [] + mkAppM ``LMonoTy.tcons #[(mkStrLit "int"), argslist, (.app (mkConst `Lambda.LTyRestrict.nodata) (mkConst `Lambda.BoundTyRestrict))] | `(lmonoty| bool) => do - let argslist ← mkListLit (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) [] - mkAppM ``LMonoTy.tcons #[(mkStrLit "bool"), argslist, (mkConst `Lambda.LTyRestrict.nodata)] + let argslist ← mkListLit (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) [] + mkAppM ``LMonoTy.tcons #[(mkStrLit "bool"), argslist, (.app (mkConst `Lambda.LTyRestrict.nodata) (mkConst `Lambda.BoundTyRestrict))] | `(lmonoty| bv1) => mkAppM ``LMonoTy.bv1 #[] | `(lmonoty| bv8) => mkAppM ``LMonoTy.bv8 #[] | `(lmonoty| bv16) => mkAppM ``LMonoTy.bv16 #[] @@ -468,8 +467,8 @@ partial def elabLMonoTy : Lean.Syntax → MetaM Expr | `(lmonoty| bv64) => mkAppM ``LMonoTy.bv64 #[] | `(lmonoty| $i:ident $args:lmonoty*) => do let args' ← go args - let argslist ← mkListLit (.app (mkConst ``LMonoTy) (mkConst `Lambda.LTyRestrict)) args'.toList - mkAppM ``LMonoTy.tcons #[(mkStrLit (toString i.getId)), argslist, (mkConst `Lambda.LTyRestrict.nodata)] + let argslist ← mkListLit (.app (mkConst ``LMonoTy) (mkConst `Lambda.BoundTyRestrict)) args'.toList + mkAppM ``LMonoTy.tcons #[(mkStrLit (toString i.getId)), argslist, (.app (mkConst `Lambda.LTyRestrict.nodata) (mkConst `Lambda.BoundTyRestrict))] | `(lmonoty| ($ty:lmonoty)) => elabLMonoTy ty | _ => throwUnsupportedSyntax where go (args : TSyntaxArray `lmonoty) : MetaM (Array Expr) := do @@ -481,31 +480,35 @@ partial def elabLMonoTy : Lean.Syntax → MetaM Expr elab "mty[" ty:lmonoty "]" : term => elabLMonoTy ty -/-- info: LMonoTy.tcons LTyRestrict "list" [LMonoTy.tcons LTyRestrict "int" [] LTyRestrict.nodata] LTyRestrict.nodata : LMonoTy -/ +/-- info: LMonoTy.tcons BoundTyRestrict "list" [LMonoTy.tcons BoundTyRestrict "int" [] LTyRestrict.nodata] + LTyRestrict.nodata : LMonoTy -/ #guard_msgs in #check mty[list int] -/-- info: LMonoTy.tcons LTyRestrict "pair" - [LMonoTy.tcons LTyRestrict "int" [] LTyRestrict.nodata, LMonoTy.tcons LTyRestrict "bool" [] LTyRestrict.nodata] +/-- info: LMonoTy.tcons BoundTyRestrict "pair" + [LMonoTy.tcons BoundTyRestrict "int" [] LTyRestrict.nodata, + LMonoTy.tcons BoundTyRestrict "bool" [] LTyRestrict.nodata] LTyRestrict.nodata : LMonoTy -/ #guard_msgs in #check mty[pair int bool] /-- -info: LMonoTy.tcons LTyRestrict "arrow" - [LMonoTy.tcons LTyRestrict "Map" [LMonoTy.ftvar LTyRestrict "k", LMonoTy.ftvar LTyRestrict "v"] LTyRestrict.nodata, - LMonoTy.tcons LTyRestrict "arrow" [LMonoTy.ftvar LTyRestrict "k", LMonoTy.ftvar LTyRestrict "v"] LTyRestrict.nodata] +info: LMonoTy.tcons BoundTyRestrict "arrow" + [LMonoTy.tcons BoundTyRestrict "Map" [LMonoTy.ftvar BoundTyRestrict "k", LMonoTy.ftvar BoundTyRestrict "v"] + LTyRestrict.nodata, + LMonoTy.tcons BoundTyRestrict "arrow" [LMonoTy.ftvar BoundTyRestrict "k", LMonoTy.ftvar BoundTyRestrict "v"] + LTyRestrict.nodata] LTyRestrict.nodata : LMonoTy -/ #guard_msgs in #check mty[(Map %k %v) → %k → %v] /-- -info: LMonoTy.tcons LTyRestrict "arrow" - [LMonoTy.ftvar LTyRestrict "a", - LMonoTy.tcons LTyRestrict "arrow" - [LMonoTy.ftvar LTyRestrict "b", - LMonoTy.tcons LTyRestrict "arrow" [LMonoTy.ftvar LTyRestrict "c", LMonoTy.ftvar LTyRestrict "d"] +info: LMonoTy.tcons BoundTyRestrict "arrow" + [LMonoTy.ftvar BoundTyRestrict "a", + LMonoTy.tcons BoundTyRestrict "arrow" + [LMonoTy.ftvar BoundTyRestrict "b", + LMonoTy.tcons BoundTyRestrict "arrow" [LMonoTy.ftvar BoundTyRestrict "c", LMonoTy.ftvar BoundTyRestrict "d"] LTyRestrict.nodata] LTyRestrict.nodata] LTyRestrict.nodata : LMonoTy @@ -534,16 +537,17 @@ partial def elabLTy : Lean.Syntax → MetaM Expr elab "t[" lsch:lty "]" : term => elabLTy lsch -/-- info: forAll ["α"] (LMonoTy.tcons LTyRestrict "myType" [LMonoTy.ftvar LTyRestrict "α"] LTyRestrict.nodata) : LTy -/ +/-- info: forAll ["α"] (LMonoTy.tcons BoundTyRestrict "myType" [LMonoTy.ftvar BoundTyRestrict "α"] LTyRestrict.nodata) : LTy -/ #guard_msgs in #check t[∀α. myType %α] /-- info: forAll ["α"] - (LMonoTy.tcons LTyRestrict "arrow" - [LMonoTy.ftvar LTyRestrict "α", - LMonoTy.tcons LTyRestrict "arrow" - [LMonoTy.ftvar LTyRestrict "α", LMonoTy.tcons LTyRestrict "int" [] LTyRestrict.nodata] LTyRestrict.nodata] + (LMonoTy.tcons BoundTyRestrict "arrow" + [LMonoTy.ftvar BoundTyRestrict "α", + LMonoTy.tcons BoundTyRestrict "arrow" + [LMonoTy.ftvar BoundTyRestrict "α", LMonoTy.tcons BoundTyRestrict "int" [] LTyRestrict.nodata] + LTyRestrict.nodata] LTyRestrict.nodata) : LTy -/ #guard_msgs in diff --git a/Strata/DL/Lambda/LTyUnify.lean b/Strata/DL/Lambda/LTyUnify.lean index 768d2045..9d8a3783 100644 --- a/Strata/DL/Lambda/LTyUnify.lean +++ b/Strata/DL/Lambda/LTyUnify.lean @@ -160,7 +160,6 @@ def LMonoTy.subst (S : Subst) (mty : LMonoTy) : LMonoTy := match mty with | .ftvar x => match S.find? x with | some sty => sty | none => mty - | .bitvec _ => mty | .tcons name ltys r => .tcons name (LMonoTys.subst S ltys) r /-- @@ -249,10 +248,6 @@ theorem LMonoTy.subst_keys_not_in_substituted_type (h : SubstWF S) : have := @Maps.find?_of_not_mem_values _ _ i _ S simp_all exact ne_of_mem_of_not_mem hid this - case bitvec n => - simp_all [LMonoTy.subst] - unfold LMonoTy.freeVars - simp case tcons name args r h1 => simp_all simp [subst] @@ -292,8 +287,6 @@ theorem LMonoTy.freeVars_of_subst_subset (S : Subst) (mty : LMonoTy) : apply @Maps.find?_mem_values _ _ x sty _ S h_find · -- Case: S.find? x = none simp [freeVars] - case bitvec n => - simp [subst] case tcons name args r ih => simp [LMonoTy.subst, LMonoTy.freeVars] induction args diff --git a/Strata/Languages/Boogie/Identifiers.lean b/Strata/Languages/Boogie/Identifiers.lean index f2955c88..a2517fc9 100644 --- a/Strata/Languages/Boogie/Identifiers.lean +++ b/Strata/Languages/Boogie/Identifiers.lean @@ -142,7 +142,7 @@ info: (Lambda.LExpr.op (BoogieIdent.unres "old") none).app open Lambda.LTy.Syntax in /-- info: Lambda.LExpr.fvar (BoogieIdent.unres "x") (some - (Lambda.LMonoTy.tcons Lambda.LTyRestrict "bool" [] + (Lambda.LMonoTy.tcons Lambda.BoundTyRestrict "bool" [] Lambda.LTyRestrict.nodata)) : Lambda.LExpr Lambda.LMonoTy (Visibility × String) -/ #guard_msgs in #check eb[(x : bool)] From 59e6eb715a2c95fe456a9287acfc358146664178 Mon Sep 17 00:00:00 2001 From: Josh Cohen Date: Fri, 10 Oct 2025 11:15:55 -0400 Subject: [PATCH 18/19] Add missing type translation in output for abs, quant, ite --- Strata/DL/Bounded/BExpr.lean | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/Strata/DL/Bounded/BExpr.lean b/Strata/DL/Bounded/BExpr.lean index 89f76dd2..c515cc4c 100644 --- a/Strata/DL/Bounded/BExpr.lean +++ b/Strata/DL/Bounded/BExpr.lean @@ -313,13 +313,13 @@ def translateBounded [Coe String Identifier] [DecidableEq Identifier] (e: LExprT | .abs e (.arrow ty .bool) => let e' := translateBounded e [] pos; let all_assume := conditionalAdd pos (makeBoundAssumption ty ++ (increaseBVars assume)) e'.assume; - ⟨.abs (addAssumptions all_assume e'.translate) (removeTyBound (.arrow ty .bool)), addBoundedWfAssume all_assume e'.wfCond, conditionalReturn pos (List.map (quantifyAssumptions ty) e'.assume)⟩ + ⟨.abs (addAssumptions all_assume e'.translate) (removeTyBound (.arrow ty .bool)), addBoundedWfAssume all_assume e'.wfCond, conditionalReturn pos (List.map (quantifyAssumptions (removeTyBound ty)) e'.assume)⟩ | .abs e (.arrow ty1 ty2) => let all_assume := makeBoundAssumption ty1 ++ (increaseBVars assume); let e' := translateBounded e all_assume pos; let res := .abs e'.translate (removeTyBound (.arrow ty1 ty2)); -- Note: don't add assumptions to e'.wfCond, already included - ⟨res, addBoundedWf all_assume (boundExprIfType ty2 e'.translate) e'.wfCond, List.map (quantifyAssumptions ty1) e'.assume⟩ + ⟨res, addBoundedWf all_assume (boundExprIfType ty2 e'.translate) e'.wfCond, List.map (quantifyAssumptions (removeTyBound ty1)) e'.assume⟩ | .abs _ _ => ⟨.const "0" .int, [], []⟩ -- error case /- Quantifiers are simpler because they are boolean-valued. ∀ (x : nat). e adds an assumption x >= 0 -> ..., while ∃ (x: nat). e results in x >= 0 ∧ .. @@ -328,13 +328,13 @@ def translateBounded [Coe String Identifier] [DecidableEq Identifier] (e: LExprT let e' := translateBounded e [] pos; let tr' := translateClean tr; let all_assume := conditionalAdd pos (makeBoundAssumption ty ++ (increaseBVars assume)) e'.assume; - ⟨.quant .all (removeTyBound ty) tr' (addAssumptions all_assume e'.translate), addBoundedWfAssume all_assume e'.wfCond, conditionalReturn pos (List.map (quantifyAssumptions ty) e'.assume)⟩ + ⟨.quant .all (removeTyBound ty) tr' (addAssumptions all_assume e'.translate), addBoundedWfAssume all_assume e'.wfCond, conditionalReturn pos (List.map (quantifyAssumptions (removeTyBound ty)) e'.assume)⟩ | .quant .exist ty tr e => let newAssumption := makeBoundAssumption ty; let e' := translateBounded e [] pos; let tr' := translateClean tr; let all_assume := conditionalAdd pos (increaseBVars assume) e'.assume; - ⟨.quant .exist (removeTyBound ty) tr' (addAssumptions all_assume (addAndExprs newAssumption e'.translate)), addBoundedWfAssume (newAssumption ++ all_assume) e'.wfCond, conditionalReturn pos (List.map (quantifyAssumptions ty) e'.assume)⟩ + ⟨.quant .exist (removeTyBound ty) tr' (addAssumptions all_assume (addAndExprs newAssumption e'.translate)), addBoundedWfAssume (newAssumption ++ all_assume) e'.wfCond, conditionalReturn pos (List.map (quantifyAssumptions (removeTyBound ty)) e'.assume)⟩ /- For if-then-else, add assumptions to the condition, which is always bool-valued. For a bool-valued result, can add the conditions freely. For a bounded-valued term, produce a wf condition proving this. -/ @@ -350,7 +350,7 @@ def translateBounded [Coe String Identifier] [DecidableEq Identifier] (e: LExprT let t' := translateBounded t assume pos; let f' := translateBounded f assume pos; -- here can add inside if positive, never add outside - let res := .ite (addAssumptions (conditionalAdd pos assume c'.assume) c'.translate) t'.translate f'.translate ty; + let res := .ite (addAssumptions (conditionalAdd pos assume c'.assume) c'.translate) t'.translate f'.translate (removeTyBound ty); ⟨res, ListSet.union [boundExprIfType ty res, c'.wfCond, t'.wfCond, f'.wfCond], conditionalReturn pos (c'.assume ++ t'.assume ++ f'.assume)⟩ /- Equality is always bool-valued, so can add assumptions @@ -359,7 +359,7 @@ def translateBounded [Coe String Identifier] [DecidableEq Identifier] (e: LExprT | .eq e1 e2 ty => let e1' := translateBounded e1 [] false; let e2' := translateBounded e2 [] false; - ⟨addAssumptions (conditionalAdd pos assume (e1'.assume ++ e2'.assume)) (.eq e1'.translate e2'.translate ty), ListSet.union [e1'.wfCond, e2'.wfCond], conditionalReturn pos (e1'.assume ++ e2'.assume)⟩ + ⟨addAssumptions (conditionalAdd pos assume (e1'.assume ++ e2'.assume)) (.eq e1'.translate e2'.translate (removeTyBound ty)), ListSet.union [e1'.wfCond, e2'.wfCond], conditionalReturn pos (e1'.assume ++ e2'.assume)⟩ | .mdata m e => let e' := translateBounded e assume pos; ⟨.mdata m e'.translate, e'.wfCond, e'.assume⟩ From 5c5fae6512cc09bb3b126cba21f8d9fda99998e7 Mon Sep 17 00:00:00 2001 From: Josh Cohen Date: Fri, 10 Oct 2025 11:20:08 -0400 Subject: [PATCH 19/19] Add copyright header to BExprTest --- StrataTest/DL/Bounded/BExprTest.lean | 8 +++++++- 1 file changed, 7 insertions(+), 1 deletion(-) diff --git a/StrataTest/DL/Bounded/BExprTest.lean b/StrataTest/DL/Bounded/BExprTest.lean index f6f77afa..f887e84b 100644 --- a/StrataTest/DL/Bounded/BExprTest.lean +++ b/StrataTest/DL/Bounded/BExprTest.lean @@ -1,8 +1,14 @@ +/- + Copyright Strata Contributors + + SPDX-License-Identifier: Apache-2.0 OR MIT +-/ + import Strata.DL.Bounded.BExpr import Strata.DL.Lambda.LExprT import Strata.DL.Lambda.LTy --- Tests +-- Tests for Bounded Expressions namespace Test open Std (ToFormat Format format)