From 828b7d09860d5d0755096c6c5af8c3b9a5954b76 Mon Sep 17 00:00:00 2001 From: utkn Date: Fri, 15 Nov 2024 11:22:26 +0100 Subject: [PATCH 01/39] First trial --- Lampe/Lampe/Ast.lean | 1 + Lampe/Lampe/Semantics.lean | 15 +++++++++++++++ 2 files changed, 16 insertions(+) diff --git a/Lampe/Lampe/Ast.lean b/Lampe/Lampe/Ast.lean index 10ddad5..dd3275c 100644 --- a/Lampe/Lampe/Ast.lean +++ b/Lampe/Lampe/Ast.lean @@ -11,6 +11,7 @@ abbrev Ident := String inductive FunctionIdent : Type where | builtin : Builtin → FunctionIdent | decl : Ident → FunctionIdent +| ref : Ref → FunctionIdent inductive Member : Tp → List Tp → Type where | head : Member tp (tp :: tps) diff --git a/Lampe/Lampe/Semantics.lean b/Lampe/Lampe/Semantics.lean index 8c7cec0..3a50cdd 100644 --- a/Lampe/Lampe/Semantics.lean +++ b/Lampe/Lampe/Semantics.lean @@ -17,6 +17,8 @@ def Env.ofModule (m : Module): Env := fun i => (m.decls.find? fun d => d.name == @[reducible] def Env.extend (Γ₁ : Env) (Γ₂ : Env) : Env := fun i => Γ₁ i <|> Γ₂ i +def Ident.ofLambda (lambdaRef : Ref) : Ident := lambdaRef.val.toSubscriptString + inductive Omni : Env → State P → Expr (Tp.denote P) tp → (Option (State P × tp.denote P) → Prop) → Prop where | litField : Q (some (st, n)) → Omni Γ st (.lit .field n) Q | litFalse : Q (some (st, false)) → Omni Γ st (.lit .bool 0) Q @@ -45,6 +47,13 @@ inductive Omni : Env → State P → Expr (Tp.denote P) tp → (Option (State P (htco : fn.outTp (hkc ▸ generics) = res) → Omni Γ st (htco ▸ fn.body _ (hkc ▸ generics) (htci ▸ args)) Q → Omni Γ st (@Expr.call _ tyKinds generics argTypes res (.decl fname) args) Q +| callLambda : + Γ (Ident.ofLambda lambdaRef) = some fn → + (hg : fn.generics = []) → + (hi : fn.inTps (hg ▸ h![]) = argTps) → + (ho : fn.outTp (hg ▸ h![]) = outTp) → + Omni Γ st (ho ▸ fn.body _ (hg ▸ h![]) (hi ▸ args)) Q → + Omni Γ st (Expr.call h![] argTps outTp (.ref lambdaRef) args) Q | loopDone : lo ≥ hi → Omni Γ st (.loop lo hi body) Q @@ -125,5 +134,11 @@ theorem Omni.frame {p Γ tp st₁ st₂} {e : Expr (Tp.denote p) tp} {Q}: constructor apply ih assumption + | callLambda ih => + intros + constructor + apply ih + all_goals (try assumption) + tauto end Lampe From d8daaa12e1a2403c29e97c6a155a1aa8689367e4 Mon Sep 17 00:00:00 2001 From: utkn Date: Fri, 15 Nov 2024 17:45:03 +0100 Subject: [PATCH 02/39] `Omni.newLambda`, `Omni.callLambda` omni rule `Expr.newLambda` expression Lambda function constructor --- Lampe/Lampe/Ast.lean | 17 +++++++++++------ Lampe/Lampe/Lambda.lean | 16 ++++++++++++++++ Lampe/Lampe/Semantics.lean | 24 +++++++++++++++++------- 3 files changed, 44 insertions(+), 13 deletions(-) create mode 100644 Lampe/Lampe/Lambda.lean diff --git a/Lampe/Lampe/Ast.lean b/Lampe/Lampe/Ast.lean index dd3275c..b730356 100644 --- a/Lampe/Lampe/Ast.lean +++ b/Lampe/Lampe/Ast.lean @@ -25,27 +25,32 @@ inductive Expr (rep : Tp → Type): Tp → Type where | ite : rep .bool → Expr rep a → Expr rep a → Expr rep a | skip : Expr rep .unit | loop : rep (.u s) → rep (.u s) → (rep (.u s) → Expr rep r) → Expr rep .unit +| newLambda : (argTps : List Tp) → (outTp : Tp) → (HList rep argTps → Expr rep outTp) → Expr rep (.ref .unit) structure Function : Type _ where generics : List Kind inTps : HList Kind.denote generics → List Tp outTp : HList Kind.denote generics → Tp - body : ∀ rep, (gs : HList Kind.denote generics) → HList rep (inTps gs) → Expr rep (outTp gs) + rep : Tp → Type + body : (gs : HList Kind.denote generics) → + HList rep (inTps gs) → + Expr rep (outTp gs) /-- Polymorphic identity -/ -example : Function := { +example {p : Prime}: Function := { generics := [.type] + rep := fun tp => tp.denote p inTps := fun h![x] => [x] outTp := fun h![x] => x - body := fun _ h![_] h![x] => .var x + body := fun h![_] h![x] => .var x } structure FunctionDecl where -name : Ident -fn : Function + name : Ident + fn : Function structure Module where -decls : List FunctionDecl + decls : List FunctionDecl structure Struct where name : String diff --git a/Lampe/Lampe/Lambda.lean b/Lampe/Lampe/Lambda.lean new file mode 100644 index 0000000..6774636 --- /dev/null +++ b/Lampe/Lampe/Lambda.lean @@ -0,0 +1,16 @@ +import Mathlib +import Lampe.Ast + +namespace Lampe + +def newLambda (argTps : List Tp) {outTp : Tp} (body : HList rep argTps → Expr rep outTp) : Function := { + generics := [] + inTps := fun _ => argTps + outTp := fun _ => outTp + rep := rep + body := fun _ => body +} + +def Ident.ofLambdaRef (lambdaRef : Ref) : Ident := toString lambdaRef.val + +end Lampe diff --git a/Lampe/Lampe/Semantics.lean b/Lampe/Lampe/Semantics.lean index 3a50cdd..5632112 100644 --- a/Lampe/Lampe/Semantics.lean +++ b/Lampe/Lampe/Semantics.lean @@ -4,7 +4,7 @@ import Lampe.Data.Field import Lampe.Tactic.IntroCases import Mathlib import Lampe.State - +import Lampe.Lambda namespace Lampe @@ -12,13 +12,11 @@ variable (P : Prime) def Env := Ident → Option Function -def Env.ofModule (m : Module): Env := fun i => (m.decls.find? fun d => d.name == i).map (·.fn) +def Env.ofModule (m : Module) : Env := fun i => (m.decls.find? fun d => d.name == i).map (·.fn) @[reducible] def Env.extend (Γ₁ : Env) (Γ₂ : Env) : Env := fun i => Γ₁ i <|> Γ₂ i -def Ident.ofLambda (lambdaRef : Ref) : Ident := lambdaRef.val.toSubscriptString - inductive Omni : Env → State P → Expr (Tp.denote P) tp → (Option (State P × tp.denote P) → Prop) → Prop where | litField : Q (some (st, n)) → Omni Γ st (.lit .field n) Q | litFalse : Q (some (st, false)) → Omni Γ st (.lit .bool 0) Q @@ -45,15 +43,21 @@ inductive Omni : Env → State P → Expr (Tp.denote P) tp → (Option (State P (hkc : fn.generics = tyKinds) → (htci : fn.inTps (hkc ▸ generics) = argTypes) → (htco : fn.outTp (hkc ▸ generics) = res) → - Omni Γ st (htco ▸ fn.body _ (hkc ▸ generics) (htci ▸ args)) Q → + (hrep : fn.rep = Tp.denote P) → + Omni Γ st (htco ▸ hrep ▸ fn.body (hkc ▸ generics) (htci ▸ hrep ▸ args)) Q → Omni Γ st (@Expr.call _ tyKinds generics argTypes res (.decl fname) args) Q | callLambda : - Γ (Ident.ofLambda lambdaRef) = some fn → + Γ (Ident.ofLambdaRef lambdaRef) = some fn → (hg : fn.generics = []) → (hi : fn.inTps (hg ▸ h![]) = argTps) → (ho : fn.outTp (hg ▸ h![]) = outTp) → - Omni Γ st (ho ▸ fn.body _ (hg ▸ h![]) (hi ▸ args)) Q → + (hrep : fn.rep = Tp.denote P) → + Omni Γ st (ho ▸ hrep ▸ fn.body (hg ▸ h![]) (hi ▸ hrep ▸ args)) Q → Omni Γ st (Expr.call h![] argTps outTp (.ref lambdaRef) args) Q +| newLambda {Q} : + Q (some (st, lambdaRef)) → + Γ (Ident.ofLambdaRef lambdaRef) = some (@newLambda (Tp.denote P ·) argTps outTp body) → + Omni Γ st (Expr.newLambda argTps outTp body) Q | loopDone : lo ≥ hi → Omni Γ st (.loop lo hi body) Q @@ -134,6 +138,12 @@ theorem Omni.frame {p Γ tp st₁ st₂} {e : Expr (Tp.denote p) tp} {Q}: constructor apply ih assumption + | newLambda => + intros + constructor + all_goals (try tauto) + constructor + tauto | callLambda ih => intros constructor From dd05821b644850499ffd4179e124c054abdfbcca Mon Sep 17 00:00:00 2001 From: utkn Date: Fri, 15 Nov 2024 19:16:47 +0100 Subject: [PATCH 03/39] lambda syntax first iteration --- Lampe/Lampe.lean | 5 +++++ Lampe/Lampe/Ast.lean | 15 ++++++++------- Lampe/Lampe/Lambda.lean | 10 ++++++---- Lampe/Lampe/Semantics.lean | 10 ++++------ Lampe/Lampe/Syntax.lean | 11 +++++++++++ 5 files changed, 34 insertions(+), 17 deletions(-) diff --git a/Lampe/Lampe.lean b/Lampe/Lampe.lean index a9cecd0..6924912 100644 --- a/Lampe/Lampe.lean +++ b/Lampe/Lampe.lean @@ -90,3 +90,8 @@ example {p Γ x y}: STHoare p Γ ⟦⟧ (simple_if_else.fn.body (Tp.denote p) h! sl contradiction . simp_all + +nr_def simple_lambda<>(x : Field) -> Field { + let foo = |a|: Field -> Field { a }; + ^foo(x) : Field; +} diff --git a/Lampe/Lampe/Ast.lean b/Lampe/Lampe/Ast.lean index b730356..ee5929d 100644 --- a/Lampe/Lampe/Ast.lean +++ b/Lampe/Lampe/Ast.lean @@ -11,7 +11,7 @@ abbrev Ident := String inductive FunctionIdent : Type where | builtin : Builtin → FunctionIdent | decl : Ident → FunctionIdent -| ref : Ref → FunctionIdent +| ref : Tp.denote p TpLambdaRef → FunctionIdent inductive Member : Tp → List Tp → Type where | head : Member tp (tp :: tps) @@ -25,24 +25,25 @@ inductive Expr (rep : Tp → Type): Tp → Type where | ite : rep .bool → Expr rep a → Expr rep a → Expr rep a | skip : Expr rep .unit | loop : rep (.u s) → rep (.u s) → (rep (.u s) → Expr rep r) → Expr rep .unit -| newLambda : (argTps : List Tp) → (outTp : Tp) → (HList rep argTps → Expr rep outTp) → Expr rep (.ref .unit) +| newLambda: (argTps : List Tp) → + (outTp : Tp) → + (HList rep argTps → Expr rep outTp) → + Expr rep (.ref .unit) structure Function : Type _ where generics : List Kind inTps : HList Kind.denote generics → List Tp outTp : HList Kind.denote generics → Tp - rep : Tp → Type - body : (gs : HList Kind.denote generics) → + body : (rep : Tp → Type) → (gs : HList Kind.denote generics) → HList rep (inTps gs) → Expr rep (outTp gs) /-- Polymorphic identity -/ -example {p : Prime}: Function := { +example : Function := { generics := [.type] - rep := fun tp => tp.denote p inTps := fun h![x] => [x] outTp := fun h![x] => x - body := fun h![_] h![x] => .var x + body := fun _ h![_] h![x] => .var x } structure FunctionDecl where diff --git a/Lampe/Lampe/Lambda.lean b/Lampe/Lampe/Lambda.lean index 6774636..c4bd2c0 100644 --- a/Lampe/Lampe/Lambda.lean +++ b/Lampe/Lampe/Lambda.lean @@ -3,14 +3,16 @@ import Lampe.Ast namespace Lampe -def newLambda (argTps : List Tp) {outTp : Tp} (body : HList rep argTps → Expr rep outTp) : Function := { +abbrev TpLambdaRef := Tp.ref .unit + +def newLambda (argTps : List Tp) (outTp : Tp) +(body : (rep : Tp → Type) → HList Kind.denote [] → HList rep argTps → Expr rep outTp) : Function := { generics := [] inTps := fun _ => argTps outTp := fun _ => outTp - rep := rep - body := fun _ => body + body := body } -def Ident.ofLambdaRef (lambdaRef : Ref) : Ident := toString lambdaRef.val +def Ident.ofLambdaRef (lambdaRef : Tp.denote p TpLambdaRef) : Ident := toString lambdaRef.val end Lampe diff --git a/Lampe/Lampe/Semantics.lean b/Lampe/Lampe/Semantics.lean index 5632112..e52b6ea 100644 --- a/Lampe/Lampe/Semantics.lean +++ b/Lampe/Lampe/Semantics.lean @@ -43,21 +43,19 @@ inductive Omni : Env → State P → Expr (Tp.denote P) tp → (Option (State P (hkc : fn.generics = tyKinds) → (htci : fn.inTps (hkc ▸ generics) = argTypes) → (htco : fn.outTp (hkc ▸ generics) = res) → - (hrep : fn.rep = Tp.denote P) → - Omni Γ st (htco ▸ hrep ▸ fn.body (hkc ▸ generics) (htci ▸ hrep ▸ args)) Q → + Omni Γ st (htco ▸ fn.body _ (hkc ▸ generics) (htci ▸ args)) Q → Omni Γ st (@Expr.call _ tyKinds generics argTypes res (.decl fname) args) Q | callLambda : Γ (Ident.ofLambdaRef lambdaRef) = some fn → (hg : fn.generics = []) → (hi : fn.inTps (hg ▸ h![]) = argTps) → (ho : fn.outTp (hg ▸ h![]) = outTp) → - (hrep : fn.rep = Tp.denote P) → - Omni Γ st (ho ▸ hrep ▸ fn.body (hg ▸ h![]) (hi ▸ hrep ▸ args)) Q → + Omni Γ st (ho ▸ fn.body _ (hg ▸ h![]) (hi ▸ args)) Q → Omni Γ st (Expr.call h![] argTps outTp (.ref lambdaRef) args) Q | newLambda {Q} : Q (some (st, lambdaRef)) → - Γ (Ident.ofLambdaRef lambdaRef) = some (@newLambda (Tp.denote P ·) argTps outTp body) → - Omni Γ st (Expr.newLambda argTps outTp body) Q + Γ (Ident.ofLambdaRef lambdaRef) = some (newLambda argTps outTp body) → + Omni Γ st (Expr.newLambda argTps outTp (body _ _)) Q | loopDone : lo ≥ hi → Omni Γ st (.loop lo hi body) Q diff --git a/Lampe/Lampe/Syntax.lean b/Lampe/Lampe/Syntax.lean index 1c34918..1ea2414 100644 --- a/Lampe/Lampe/Syntax.lean +++ b/Lampe/Lampe/Syntax.lean @@ -98,6 +98,8 @@ syntax "if" nr_expr nr_expr ("else" nr_expr)? : nr_expr syntax "for" ident "in" nr_expr ".." nr_expr nr_expr : nr_expr syntax "(" nr_expr ")" : nr_expr syntax "*(" nr_expr ")" : nr_expr +syntax "|" ident,* "|:" nr_type,* "->" nr_type nr_expr : nr_expr +syntax "^" nr_ident "(" nr_expr,* ")" ":" nr_type : nr_expr def Expr.letMutIn (definition : Expr rep tp) (body : rep tp.ref → Expr rep tp'): Expr rep tp' := let refDef := Expr.letIn definition fun v => Expr.call h![] _ (tp.ref) (.builtin .ref) h![v] @@ -196,6 +198,15 @@ partial def mkExpr [MonadSyntax m] (e : TSyntax `nr_expr) (vname : Option Lean.I mkExpr cond none fun cond => do let mainBody ← mkExpr mainBody none fun x => ``(Lampe.Expr.var $x) wrapSimple (←`(Lampe.Expr.ite $cond $mainBody (Lampe.Expr.skip))) vname k +| `(nr_expr| |$args,*|: $argTps,* -> $outTp $lambdaBody) => do + let outTp ← mkNrType outTp + let argTps ← mkListLit (← argTps.getElems.toList.mapM mkNrType) + let args ← mkHListLit args.getElems.toList + let body ← mkExpr lambdaBody none fun x => ``(Lampe.Expr.var $x) + wrapSimple (←`(Lampe.Expr.newLambda $argTps $outTp (fun $args => $body))) vname k +| `(nr_expr| ^ $i:ident ($args,*): $tp) => do + mkArgs args.getElems.toList fun argVals => do + wrapSimple (←`(Lampe.Expr.call h![] _ $(←mkNrType tp) (.ref $i) $(←mkHListLit argVals))) vname k | _ => throwUnsupportedSyntax end From ed4204edb9b541eb491e4c17e084fc47eb455925 Mon Sep 17 00:00:00 2001 From: utkn Date: Fri, 15 Nov 2024 21:32:45 +0100 Subject: [PATCH 04/39] Lambda parsing complete. Intro rules (first iteration) --- Lampe/Lampe.lean | 5 +++++ Lampe/Lampe/Ast.lean | 29 +++++++++++++++++-------- Lampe/Lampe/Hoare/SepTotal.lean | 27 +++++++++++++++++++++++ Lampe/Lampe/Lambda.lean | 18 --------------- Lampe/Lampe/Semantics.lean | 5 ++--- Lampe/Lampe/Syntax.lean | 4 ++-- Lampe/Lampe/Tactic/SeparationLogic.lean | 6 +++++ 7 files changed, 62 insertions(+), 32 deletions(-) delete mode 100644 Lampe/Lampe/Lambda.lean diff --git a/Lampe/Lampe.lean b/Lampe/Lampe.lean index 6924912..8058bb3 100644 --- a/Lampe/Lampe.lean +++ b/Lampe/Lampe.lean @@ -95,3 +95,8 @@ nr_def simple_lambda<>(x : Field) -> Field { let foo = |a|: Field -> Field { a }; ^foo(x) : Field; } + +example {p Γ x} : STHoare p Γ ⟦⟧ (simple_lambda.fn.body (Tp.denote p) h![] h![x]) fun v => v = x := by + simp only [simple_lambda] + steps + all_goals sorry diff --git a/Lampe/Lampe/Ast.lean b/Lampe/Lampe/Ast.lean index ee5929d..aafefac 100644 --- a/Lampe/Lampe/Ast.lean +++ b/Lampe/Lampe/Ast.lean @@ -8,27 +8,30 @@ namespace Lampe abbrev Ident := String -inductive FunctionIdent : Type where +/-- A reference to a lambda is represented as a reference to a unit type -/ +abbrev Tp.lambdaRef := Tp.ref .unit + +/-- Converts a lambda reference to a function identity -/ +def Ident.ofLambdaRef (lambdaRef : Ref) : Ident := toString lambdaRef.val + +inductive FunctionIdent {rep : Tp → Type} : Type where | builtin : Builtin → FunctionIdent -| decl : Ident → FunctionIdent -| ref : Tp.denote p TpLambdaRef → FunctionIdent +| decl : Ident → FunctionIdent -- a function declared at the module level +| lambda : rep .lambdaRef → FunctionIdent inductive Member : Tp → List Tp → Type where | head : Member tp (tp :: tps) | tail : Member tp tps → Member tp (tp' :: tps) -inductive Expr (rep : Tp → Type): Tp → Type where +inductive Expr (rep : Tp → Type) : Tp → Type where | lit : (tp : Tp) → Nat → Expr rep tp | var : rep tp → Expr rep tp | letIn : Expr rep t₁ → (rep t₁ → Expr rep t₂) → Expr rep t₂ -| call : HList Kind.denote tyKinds → (argTypes : List Tp) → (res : Tp) → FunctionIdent → HList rep argTypes → Expr rep res +| call : HList Kind.denote tyKinds → (argTypes : List Tp) → (res : Tp) → @FunctionIdent rep → HList rep argTypes → Expr rep res | ite : rep .bool → Expr rep a → Expr rep a → Expr rep a | skip : Expr rep .unit | loop : rep (.u s) → rep (.u s) → (rep (.u s) → Expr rep r) → Expr rep .unit -| newLambda: (argTps : List Tp) → - (outTp : Tp) → - (HList rep argTps → Expr rep outTp) → - Expr rep (.ref .unit) +| lambda : (argTps : List Tp) → (outTp : Tp) → (HList rep argTps → Expr rep outTp) → Expr rep .lambdaRef structure Function : Type _ where generics : List Kind @@ -46,6 +49,14 @@ example : Function := { body := fun _ h![_] h![x] => .var x } +def newLambda (argTps : List Tp) (outTp : Tp) +(body : (rep : Tp → Type) → HList Kind.denote [] → HList rep argTps → Expr rep outTp) : Function := { + generics := [] + inTps := fun _ => argTps + outTp := fun _ => outTp + body := body +} + structure FunctionDecl where name : Ident fn : Function diff --git a/Lampe/Lampe/Hoare/SepTotal.lean b/Lampe/Lampe/Hoare/SepTotal.lean index fedea4e..de3ce14 100644 --- a/Lampe/Lampe/Hoare/SepTotal.lean +++ b/Lampe/Lampe/Hoare/SepTotal.lean @@ -350,4 +350,31 @@ theorem skip_intro : . apply SLP.ent_star_top tauto +theorem lambda_intro : + Γ (Ident.ofLambdaRef lambdaRef) = some (newLambda argTps outTp body) → + STHoare p Γ ⟦⟧ (.lambda argTps outTp (body _ h![])) fun v => v = lambdaRef := by + unfold STHoare + intros + unfold THoare + intros + constructor + all_goals (try tauto) + simp only + apply SLP.ent_star_top + assumption + +theorem lambdaCall_intro : + Γ (Ident.ofLambdaRef lambdaRef) = some fn → + (hg : fn.generics = []) → + (hi : fn.inTps (hg ▸ h![]) = argTps) → + (ho : fn.outTp (hg ▸ h![]) = outTp) → + STHoare p Γ ⟦⟧ (ho ▸ fn.body _ (hg ▸ h![]) (hi ▸ args)) Q → + STHoare p Γ ⟦⟧ (.call h![] argTps outTp (.lambda lambdaRef) args) Q := by + unfold STHoare + intros + unfold THoare + intros + constructor + <;> tauto + end Lampe.STHoare diff --git a/Lampe/Lampe/Lambda.lean b/Lampe/Lampe/Lambda.lean deleted file mode 100644 index c4bd2c0..0000000 --- a/Lampe/Lampe/Lambda.lean +++ /dev/null @@ -1,18 +0,0 @@ -import Mathlib -import Lampe.Ast - -namespace Lampe - -abbrev TpLambdaRef := Tp.ref .unit - -def newLambda (argTps : List Tp) (outTp : Tp) -(body : (rep : Tp → Type) → HList Kind.denote [] → HList rep argTps → Expr rep outTp) : Function := { - generics := [] - inTps := fun _ => argTps - outTp := fun _ => outTp - body := body -} - -def Ident.ofLambdaRef (lambdaRef : Tp.denote p TpLambdaRef) : Ident := toString lambdaRef.val - -end Lampe diff --git a/Lampe/Lampe/Semantics.lean b/Lampe/Lampe/Semantics.lean index e52b6ea..777fe8e 100644 --- a/Lampe/Lampe/Semantics.lean +++ b/Lampe/Lampe/Semantics.lean @@ -4,7 +4,6 @@ import Lampe.Data.Field import Lampe.Tactic.IntroCases import Mathlib import Lampe.State -import Lampe.Lambda namespace Lampe @@ -51,11 +50,11 @@ inductive Omni : Env → State P → Expr (Tp.denote P) tp → (Option (State P (hi : fn.inTps (hg ▸ h![]) = argTps) → (ho : fn.outTp (hg ▸ h![]) = outTp) → Omni Γ st (ho ▸ fn.body _ (hg ▸ h![]) (hi ▸ args)) Q → - Omni Γ st (Expr.call h![] argTps outTp (.ref lambdaRef) args) Q + Omni Γ st (Expr.call h![] argTps outTp (.lambda lambdaRef) args) Q | newLambda {Q} : Q (some (st, lambdaRef)) → Γ (Ident.ofLambdaRef lambdaRef) = some (newLambda argTps outTp body) → - Omni Γ st (Expr.newLambda argTps outTp (body _ _)) Q + Omni Γ st (Expr.lambda argTps outTp (body _ _)) Q | loopDone : lo ≥ hi → Omni Γ st (.loop lo hi body) Q diff --git a/Lampe/Lampe/Syntax.lean b/Lampe/Lampe/Syntax.lean index 1ea2414..d344258 100644 --- a/Lampe/Lampe/Syntax.lean +++ b/Lampe/Lampe/Syntax.lean @@ -203,10 +203,10 @@ partial def mkExpr [MonadSyntax m] (e : TSyntax `nr_expr) (vname : Option Lean.I let argTps ← mkListLit (← argTps.getElems.toList.mapM mkNrType) let args ← mkHListLit args.getElems.toList let body ← mkExpr lambdaBody none fun x => ``(Lampe.Expr.var $x) - wrapSimple (←`(Lampe.Expr.newLambda $argTps $outTp (fun $args => $body))) vname k + wrapSimple (←`(Lampe.Expr.lambda $argTps $outTp (fun $args => $body))) vname k | `(nr_expr| ^ $i:ident ($args,*): $tp) => do mkArgs args.getElems.toList fun argVals => do - wrapSimple (←`(Lampe.Expr.call h![] _ $(←mkNrType tp) (.ref $i) $(←mkHListLit argVals))) vname k + wrapSimple (←`(Lampe.Expr.call h![] _ $(←mkNrType tp) (.lambda $i) $(←mkHListLit argVals))) vname k | _ => throwUnsupportedSyntax end diff --git a/Lampe/Lampe/Tactic/SeparationLogic.lean b/Lampe/Lampe/Tactic/SeparationLogic.lean index 996ebfe..19883ea 100644 --- a/Lampe/Lampe/Tactic/SeparationLogic.lean +++ b/Lampe/Lampe/Tactic/SeparationLogic.lean @@ -452,6 +452,8 @@ macro "stephelper1" : tactic => `(tactic|( | apply fresh_intro | apply assert_intro | apply skip_intro + | apply lambdaCall_intro + | apply lambda_intro -- memory builtins | apply var_intro | apply ref_intro @@ -503,6 +505,8 @@ macro "stephelper2" : tactic => `(tactic|( | apply consequence_frame_left fresh_intro | apply consequence_frame_left Lampe.STHoare.litU_intro | apply consequence_frame_left assert_intro + -- | apply consequence_frame_left lambdaCall_intro + -- | apply consequence_frame_left lambda_intro -- | apply consequence_frame_left skip_intro -- memory builtins | apply consequence_frame_left var_intro @@ -557,6 +561,8 @@ macro "stephelper3" : tactic => `(tactic|( | apply ramified_frame_top Lampe.STHoare.litU_intro | apply ramified_frame_top assert_intro | apply ramified_frame_top skip_intro + | apply ramified_frame_top lambdaCall_intro + | apply ramified_frame_top lambda_intro -- memory builtins | apply ramified_frame_top var_intro | apply ramified_frame_top ref_intro From 6c41c454117bc507724784144729bf4032cd77e6 Mon Sep 17 00:00:00 2001 From: utkn Date: Sat, 16 Nov 2024 21:56:14 +0100 Subject: [PATCH 05/39] minor updates --- Lampe/Lampe/Hoare/Builtins.lean | 45 +++++++++++++++++++++ Lampe/Lampe/Hoare/SepTotal.lean | 71 ++++----------------------------- Lampe/Lampe/State.lean | 4 +- 3 files changed, 54 insertions(+), 66 deletions(-) diff --git a/Lampe/Lampe/Hoare/Builtins.lean b/Lampe/Lampe/Hoare/Builtins.lean index 4e391a6..68c34f1 100644 --- a/Lampe/Lampe/Hoare/Builtins.lean +++ b/Lampe/Lampe/Hoare/Builtins.lean @@ -2,6 +2,51 @@ import Lampe.Hoare.SepTotal namespace Lampe.STHoare +/-- +Introduction rule for pure builtins. +-/ +theorem pureBuiltin_intro {A : Type} {a : A} {sgn desc args} : + STHoare p Γ + ⟦⟧ + (.call h![] (sgn a).fst (sgn a).snd (.builtin (Builtin.newGenericPureBuiltin sgn desc)) args) + (fun v => ∃h, (v = (desc a args).snd h)) := by + unfold STHoare + intro H + unfold THoare + intros st p + constructor + cases em (desc a args).fst + . apply Builtin.genericPureOmni.ok + . simp_all only [SLP.true_star, exists_const] + apply SLP.ent_star_top + assumption + . tauto + . apply Builtin.genericPureOmni.err + . tauto + . simp + +lemma pureBuiltin_intro_consequence + {A : Type} {a : A} {sgn desc args} {Q : Tp.denote p outTp → Prop} + (h1 : argTps = (sgn a).fst) + (h2 : outTp = (sgn a).snd) + (hp : (h: (desc a (h1 ▸ args)).fst) → Q (h2 ▸ (desc a (h1 ▸ args)).snd h)) + : STHoare p Γ ⟦⟧ + (.call h![] argTps outTp (.builtin (Builtin.newGenericPureBuiltin sgn desc)) args) + fun v => Q v := by + subst_vars + dsimp only at * + apply ramified_frame_top + apply pureBuiltin_intro + simp only [SLP.true_star] + apply SLP.forall_right + intro + apply SLP.wand_intro + simp only [SLP.true_star] + apply SLP.pure_left' + rintro ⟨_, _⟩ + simp_all [SLP.entails_top] + + -- Arithmetics theorem uAdd_intro : STHoarePureBuiltin p Γ Builtin.uAdd (by tauto) h![a, b] := by diff --git a/Lampe/Lampe/Hoare/SepTotal.lean b/Lampe/Lampe/Hoare/SepTotal.lean index de3ce14..562a085 100644 --- a/Lampe/Lampe/Hoare/SepTotal.lean +++ b/Lampe/Lampe/Hoare/SepTotal.lean @@ -130,50 +130,6 @@ lemma Finmap.union_singleton [DecidableEq α] {β : α → Type u} {r : α} {v v · simp_all [Finmap.lookup_eq_none, eq_comm] simp_all [eq_comm] -/-- -Introduction rule for pure builtins. --/ -theorem pureBuiltin_intro {A : Type} {a : A} {sgn desc args} : - STHoare p Γ - ⟦⟧ - (.call h![] (sgn a).fst (sgn a).snd (.builtin (Builtin.newGenericPureBuiltin sgn desc)) args) - (fun v => ∃h, (v = (desc a args).snd h)) := by - unfold STHoare - intro H - unfold THoare - intros st p - constructor - cases em (desc a args).fst - . apply Builtin.genericPureOmni.ok - . simp_all only [SLP.true_star, exists_const] - apply SLP.ent_star_top - assumption - . tauto - . apply Builtin.genericPureOmni.err - . tauto - . simp - -lemma pureBuiltin_intro_consequence - {A : Type} {a : A} {sgn desc args} {Q : Tp.denote p outTp → Prop} - (h1 : argTps = (sgn a).fst) - (h2 : outTp = (sgn a).snd) - (hp : (h: (desc a (h1 ▸ args)).fst) → Q (h2 ▸ (desc a (h1 ▸ args)).snd h)) - : STHoare p Γ ⟦⟧ - (.call h![] argTps outTp (.builtin (Builtin.newGenericPureBuiltin sgn desc)) args) - fun v => Q v := by - subst_vars - dsimp only at * - apply ramified_frame_top - apply pureBuiltin_intro - simp only [SLP.true_star] - apply SLP.forall_right - intro - apply SLP.wand_intro - simp only [SLP.true_star] - apply SLP.pure_left' - rintro ⟨_, _⟩ - simp_all [SLP.entails_top] - theorem fresh_intro : STHoare p Γ ⟦⟧ @@ -350,31 +306,18 @@ theorem skip_intro : . apply SLP.ent_star_top tauto -theorem lambda_intro : - Γ (Ident.ofLambdaRef lambdaRef) = some (newLambda argTps outTp body) → - STHoare p Γ ⟦⟧ (.lambda argTps outTp (body _ h![])) fun v => v = lambdaRef := by - unfold STHoare - intros - unfold THoare - intros - constructor - all_goals (try tauto) - simp only - apply SLP.ent_star_top - assumption +theorem lambda_intro {ls : LambdaStore} : + ls lambdaRef = some (newLambda argTps outTp body) → + STHoare p Γ ⟦⟧ (.lambda argTps outTp (body _ h![])) (fun r => r = lambdaRef) := by + sorry -theorem lambdaCall_intro : - Γ (Ident.ofLambdaRef lambdaRef) = some fn → +theorem lambdaCall_intro {ls : LambdaStore} : + ls lambdaRef = some fn → (hg : fn.generics = []) → (hi : fn.inTps (hg ▸ h![]) = argTps) → (ho : fn.outTp (hg ▸ h![]) = outTp) → STHoare p Γ ⟦⟧ (ho ▸ fn.body _ (hg ▸ h![]) (hi ▸ args)) Q → STHoare p Γ ⟦⟧ (.call h![] argTps outTp (.lambda lambdaRef) args) Q := by - unfold STHoare - intros - unfold THoare - intros - constructor - <;> tauto + sorry end Lampe.STHoare diff --git a/Lampe/Lampe/State.lean b/Lampe/Lampe/State.lean index d5cda8e..56ee060 100644 --- a/Lampe/Lampe/State.lean +++ b/Lampe/Lampe/State.lean @@ -11,9 +11,9 @@ lemma Finmap.singleton_disjoint_of_not_mem (hp : ref ∉ s): namespace Lampe -def AnyValue (P : Prime) := (tp : Tp) × tp.denote P +def AnyValue (p : Prime) := (tp : Tp) × tp.denote p -abbrev State (P : Prime) := Finmap (fun (_ : Ref) => AnyValue P) +abbrev State (p : Prime) := Finmap (fun (_ : Ref) => AnyValue p) namespace State From 9c21543acd21d081690cfd12aa9d0e4f5de78916 Mon Sep 17 00:00:00 2001 From: utkn Date: Thu, 21 Nov 2024 09:06:46 +0300 Subject: [PATCH 06/39] Introduced a new state that also contains closures --- Lampe/Lampe/Ast.lean | 2 +- Lampe/Lampe/Builtin/Basic.lean | 21 +++-- Lampe/Lampe/Hoare/SepTotal.lean | 14 --- Lampe/Lampe/Hoare/Total.lean | 6 +- Lampe/Lampe/Semantics.lean | 115 +++++++++++------------- Lampe/Lampe/SeparationLogic.lean | 4 +- Lampe/Lampe/State.lean | 25 +++--- Lampe/Lampe/Tactic/SeparationLogic.lean | 8 +- Lampe/Lampe/ValHeap.lean | 18 ++++ 9 files changed, 105 insertions(+), 108 deletions(-) create mode 100644 Lampe/Lampe/ValHeap.lean diff --git a/Lampe/Lampe/Ast.lean b/Lampe/Lampe/Ast.lean index aafefac..018ffeb 100644 --- a/Lampe/Lampe/Ast.lean +++ b/Lampe/Lampe/Ast.lean @@ -1,7 +1,7 @@ import Mathlib import Lampe.Tp import Lampe.Data.HList -import Lampe.State +import Lampe.ValHeap import Lampe.Builtin namespace Lampe diff --git a/Lampe/Lampe/Builtin/Basic.lean b/Lampe/Lampe/Builtin/Basic.lean index a7c6eb4..24f56bc 100644 --- a/Lampe/Lampe/Builtin/Basic.lean +++ b/Lampe/Lampe/Builtin/Basic.lean @@ -1,4 +1,4 @@ -import Lampe.State +import Lampe.ValHeap import Lampe.Data.Field import Lampe.Data.HList import Lampe.SeparationLogic @@ -8,22 +8,25 @@ import Mathlib namespace Lampe abbrev Builtin.Omni := ∀(P:Prime), - State P → + ValHeap P → (argTps : List Tp) → (outTp : Tp) → HList (Tp.denote P) argTps → - (Option (State P × Tp.denote P outTp) → Prop) → + (Option (ValHeap P × Tp.denote P outTp) → Prop) → Prop def Builtin.omni_conseq (omni : Builtin.Omni) : Prop := - ∀{P st argTps outTp args Q Q'}, omni P st argTps outTp args Q → (∀ r, Q r → Q' r) → omni P st argTps outTp args Q' + ∀ {P st argTps outTp args Q Q'}, + omni P st argTps outTp args Q → (∀ r, Q r → Q' r) → omni P st argTps outTp args Q' def Builtin.omni_frame (omni : Builtin.Omni) : Prop := - ∀{P st₁ st₂ argTps outTp args Q}, omni P st₁ argTps outTp args Q → st₁.Disjoint st₂ → - omni P (st₁ ∪ st₂) argTps outTp args (fun r => match r with - | some (st, v) => ((fun st => Q (some (st, v))) ⋆ (fun st => st = st₂)) st - | none => Q none - ) + ∀ {P st₁ st₂ argTps outTp args Q}, + omni P st₁ argTps outTp args Q → + st₁.Disjoint st₂ → + omni P (st₁ ∪ st₂) argTps outTp args (fun r => match r with + | some (st, v) => ((fun st => Q (some (st, v))) ⋆ (fun st => st = st₂)) st + | none => Q none + ) structure Builtin where omni : Builtin.Omni diff --git a/Lampe/Lampe/Hoare/SepTotal.lean b/Lampe/Lampe/Hoare/SepTotal.lean index 562a085..5e1fe64 100644 --- a/Lampe/Lampe/Hoare/SepTotal.lean +++ b/Lampe/Lampe/Hoare/SepTotal.lean @@ -306,18 +306,4 @@ theorem skip_intro : . apply SLP.ent_star_top tauto -theorem lambda_intro {ls : LambdaStore} : - ls lambdaRef = some (newLambda argTps outTp body) → - STHoare p Γ ⟦⟧ (.lambda argTps outTp (body _ h![])) (fun r => r = lambdaRef) := by - sorry - -theorem lambdaCall_intro {ls : LambdaStore} : - ls lambdaRef = some fn → - (hg : fn.generics = []) → - (hi : fn.inTps (hg ▸ h![]) = argTps) → - (ho : fn.outTp (hg ▸ h![]) = outTp) → - STHoare p Γ ⟦⟧ (ho ▸ fn.body _ (hg ▸ h![]) (hi ▸ args)) Q → - STHoare p Γ ⟦⟧ (.call h![] argTps outTp (.lambda lambdaRef) args) Q := by - sorry - end Lampe.STHoare diff --git a/Lampe/Lampe/Hoare/Total.lean b/Lampe/Lampe/Hoare/Total.lean index bfa4866..85052f3 100644 --- a/Lampe/Lampe/Hoare/Total.lean +++ b/Lampe/Lampe/Hoare/Total.lean @@ -12,9 +12,9 @@ def THoare (P : SLP p) (e : Expr (Tp.denote p) tp) (Q : (tp.denote p) → SLP p): Prop := - ∀st, P st → Omni p Γ st e (fun r => match r with + ∀st, P st.vals → Omni p Γ st e (fun r => match r with | none => True - | some (st', v) => Q v st') + | some (st', v) => Q v st'.vals) namespace THoare @@ -50,7 +50,7 @@ theorem assert_intro {v: Bool} (h : v ⋆ P ⊢ Q ()): cases v · apply Builtin.genericPureOmni.err · simp - · simp + · tauto · exact () · apply Builtin.genericPureOmni.ok · simp_all; tauto diff --git a/Lampe/Lampe/Semantics.lean b/Lampe/Lampe/Semantics.lean index 777fe8e..d4ae593 100644 --- a/Lampe/Lampe/Semantics.lean +++ b/Lampe/Lampe/Semantics.lean @@ -7,8 +7,6 @@ import Lampe.State namespace Lampe -variable (P : Prime) - def Env := Ident → Option Function def Env.ofModule (m : Module) : Env := fun i => (m.decls.find? fun d => d.name == i).map (·.fn) @@ -16,52 +14,52 @@ def Env.ofModule (m : Module) : Env := fun i => (m.decls.find? fun d => d.name = @[reducible] def Env.extend (Γ₁ : Env) (Γ₂ : Env) : Env := fun i => Γ₁ i <|> Γ₂ i -inductive Omni : Env → State P → Expr (Tp.denote P) tp → (Option (State P × tp.denote P) → Prop) → Prop where -| litField : Q (some (st, n)) → Omni Γ st (.lit .field n) Q -| litFalse : Q (some (st, false)) → Omni Γ st (.lit .bool 0) Q -| litTrue : Q (some (st, true)) → Omni Γ st (.lit .bool 1) Q -| litU {Q} : Q (some (st, ↑n)) → Omni Γ st (.lit (.u s) n) Q -| var : Q (some (st, v)) → Omni Γ st (.var v) Q -| skip : Q (some (st, ())) → Omni Γ st (.skip) Q +inductive Omni : (p : Prime) → Env → State p → Expr (Tp.denote p) tp → (Option (State p × Tp.denote p tp) → Prop) → Prop where +| litField {Q} : Q (some (st, n)) → Omni p Γ st (.lit .field n) Q +| litFalse {Q} : Q (some (st, false)) → Omni p Γ st (.lit .bool 0) Q +| litTrue {Q} : Q (some (st, true)) → Omni p Γ st (.lit .bool 1) Q +| litU {Q} : Q (some (st, ↑n)) → Omni p Γ st (.lit (.u s) n) Q +| var {Q} : Q (some (st, v)) → Omni p Γ st (.var v) Q +| skip {Q} : Q (some (st, ())) → Omni p Γ st (.skip) Q | iteTrue {mainBranch elseBranch} : - Omni Γ st mainBranch Q → - Omni Γ st (Expr.ite true mainBranch elseBranch) Q + Omni p Γ st mainBranch Q → + Omni p Γ st (Expr.ite true mainBranch elseBranch) Q | iteFalse {mainBranch elseBranch} : - Omni Γ st elseBranch Q → - Omni Γ st (Expr.ite false mainBranch elseBranch) Q + Omni p Γ st elseBranch Q → + Omni p Γ st (Expr.ite false mainBranch elseBranch) Q | letIn : - Omni Γ st e Q₁ → - (∀v s, Q₁ (some (s, v)) → Omni Γ s (b v) Q) → + Omni p Γ ⟨vh₁, cls⟩ e Q₁ → + (∀v vh, Q₁ (some (⟨vh, cls⟩, v)) → Omni p Γ ⟨vh, cls⟩ (b v) Q) → (Q₁ none → Q none) → - Omni Γ st (.letIn e b) Q -| callBuiltin {Q} : - (b.omni P st argTypes resType args Q) → - Omni Γ st (Expr.call h![] argTypes resType (.builtin b) args) Q + Omni p Γ ⟨vh₁, cls⟩ (.letIn e b) Q +| callBuiltin {Q} {vh₁ cls} : + (b.omni p vh₁ argTypes resType args (mapToValHeapCondition Q cls)) → + Omni p Γ ⟨vh₁, cls⟩ (Expr.call h![] argTypes resType (.builtin b) args) Q | callDecl: Γ fname = some fn → (hkc : fn.generics = tyKinds) → (htci : fn.inTps (hkc ▸ generics) = argTypes) → (htco : fn.outTp (hkc ▸ generics) = res) → - Omni Γ st (htco ▸ fn.body _ (hkc ▸ generics) (htci ▸ args)) Q → - Omni Γ st (@Expr.call _ tyKinds generics argTypes res (.decl fname) args) Q -| callLambda : - Γ (Ident.ofLambdaRef lambdaRef) = some fn → - (hg : fn.generics = []) → - (hi : fn.inTps (hg ▸ h![]) = argTps) → - (ho : fn.outTp (hg ▸ h![]) = outTp) → - Omni Γ st (ho ▸ fn.body _ (hg ▸ h![]) (hi ▸ args)) Q → - Omni Γ st (Expr.call h![] argTps outTp (.lambda lambdaRef) args) Q -| newLambda {Q} : - Q (some (st, lambdaRef)) → - Γ (Ident.ofLambdaRef lambdaRef) = some (newLambda argTps outTp body) → - Omni Γ st (Expr.lambda argTps outTp (body _ _)) Q + Omni p Γ st (htco ▸ fn.body _ (hkc ▸ generics) (htci ▸ args)) Q → + Omni p Γ st (@Expr.call _ tyKinds generics argTypes res (.decl fname) args) Q +-- | callLambda : +-- Γ (Ident.ofLambdaRef lambdaRef) = some fn → +-- (hg : fn.generics = []) → +-- (hi : fn.inTps (hg ▸ h![]) = argTps) → +-- (ho : fn.outTp (hg ▸ h![]) = outTp) → +-- Omni p Γ st (ho ▸ fn.body _ (hg ▸ h![]) (hi ▸ args)) Q → +-- Omni p Γ st (Expr.call h![] argTps outTp (.lambda lambdaRef) args) Q +-- | newLambda {Q} : +-- Q (some (st, lambdaRef)) → +-- Γ (Ident.ofLambdaRef lambdaRef) = some (newLambda argTps outTp body) → +-- Omni p Γ st (Expr.lambda argTps outTp (body _ _)) Q | loopDone : lo ≥ hi → - Omni Γ st (.loop lo hi body) Q + Omni p Γ st (.loop lo hi body) Q | loopNext {s} {lo hi : U s} {body} : lo < hi → - Omni Γ st (.letIn (body lo) (fun _ => .loop (lo + 1) hi body)) Q → - Omni Γ st (.loop lo hi body) Q + Omni p Γ st (.letIn (body lo) (fun _ => .loop (lo + 1) hi body)) Q → + Omni p Γ st (.loop lo hi body) Q theorem Omni.consequence {p Γ st tp} {e : Expr (Tp.denote p) tp} {Q Q'}: Omni p Γ st e Q → @@ -81,41 +79,48 @@ theorem Omni.consequence {p Γ st tp} {e : Expr (Tp.denote p) tp} {Q Q'}: apply loopNext (by assumption) tauto -theorem Omni.frame {p Γ tp st₁ st₂} {e : Expr (Tp.denote p) tp} {Q}: - Omni p Γ st₁ e Q → - st₁.Disjoint st₂ → - Omni p Γ (st₁ ∪ st₂) e (fun st => match st with - | some (st', v) => ((fun st => Q (some (st, v))) ⋆ (fun st => st = st₂)) st' +theorem Omni.frame {p Γ tp} {vh₁ vh₂ : ValHeap p} {cls : Closures} {e : Expr (Tp.denote p) tp} {Q}: + Omni p Γ ⟨vh₁, cls⟩ e Q → + vh₁.Disjoint vh₂ → + Omni p Γ ⟨(vh₁ ∪ vh₂), cls⟩ e (fun stv => match stv with + | some (st, v) => ((fun vals => Q (some (⟨vals, cls⟩, v))) ⋆ (fun vals => vals = vh₂)) st.vals | none => Q none ) := by intro h + generalize hc : State.mk vh₁ cls = st₁ + rw [hc] at h induction h with | litField hq - | skip + | skip hq | litFalse hq | litTrue hq - | litU _ + | litU hq | var hq => intro + subst hc constructor - repeat apply Exists.intro - tauto + exists vh₁, vh₂ | letIn _ _ hN ihE ihB => intro constructor apply ihE assumption + assumption · intro _ _ h cases h casesm* ∃ _, _, _∧_ - subst_vars + simp_all only apply ihB - assumption - assumption + simp_all only [State.mk.injEq] + tauto + . aesop + . assumption · simp_all | callBuiltin hq => cases_type Builtin - tauto + intro + constructor + aesop | callDecl _ _ _ _ _ ih => intro constructor @@ -134,18 +139,6 @@ theorem Omni.frame {p Γ tp st₁ st₂} {e : Expr (Tp.denote p) tp} {Q}: intro constructor apply ih - assumption - | newLambda => - intros - constructor - all_goals (try tauto) - constructor - tauto - | callLambda ih => - intros - constructor - apply ih - all_goals (try assumption) - tauto + <;> tauto end Lampe diff --git a/Lampe/Lampe/SeparationLogic.lean b/Lampe/Lampe/SeparationLogic.lean index 2f6b19e..d9d7444 100644 --- a/Lampe/Lampe/SeparationLogic.lean +++ b/Lampe/Lampe/SeparationLogic.lean @@ -1,10 +1,10 @@ -import Lampe.State +import Lampe.ValHeap import Lampe.Tactic.IntroCases namespace Lampe -def SLP (p : Prime) : Type := State p → Prop +def SLP (p : Prime) : Type := ValHeap p → Prop namespace SLP diff --git a/Lampe/Lampe/State.lean b/Lampe/Lampe/State.lean index 56ee060..394a638 100644 --- a/Lampe/Lampe/State.lean +++ b/Lampe/Lampe/State.lean @@ -1,22 +1,19 @@ -import Lampe.Tp import Mathlib - -lemma Finmap.insert_eq_singleton_union [DecidableEq α] {ref : α}: - m.insert ref v = Finmap.singleton ref v ∪ m := by rfl - -@[simp] -lemma Finmap.singleton_disjoint_of_not_mem (hp : ref ∉ s): - Finmap.Disjoint (Finmap.singleton ref v) s := by - simp_all [Finmap.Disjoint] +import Lampe.ValHeap +import Lampe.Ast namespace Lampe -def AnyValue (p : Prime) := (tp : Tp) × tp.denote p - -abbrev State (p : Prime) := Finmap (fun (_ : Ref) => AnyValue p) +abbrev Closures := Finmap fun (_ : Ref) => Function -namespace State +structure State (p : Prime) where + vals : ValHeap p + funcs : Closures -end State +@[reducible] +def mapToValHeapCondition + (Q : Option (State p × T) → Prop) + (closures : Closures) : Option (ValHeap p × T) → Prop := + fun vv => Q (vv.map fun (vals, v) => ⟨⟨vals, closures⟩, v⟩) end Lampe diff --git a/Lampe/Lampe/Tactic/SeparationLogic.lean b/Lampe/Lampe/Tactic/SeparationLogic.lean index 19883ea..325989c 100644 --- a/Lampe/Lampe/Tactic/SeparationLogic.lean +++ b/Lampe/Lampe/Tactic/SeparationLogic.lean @@ -452,8 +452,8 @@ macro "stephelper1" : tactic => `(tactic|( | apply fresh_intro | apply assert_intro | apply skip_intro - | apply lambdaCall_intro - | apply lambda_intro + -- | apply lambdaCall_intro + -- | apply lambda_intro -- memory builtins | apply var_intro | apply ref_intro @@ -561,8 +561,8 @@ macro "stephelper3" : tactic => `(tactic|( | apply ramified_frame_top Lampe.STHoare.litU_intro | apply ramified_frame_top assert_intro | apply ramified_frame_top skip_intro - | apply ramified_frame_top lambdaCall_intro - | apply ramified_frame_top lambda_intro + -- | apply ramified_frame_top lambdaCall_intro + -- | apply ramified_frame_top lambda_intro -- memory builtins | apply ramified_frame_top var_intro | apply ramified_frame_top ref_intro diff --git a/Lampe/Lampe/ValHeap.lean b/Lampe/Lampe/ValHeap.lean new file mode 100644 index 0000000..14cdf1f --- /dev/null +++ b/Lampe/Lampe/ValHeap.lean @@ -0,0 +1,18 @@ +import Mathlib +import Lampe.Tp + +lemma Finmap.insert_eq_singleton_union [DecidableEq α] {ref : α}: + m.insert ref v = Finmap.singleton ref v ∪ m := by rfl + +@[simp] +lemma Finmap.singleton_disjoint_of_not_mem (hp : ref ∉ s): + Finmap.Disjoint (Finmap.singleton ref v) s := by + simp_all [Finmap.Disjoint] + +namespace Lampe + +def AnyValue (p : Prime) := (tp : Tp) × tp.denote p + +abbrev ValHeap (p : Prime) := Finmap (fun (_ : Ref) => AnyValue p) + +end Lampe From 588ebe3329b6ca34cdf8b3be081c40642051dec4 Mon Sep 17 00:00:00 2001 From: utkn Date: Thu, 21 Nov 2024 09:12:31 +0300 Subject: [PATCH 07/39] Proof fixes --- Lampe/Lampe/Hoare/Builtins.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Lampe/Lampe/Hoare/Builtins.lean b/Lampe/Lampe/Hoare/Builtins.lean index 68c34f1..f022407 100644 --- a/Lampe/Lampe/Hoare/Builtins.lean +++ b/Lampe/Lampe/Hoare/Builtins.lean @@ -17,13 +17,13 @@ theorem pureBuiltin_intro {A : Type} {a : A} {sgn desc args} : constructor cases em (desc a args).fst . apply Builtin.genericPureOmni.ok - . simp_all only [SLP.true_star, exists_const] + . simp_all only [mapToValHeapCondition, SLP.true_star, exists_const] apply SLP.ent_star_top - assumption + simp_all only [SLP.true_star, exists_const] . tauto . apply Builtin.genericPureOmni.err . tauto - . simp + . simp_all [mapToValHeapCondition] lemma pureBuiltin_intro_consequence {A : Type} {a : A} {sgn desc args} {Q : Tp.denote p outTp → Prop} From 2b29e38751c9c11b6c8ec9a9a6916ed526f7ae55 Mon Sep 17 00:00:00 2001 From: utkn Date: Thu, 21 Nov 2024 09:29:47 +0300 Subject: [PATCH 08/39] Lambda semantics --- Lampe/Lampe/Semantics.lean | 30 +++++++++++++++---------- Lampe/Lampe/State.lean | 4 ++-- Lampe/Lampe/Tactic/SeparationLogic.lean | 6 ----- 3 files changed, 20 insertions(+), 20 deletions(-) diff --git a/Lampe/Lampe/Semantics.lean b/Lampe/Lampe/Semantics.lean index d4ae593..4221781 100644 --- a/Lampe/Lampe/Semantics.lean +++ b/Lampe/Lampe/Semantics.lean @@ -35,24 +35,24 @@ inductive Omni : (p : Prime) → Env → State p → Expr (Tp.denote p) tp → ( | callBuiltin {Q} {vh₁ cls} : (b.omni p vh₁ argTypes resType args (mapToValHeapCondition Q cls)) → Omni p Γ ⟨vh₁, cls⟩ (Expr.call h![] argTypes resType (.builtin b) args) Q -| callDecl: +| callDecl : Γ fname = some fn → (hkc : fn.generics = tyKinds) → (htci : fn.inTps (hkc ▸ generics) = argTypes) → (htco : fn.outTp (hkc ▸ generics) = res) → Omni p Γ st (htco ▸ fn.body _ (hkc ▸ generics) (htci ▸ args)) Q → Omni p Γ st (@Expr.call _ tyKinds generics argTypes res (.decl fname) args) Q --- | callLambda : --- Γ (Ident.ofLambdaRef lambdaRef) = some fn → --- (hg : fn.generics = []) → --- (hi : fn.inTps (hg ▸ h![]) = argTps) → --- (ho : fn.outTp (hg ▸ h![]) = outTp) → --- Omni p Γ st (ho ▸ fn.body _ (hg ▸ h![]) (hi ▸ args)) Q → --- Omni p Γ st (Expr.call h![] argTps outTp (.lambda lambdaRef) args) Q --- | newLambda {Q} : --- Q (some (st, lambdaRef)) → --- Γ (Ident.ofLambdaRef lambdaRef) = some (newLambda argTps outTp body) → --- Omni p Γ st (Expr.lambda argTps outTp (body _ _)) Q +| callLambda {cls : Closures} : + cls.get? lambdaRef.val = some fn → + (hg : fn.generics = []) → + (hi : fn.inTps (hg ▸ h![]) = argTps) → + (ho : fn.outTp (hg ▸ h![]) = outTp) → + Omni p Γ st (ho ▸ fn.body _ (hg ▸ h![]) (hi ▸ args)) Q → + Omni p Γ st (Expr.call h![] argTps outTp (.lambda lambdaRef) args) Q +| newLambda {Q} : + cls' = cls ++ [(newLambda argTps outTp body)] → + Omni p Γ ⟨vh, cls'⟩ (Expr.lit .lambdaRef cls.length) Q → + Omni p Γ ⟨vh, cls⟩ (Expr.lambda argTps outTp (body _ _)) Q | loopDone : lo ≥ hi → Omni p Γ st (.loop lo hi body) Q @@ -140,5 +140,11 @@ theorem Omni.frame {p Γ tp} {vh₁ vh₂ : ValHeap p} {cls : Closures} {e : Ex constructor apply ih <;> tauto + | newLambda => + intro + tauto + | callLambda => + intro + constructor <;> try tauto end Lampe diff --git a/Lampe/Lampe/State.lean b/Lampe/Lampe/State.lean index 394a638..df41b1d 100644 --- a/Lampe/Lampe/State.lean +++ b/Lampe/Lampe/State.lean @@ -4,11 +4,11 @@ import Lampe.Ast namespace Lampe -abbrev Closures := Finmap fun (_ : Ref) => Function +abbrev Closures := List Function structure State (p : Prime) where vals : ValHeap p - funcs : Closures + cls : Closures @[reducible] def mapToValHeapCondition diff --git a/Lampe/Lampe/Tactic/SeparationLogic.lean b/Lampe/Lampe/Tactic/SeparationLogic.lean index 325989c..996ebfe 100644 --- a/Lampe/Lampe/Tactic/SeparationLogic.lean +++ b/Lampe/Lampe/Tactic/SeparationLogic.lean @@ -452,8 +452,6 @@ macro "stephelper1" : tactic => `(tactic|( | apply fresh_intro | apply assert_intro | apply skip_intro - -- | apply lambdaCall_intro - -- | apply lambda_intro -- memory builtins | apply var_intro | apply ref_intro @@ -505,8 +503,6 @@ macro "stephelper2" : tactic => `(tactic|( | apply consequence_frame_left fresh_intro | apply consequence_frame_left Lampe.STHoare.litU_intro | apply consequence_frame_left assert_intro - -- | apply consequence_frame_left lambdaCall_intro - -- | apply consequence_frame_left lambda_intro -- | apply consequence_frame_left skip_intro -- memory builtins | apply consequence_frame_left var_intro @@ -561,8 +557,6 @@ macro "stephelper3" : tactic => `(tactic|( | apply ramified_frame_top Lampe.STHoare.litU_intro | apply ramified_frame_top assert_intro | apply ramified_frame_top skip_intro - -- | apply ramified_frame_top lambdaCall_intro - -- | apply ramified_frame_top lambda_intro -- memory builtins | apply ramified_frame_top var_intro | apply ramified_frame_top ref_intro From 5741a355a9fa0b851e2656b32b9630a39fe70f17 Mon Sep 17 00:00:00 2001 From: utkn Date: Thu, 21 Nov 2024 14:28:51 +0300 Subject: [PATCH 09/39] various changes --- Lampe/Lampe/Ast.lean | 8 ++++---- Lampe/Lampe/Hoare/Builtins.lean | 10 +++++++--- Lampe/Lampe/Hoare/Total.lean | 17 ----------------- Lampe/Lampe/Semantics.lean | 9 ++------- 4 files changed, 13 insertions(+), 31 deletions(-) diff --git a/Lampe/Lampe/Ast.lean b/Lampe/Lampe/Ast.lean index 018ffeb..2922eb0 100644 --- a/Lampe/Lampe/Ast.lean +++ b/Lampe/Lampe/Ast.lean @@ -11,9 +11,6 @@ abbrev Ident := String /-- A reference to a lambda is represented as a reference to a unit type -/ abbrev Tp.lambdaRef := Tp.ref .unit -/-- Converts a lambda reference to a function identity -/ -def Ident.ofLambdaRef (lambdaRef : Ref) : Ident := toString lambdaRef.val - inductive FunctionIdent {rep : Tp → Type} : Type where | builtin : Builtin → FunctionIdent | decl : Ident → FunctionIdent -- a function declared at the module level @@ -49,8 +46,11 @@ example : Function := { body := fun _ h![_] h![x] => .var x } +abbrev Lambda (argTps outTp) := + (rep : Tp → Type) → HList Kind.denote [] → HList rep argTps → Expr rep outTp + def newLambda (argTps : List Tp) (outTp : Tp) -(body : (rep : Tp → Type) → HList Kind.denote [] → HList rep argTps → Expr rep outTp) : Function := { +(body : Lambda argTps outTp) : Function := { generics := [] inTps := fun _ => argTps outTp := fun _ => outTp diff --git a/Lampe/Lampe/Hoare/Builtins.lean b/Lampe/Lampe/Hoare/Builtins.lean index f022407..46040cd 100644 --- a/Lampe/Lampe/Hoare/Builtins.lean +++ b/Lampe/Lampe/Hoare/Builtins.lean @@ -409,9 +409,13 @@ theorem writeRef_intro: -- Misc theorem assert_intro : STHoarePureBuiltin p Γ Builtin.assert (by tauto) h![a] (a := ()) := by - unfold STHoarePureBuiltin + apply pureBuiltin_intro_consequence <;> tauto + tauto + +theorem newLambda_intro {cl : Closures} : + STHoare p Γ P (.lambda argTps outTp body) Q := by + unfold STHoare intro H - apply THoare.assert_intro - simp [SLP.entails_self, SLP.star_mono_l] + sorry end Lampe.STHoare diff --git a/Lampe/Lampe/Hoare/Total.lean b/Lampe/Lampe/Hoare/Total.lean index 85052f3..5bfb9c9 100644 --- a/Lampe/Lampe/Hoare/Total.lean +++ b/Lampe/Lampe/Hoare/Total.lean @@ -40,23 +40,6 @@ theorem var_intro {v} {P : Tp.denote p tp → SLP p}: THoare p Γ (P v) (.var v) P := by tauto -/-- [TODO] there's probably a generic lemma for pure builtins to abstract this proof structure? -/ -theorem assert_intro {v: Bool} (h : v ⋆ P ⊢ Q ()): - THoare p Γ P (.call h![] [.bool] .unit (.builtin Builtin.assert) h![v]) Q := by - unfold THoare - intros - constructor - simp only [Builtin.assert] - cases v - · apply Builtin.genericPureOmni.err - · simp - · tauto - · exact () - · apply Builtin.genericPureOmni.ok - · simp_all; tauto - · exact () - · simp_all - theorem letIn_intro {P Q} (h₁ : THoare p Γ H e₁ P) (h₂ : ∀v, THoare p Γ (P v) (e₂ v) Q): diff --git a/Lampe/Lampe/Semantics.lean b/Lampe/Lampe/Semantics.lean index 4221781..11feccc 100644 --- a/Lampe/Lampe/Semantics.lean +++ b/Lampe/Lampe/Semantics.lean @@ -18,6 +18,7 @@ inductive Omni : (p : Prime) → Env → State p → Expr (Tp.denote p) tp → ( | litField {Q} : Q (some (st, n)) → Omni p Γ st (.lit .field n) Q | litFalse {Q} : Q (some (st, false)) → Omni p Γ st (.lit .bool 0) Q | litTrue {Q} : Q (some (st, true)) → Omni p Γ st (.lit .bool 1) Q +| litRef {Q} : Q (some (st, ⟨r⟩)) → Omni p Γ st (.lit (.ref tp) r) Q | litU {Q} : Q (some (st, ↑n)) → Omni p Γ st (.lit (.u s) n) Q | var {Q} : Q (some (st, v)) → Omni p Γ st (.var v) Q | skip {Q} : Q (some (st, ())) → Omni p Γ st (.skip) Q @@ -49,10 +50,6 @@ inductive Omni : (p : Prime) → Env → State p → Expr (Tp.denote p) tp → ( (ho : fn.outTp (hg ▸ h![]) = outTp) → Omni p Γ st (ho ▸ fn.body _ (hg ▸ h![]) (hi ▸ args)) Q → Omni p Γ st (Expr.call h![] argTps outTp (.lambda lambdaRef) args) Q -| newLambda {Q} : - cls' = cls ++ [(newLambda argTps outTp body)] → - Omni p Γ ⟨vh, cls'⟩ (Expr.lit .lambdaRef cls.length) Q → - Omni p Γ ⟨vh, cls⟩ (Expr.lambda argTps outTp (body _ _)) Q | loopDone : lo ≥ hi → Omni p Γ st (.loop lo hi body) Q @@ -95,6 +92,7 @@ theorem Omni.frame {p Γ tp} {vh₁ vh₂ : ValHeap p} {cls : Closures} {e : Ex | litFalse hq | litTrue hq | litU hq + | litRef hq | var hq => intro subst hc @@ -140,9 +138,6 @@ theorem Omni.frame {p Γ tp} {vh₁ vh₂ : ValHeap p} {cls : Closures} {e : Ex constructor apply ih <;> tauto - | newLambda => - intro - tauto | callLambda => intro constructor <;> try tauto From f488e5e7bd8d6dbb160d02af96fa835e14c7710f Mon Sep 17 00:00:00 2001 From: utkn Date: Thu, 21 Nov 2024 19:36:55 +0300 Subject: [PATCH 10/39] SLH typeclass --- Lampe/Lampe/Ast.lean | 3 +- Lampe/Lampe/Basic.lean | 3 +- Lampe/Lampe/Builtin/Basic.lean | 3 +- Lampe/Lampe/Builtin/Memory.lean | 2 +- Lampe/Lampe/Hoare/Builtins.lean | 58 +++-- Lampe/Lampe/Hoare/SepTotal.lean | 22 +- Lampe/Lampe/Hoare/Total.lean | 20 +- Lampe/Lampe/Semantics.lean | 60 ++--- Lampe/Lampe/SeparationLogic.lean | 269 ----------------------- Lampe/Lampe/SeparationLogic/SLH.lean | 58 +++++ Lampe/Lampe/SeparationLogic/SLP.lean | 263 ++++++++++++++++++++++ Lampe/Lampe/SeparationLogic/State.lean | 103 +++++++++ Lampe/Lampe/SeparationLogic/ValHeap.lean | 34 +++ Lampe/Lampe/State.lean | 19 -- Lampe/Lampe/Tactic/SeparationLogic.lean | 45 ++-- Lampe/Lampe/ValHeap.lean | 18 -- 16 files changed, 573 insertions(+), 407 deletions(-) delete mode 100644 Lampe/Lampe/SeparationLogic.lean create mode 100644 Lampe/Lampe/SeparationLogic/SLH.lean create mode 100644 Lampe/Lampe/SeparationLogic/SLP.lean create mode 100644 Lampe/Lampe/SeparationLogic/State.lean create mode 100644 Lampe/Lampe/SeparationLogic/ValHeap.lean delete mode 100644 Lampe/Lampe/State.lean delete mode 100644 Lampe/Lampe/ValHeap.lean diff --git a/Lampe/Lampe/Ast.lean b/Lampe/Lampe/Ast.lean index 2922eb0..a9c7ad2 100644 --- a/Lampe/Lampe/Ast.lean +++ b/Lampe/Lampe/Ast.lean @@ -1,7 +1,8 @@ import Mathlib + import Lampe.Tp import Lampe.Data.HList -import Lampe.ValHeap +import Lampe.SeparationLogic.ValHeap import Lampe.Builtin namespace Lampe diff --git a/Lampe/Lampe/Basic.lean b/Lampe/Lampe/Basic.lean index d35b344..149b4f4 100644 --- a/Lampe/Lampe/Basic.lean +++ b/Lampe/Lampe/Basic.lean @@ -1,8 +1,7 @@ import Lampe.Semantics -import Lampe.SeparationLogic import Lampe.Syntax import Lampe.Ast import Lampe.Tp import Lampe.Hoare.Total import Lampe.Hoare.SepTotal -import Lampe.Tactic.SeparationLogic +-- import Lampe.Tactic.SeparationLogic diff --git a/Lampe/Lampe/Builtin/Basic.lean b/Lampe/Lampe/Builtin/Basic.lean index 24f56bc..051d5d9 100644 --- a/Lampe/Lampe/Builtin/Basic.lean +++ b/Lampe/Lampe/Builtin/Basic.lean @@ -1,7 +1,6 @@ -import Lampe.ValHeap +import Lampe.SeparationLogic.ValHeap import Lampe.Data.Field import Lampe.Data.HList -import Lampe.SeparationLogic import Lampe.Builtin.Helpers import Mathlib diff --git a/Lampe/Lampe/Builtin/Memory.lean b/Lampe/Lampe/Builtin/Memory.lean index c788c50..a55ced4 100644 --- a/Lampe/Lampe/Builtin/Memory.lean +++ b/Lampe/Lampe/Builtin/Memory.lean @@ -27,7 +27,7 @@ def ref : Builtin := { simp_all rfl simp [Finmap.insert_union] - simp_all [Finmap.insert_eq_singleton_union, Finmap.disjoint_union_left] + simp_all [Finmap.insert_eq_singleton_union, SLH.disjoint, Finmap.disjoint_union_left] } inductive readRefOmni : Omni where diff --git a/Lampe/Lampe/Hoare/Builtins.lean b/Lampe/Lampe/Hoare/Builtins.lean index 46040cd..f4a4191 100644 --- a/Lampe/Lampe/Hoare/Builtins.lean +++ b/Lampe/Lampe/Hoare/Builtins.lean @@ -20,6 +20,7 @@ theorem pureBuiltin_intro {A : Type} {a : A} {sgn desc args} : . simp_all only [mapToValHeapCondition, SLP.true_star, exists_const] apply SLP.ent_star_top simp_all only [SLP.true_star, exists_const] + tauto . tauto . apply Builtin.genericPureOmni.err . tauto @@ -356,12 +357,19 @@ theorem ref_intro: apply THoare.consequence ?_ THoare.ref_intro (fun _ => SLP.entails_self) simp only [SLP.true_star] intro st hH r hr - exists (Finmap.singleton r ⟨tp, v⟩ ∪ st), ∅ - apply And.intro (by simp) - apply And.intro (by simp [Finmap.insert_eq_singleton_union]) - apply And.intro ?_ (by simp) - exists (Finmap.singleton r ⟨tp, v⟩), st - simp_all [SLP.singleton] + exists (⟨Finmap.singleton r ⟨tp, v⟩, st.closures⟩ ∪ st), ∅ + apply And.intro (by rw [SLH.disjoint_symm_iff]; apply SLH.disjoint_empty) + constructor + . simp only [State.insertVal, Finmap.insert_eq_singleton_union, SLH_union_empty] + simp only [State.union_parts, Finmap.union_self] + . apply And.intro ?_ (by simp) + exists (⟨Finmap.singleton r ⟨tp, v⟩, ∅⟩), st + constructor + . simp only [SLH.disjoint] + apply And.intro (by simp [Finmap.singleton_disjoint_of_not_mem hr]) (by tauto) + . simp_all only + apply And.intro _ (by trivial) + simp only [State.union_parts, Finmap.empty_union, Finmap.union_self] theorem readRef_intro: STHoare p Γ @@ -375,14 +383,15 @@ theorem readRef_intro: intro st rintro ⟨_, _, _, _, hs, _⟩ subst_vars - apply And.intro (by simp; rfl) - simp only [SLP.true_star, SLP.star_assoc] - exists (Finmap.singleton r ⟨tp, v⟩), ?_ - apply And.intro (by assumption) - apply And.intro rfl - apply And.intro (by simp [SLP.singleton]) - apply SLP.ent_star_top - assumption + all_goals sorry + -- apply And.intro (by simp; rfl) + -- simp only [SLP.true_star, SLP.star_assoc] + -- exists (Finmap.singleton r ⟨tp, v⟩), ?_ + -- apply And.intro (by assumption) + -- apply And.intro rfl + -- apply And.intro (by simp [SLP.singleton]) + -- apply SLP.ent_star_top + -- assumption theorem writeRef_intro: STHoare p Γ @@ -394,17 +403,18 @@ theorem writeRef_intro: apply THoare.consequence ?_ THoare.writeRef_intro (fun _ => SLP.entails_self) intro st rintro ⟨_, _, _, _, hs, _⟩ - simp only [SLP.singleton] at hs + simp only [State.singleton] at hs subst_vars - apply And.intro (by simp) - simp only - simp only [Finmap.insert_eq_singleton_union, ←Finmap.union_assoc, Finmap.union_singleton, SLP.star_assoc] - use Finmap.singleton r ⟨tp, v'⟩, ?_ - apply And.intro (by assumption) - apply And.intro rfl - apply And.intro (by simp [SLP.singleton]) - apply SLP.ent_star_top - assumption + sorry + -- apply And.intro (by simp) + -- simp only + -- simp only [Finmap.insert_eq_singleton_union, ←Finmap.union_assoc, Finmap.union_singleton, SLP.star_assoc] + -- use Finmap.singleton r ⟨tp, v'⟩, ?_ + -- apply And.intro (by assumption) + -- apply And.intro rfl + -- apply And.intro (by simp [SLP.singleton]) + -- apply SLP.ent_star_top + -- assumption -- Misc diff --git a/Lampe/Lampe/Hoare/SepTotal.lean b/Lampe/Lampe/Hoare/SepTotal.lean index 5e1fe64..00e6211 100644 --- a/Lampe/Lampe/Hoare/SepTotal.lean +++ b/Lampe/Lampe/Hoare/SepTotal.lean @@ -1,7 +1,6 @@ import Lampe.Ast import Lampe.Tp import Lampe.Semantics -import Lampe.SeparationLogic import Lampe.Hoare.Total namespace Lampe @@ -22,7 +21,7 @@ An intuitive way of looking at this is thinking in terms of "knowledge discovery For example, if the operation `a + b` succeeds, then we know that it evaluates to `v = a + b` **and** `a + b < 2^32`, i.e., no overflow has happened. Then, we would define the post-condition such that `Q v = (v = a + b) ∧ (a + b < 2^32)`. -/ -def STHoare p Γ P e (Q : Tp.denote p tp → SLP p) +def STHoare p Γ P e (Q : Tp.denote p tp → SLP (State p)) := ∀H, THoare p Γ (P ⋆ H) e (fun v => ((Q v) ⋆ H) ⋆ ⊤) abbrev STHoarePureBuiltin p (Γ : Env) @@ -44,7 +43,7 @@ theorem frame (h_hoare: STHoare p Γ P e Q): STHoare p Γ (P ⋆ H) e (fun v => tauto } -theorem consequence {p tp} {e : Expr (Tp.denote p) tp} {H₁ H₂} {Q₁ Q₂ : Tp.denote p tp → SLP p} +theorem consequence {p tp} {e : Expr (Tp.denote p) tp} {H₁ H₂} {Q₁ Q₂ : Tp.denote p tp → SLP (State p)} (h_pre_conseq : H₂ ⊢ H₁) (h_post_conseq : ∀ v, Q₁ v ⋆ ⊤ ⊢ Q₂ v ⋆ ⊤) (h_hoare: STHoare p Γ H₁ e Q₁): @@ -58,7 +57,7 @@ theorem consequence {p tp} {e : Expr (Tp.denote p) tp} {H₁ H₂} {Q₁ Q₂ : apply SLP.star_mono_l apply_assumption -theorem ramified_frame_top {Q₁ Q₂ : Tp.denote p tp → SLP p} +theorem ramified_frame_top {Q₁ Q₂ : Tp.denote p tp → SLP (State p)} (h_hoare: STHoare p Γ H₁ e Q₁) (h_ent: H₂ ⊢ H₁ ⋆ (∀∀v, Q₁ v -⋆ (Q₂ v ⋆ ⊤))): STHoare p Γ H₂ e Q₂ := by @@ -78,7 +77,7 @@ theorem ramified_frame_top {Q₁ Q₂ : Tp.denote p tp → SLP p} apply SLP.wand_cancel simp [SLP.entails_self] -theorem consequence_frame_left {H H₁ H₂ : SLP p} +theorem consequence_frame_left {H H₁ H₂ : SLP (State p)} (h_hoare: STHoare p Γ H₁ e Q) (h_ent : H ⊢ (H₁ ⋆ H₂)): STHoare p Γ H e (fun v => Q v ⋆ H₂) := by @@ -99,7 +98,7 @@ theorem var_intro {v : Tp.denote p tp}: apply THoare.consequence ?_ THoare.var_intro (fun _ => SLP.entails_self) simp -theorem letIn_intro {tp} {P} {Q : Tp.denote p tp → SLP p} {e₁ e₂} +theorem letIn_intro {tp} {P} {Q : Tp.denote p tp → SLP (State p)} {e₁ e₂} (h_first: STHoare p Γ P e₁ Q) (h_rest: ∀v, STHoare p Γ (Q v) (e₂ v) R): STHoare p Γ P (Expr.letIn e₁ e₂) R := by @@ -119,7 +118,8 @@ lemma Finmap.empty_disjoint: Finmap.Disjoint st ∅ := by rw [Finmap.Disjoint.symm_iff] simp [Finmap.disjoint_empty] -lemma Finmap.union_singleton [DecidableEq α] {β : α → Type u} {r : α} {v v' : β r} : Finmap.singleton r v ∪ Finmap.singleton r v' = Finmap.singleton r v := by +lemma Finmap.union_singleton [DecidableEq α] {β : α → Type u} {r : α} {v v' : β r} : + Finmap.singleton r v ∪ Finmap.singleton r v' = Finmap.singleton r v := by apply Finmap.ext_lookup intro x cases Decidable.em (r = x) @@ -130,8 +130,8 @@ lemma Finmap.union_singleton [DecidableEq α] {β : α → Type u} {r : α} {v v · simp_all [Finmap.lookup_eq_none, eq_comm] simp_all [eq_comm] -theorem fresh_intro - : STHoare p Γ +theorem fresh_intro : + STHoare p Γ ⟦⟧ (.call h![] [] tp (.builtin .fresh) h![]) (fun _ => ⟦⟧) := by @@ -192,7 +192,7 @@ theorem loopNext_intro {lo hi : U s} : apply letIn_intro all_goals tauto -lemma inv_congr (Inv : (i : U s) → (lo ≤ i) → (i ≤ hi) → SLP p) {i j hlo hhi} (hEq : i = j): +lemma inv_congr (Inv : (i : U s) → (lo ≤ i) → (i ≤ hi) → SLP (State p)) {i j hlo hhi} (hEq : i = j): Inv i hlo hhi = Inv j (hEq ▸ hlo) (hEq ▸ hhi) := by cases hEq rfl @@ -214,7 +214,7 @@ lemma U.le_plus_one_of_lt {i j : U s} (h: i < j): i + 1 ≤ j := by linarith ) -theorem loop_inv_intro (Inv : (i : U s) → (lo ≤ i) → (i ≤ hi) → SLP p) {body : U s → Expr (Tp.denote p) tp}: +theorem loop_inv_intro (Inv : (i : U s) → (lo ≤ i) → (i ≤ hi) → SLP (State p)) {body : U s → Expr (Tp.denote p) tp}: (∀i, (hlo: lo ≤ i) → (hhi: i < hi) → STHoare p Γ (Inv i hlo (BitVec.le_of_lt hhi)) (body i) (fun _ => Inv (i + 1) (BitVec.le_trans hlo (U.le_add_one_of_exists_lt hhi)) (U.le_plus_one_of_lt hhi))) → STHoare p Γ (∃∃h, Inv lo BitVec.le_refl h) (.loop lo hi body) (fun _ => ∃∃h, Inv hi h BitVec.le_refl) := by cases BitVec.le_or_lt lo hi with diff --git a/Lampe/Lampe/Hoare/Total.lean b/Lampe/Lampe/Hoare/Total.lean index 5bfb9c9..debd352 100644 --- a/Lampe/Lampe/Hoare/Total.lean +++ b/Lampe/Lampe/Hoare/Total.lean @@ -1,7 +1,7 @@ import Lampe.Ast import Lampe.Tp import Lampe.Semantics -import Lampe.SeparationLogic +import Lampe.SeparationLogic.SLH namespace Lampe @@ -9,18 +9,18 @@ def THoare {tp : Tp} (p : Prime) (Γ : Env) - (P : SLP p) + (P : SLP (State p)) (e : Expr (Tp.denote p) tp) - (Q : (tp.denote p) → SLP p): Prop := - ∀st, P st.vals → Omni p Γ st e (fun r => match r with + (Q : (tp.denote p) → SLP (State p)) : Prop := + ∀st, P st → Omni p Γ st e (fun r => match r with | none => True - | some (st', v) => Q v st'.vals) + | some (st', v) => Q v st') namespace THoare variable {p : Prime} - {H H₁ H₂ : SLP p} + {H H₁ H₂ : SLP (State p)} {Γ : Env} {tp : Tp} {e : Expr (Tp.denote p) tp} @@ -36,7 +36,7 @@ theorem consequence {Q₁ Q₂} any_goals tauto rintro (_|_) <;> tauto -theorem var_intro {v} {P : Tp.denote p tp → SLP p}: +theorem var_intro {v} {P : Tp.denote p tp → SLP (State p)}: THoare p Γ (P v) (.var v) P := by tauto @@ -50,7 +50,7 @@ theorem letIn_intro {P Q} theorem ref_intro {v}: THoare p Γ - (fun st => ∀r, r ∉ st → P r (st.insert r ⟨tp, v⟩)) + (fun st => ∀r, r ∉ st → P r ⟨(st.vals.insert r ⟨tp, v⟩), st.closures⟩) (.call h![] [tp] (Tp.ref tp) (.builtin .ref) h![v]) P := by unfold THoare @@ -61,7 +61,7 @@ theorem ref_intro {v}: theorem readRef_intro {ref}: THoare p Γ - (fun st => st.lookup ref = some ⟨tp, v⟩ ∧ P v st) + (fun st => st.vals.lookup ref = some ⟨tp, v⟩ ∧ P v st) (.call h![] [tp.ref] tp (.builtin .readRef) h![ref]) P := by unfold THoare @@ -73,7 +73,7 @@ theorem readRef_intro {ref}: theorem writeRef_intro {ref v}: THoare p Γ - (fun st => ref ∈ st ∧ P () (st.insert ref ⟨tp, v⟩)) + (fun st => ref ∈ st ∧ P () ⟨(st.vals.insert ref ⟨tp, v⟩), st.closures⟩) (.call h![] [tp.ref, tp] .unit (.builtin .writeRef) h![ref, v]) P := by unfold THoare diff --git a/Lampe/Lampe/Semantics.lean b/Lampe/Lampe/Semantics.lean index 11feccc..a10f366 100644 --- a/Lampe/Lampe/Semantics.lean +++ b/Lampe/Lampe/Semantics.lean @@ -1,9 +1,10 @@ +import Mathlib + import Lampe.Ast import Lampe.Tp import Lampe.Data.Field import Lampe.Tactic.IntroCases -import Mathlib -import Lampe.State +import Lampe.SeparationLogic.State namespace Lampe @@ -29,13 +30,13 @@ inductive Omni : (p : Prime) → Env → State p → Expr (Tp.denote p) tp → ( Omni p Γ st elseBranch Q → Omni p Γ st (Expr.ite false mainBranch elseBranch) Q | letIn : - Omni p Γ ⟨vh₁, cls⟩ e Q₁ → - (∀v vh, Q₁ (some (⟨vh, cls⟩, v)) → Omni p Γ ⟨vh, cls⟩ (b v) Q) → + Omni p Γ st e Q₁ → + (∀v st, Q₁ (some (st, v)) → Omni p Γ st (b v) Q) → (Q₁ none → Q none) → - Omni p Γ ⟨vh₁, cls⟩ (.letIn e b) Q -| callBuiltin {Q} {vh₁ cls} : - (b.omni p vh₁ argTypes resType args (mapToValHeapCondition Q cls)) → - Omni p Γ ⟨vh₁, cls⟩ (Expr.call h![] argTypes resType (.builtin b) args) Q + Omni p Γ st (.letIn e b) Q +| callBuiltin {Q} : + (b.omni p st.vals argTypes resType args (mapToValHeapCondition st.closures Q)) → + Omni p Γ st (Expr.call h![] argTypes resType (.builtin b) args) Q | callDecl : Γ fname = some fn → (hkc : fn.generics = tyKinds) → @@ -44,7 +45,7 @@ inductive Omni : (p : Prime) → Env → State p → Expr (Tp.denote p) tp → ( Omni p Γ st (htco ▸ fn.body _ (hkc ▸ generics) (htci ▸ args)) Q → Omni p Γ st (@Expr.call _ tyKinds generics argTypes res (.decl fname) args) Q | callLambda {cls : Closures} : - cls.get? lambdaRef.val = some fn → + cls.lookup lambdaRef = some fn → (hg : fn.generics = []) → (hi : fn.inTps (hg ▸ h![]) = argTps) → (ho : fn.outTp (hg ▸ h![]) = outTp) → @@ -70,22 +71,22 @@ theorem Omni.consequence {p Γ st tp} {e : Expr (Tp.denote p) tp} {Q Q'}: ) case callBuiltin => cases_type Builtin + intros + constructor tauto case loopNext => intro apply loopNext (by assumption) tauto -theorem Omni.frame {p Γ tp} {vh₁ vh₂ : ValHeap p} {cls : Closures} {e : Expr (Tp.denote p) tp} {Q}: - Omni p Γ ⟨vh₁, cls⟩ e Q → - vh₁.Disjoint vh₂ → - Omni p Γ ⟨(vh₁ ∪ vh₂), cls⟩ e (fun stv => match stv with - | some (st, v) => ((fun vals => Q (some (⟨vals, cls⟩, v))) ⋆ (fun vals => vals = vh₂)) st.vals +theorem Omni.frame {p Γ tp} {st₁ st₂ : State p} {cls : Closures} {e : Expr (Tp.denote p) tp} {Q}: + Omni p Γ st₁ e Q → + SLH.disjoint st₁ st₂ → + Omni p Γ (st₁ ∪ st₂) e (fun stv => match stv with + | some (st', v) => ((fun st => Q (some (st, v))) ⋆ (fun st => st = st₂)) st' | none => Q none ) := by intro h - generalize hc : State.mk vh₁ cls = st₁ - rw [hc] at h induction h with | litField hq | skip hq @@ -95,30 +96,33 @@ theorem Omni.frame {p Γ tp} {vh₁ vh₂ : ValHeap p} {cls : Closures} {e : Ex | litRef hq | var hq => intro - subst hc constructor - exists vh₁, vh₂ + simp only + repeat apply Exists.intro + tauto | letIn _ _ hN ihE ihB => intro constructor apply ihE assumption - assumption · intro _ _ h cases h casesm* ∃ _, _, _∧_ - simp_all only + subst_vars apply ihB - simp_all only [State.mk.injEq] - tauto - . aesop - . assumption + assumption + assumption · simp_all | callBuiltin hq => - cases_type Builtin - intro + rename Builtin => b + intros constructor - aesop + have uni_vals {p' : Prime} {a b : State p'} : (a ∪ b).vals = (a.vals ∪ b.vals) := by rfl + have uni_cls {p' : Prime} {a b : State p'} : (a ∪ b).closures = (a.closures ∪ b.closures) := by rfl + rw [uni_vals, uni_cls] + unfold mapToValHeapCondition + simp_all + sorry | callDecl _ _ _ _ _ ih => intro constructor @@ -137,7 +141,7 @@ theorem Omni.frame {p Γ tp} {vh₁ vh₂ : ValHeap p} {cls : Closures} {e : Ex intro constructor apply ih - <;> tauto + tauto | callLambda => intro constructor <;> try tauto diff --git a/Lampe/Lampe/SeparationLogic.lean b/Lampe/Lampe/SeparationLogic.lean deleted file mode 100644 index d9d7444..0000000 --- a/Lampe/Lampe/SeparationLogic.lean +++ /dev/null @@ -1,269 +0,0 @@ -import Lampe.ValHeap -import Lampe.Tactic.IntroCases - - -namespace Lampe - -def SLP (p : Prime) : Type := ValHeap p → Prop - -namespace SLP - -def star {p} (lhs rhs : SLP p) : SLP p := fun st => - ∃ st₁ st₂, Finmap.Disjoint st₁ st₂ ∧ st = st₁ ∪ st₂ ∧ lhs st₁ ∧ rhs st₂ - -def lift {p} (pr : Prop) : SLP p := fun st => pr ∧ st = ∅ - -def singleton {p} (r : Ref) (v : AnyValue p) : SLP p := fun st => st = Finmap.singleton r v - -def wand {p} (lhs rhs : SLP p) : SLP p := - fun st => ∀st', st.Disjoint st' → lhs st' → rhs (st ∪ st') - -def top {p} : SLP p := fun _ => True - -def entails {p} (a b : SLP p) := ∀st, a st → b st - -def forall' {p} (f : α → SLP p) : SLP p := fun st => ∀v, f v st -def exists' {p} (f : α → SLP p) : SLP p := fun st => ∃v, f v st - -instance {p}: Coe Prop (SLP p) := ⟨lift⟩ - -notation:max "⊤" => top - -notation:max "⟦" x "⟧" => lift x - -notation:max "⟦⟧" => ⟦True⟧ - -infixr:35 " ⋆ " => star - -infixr:30 " -⋆ " => wand - -infix:10 " ⊢ " => entails - -notation:max "[" l " ↦ " r "]" => singleton l r - -open Lean.TSyntax.Compat in -macro "∃∃" xs:Lean.explicitBinders ", " b:term : term => Lean.expandExplicitBinders ``exists' xs b - -open Lean.TSyntax.Compat in -macro "∀∀" xs:Lean.explicitBinders ", " b:term : term => Lean.expandExplicitBinders ``forall' xs b - -theorem entails_trans {p} {P Q R : SLP p}: (P ⊢ Q) → (Q ⊢ R) → (P ⊢ R) := by tauto - -section basic - -variable {p : Prime} - -@[simp] -theorem apply_top : ⊤ st := by trivial - -theorem forall_left {a} {P : α → SLP p} : (P a ⊢ Q) → ((∀∀a, P a) ⊢ Q) := by - unfold forall' - tauto - -theorem forall_right {H' : α → SLP p}: (∀x, H ⊢ H' x) → (H ⊢ ∀∀x, H' x) := by - unfold forall' entails - tauto - -theorem pure_left: (P → (H ⊢ H')) → (P ⋆ H ⊢ H') := by - unfold star entails lift - intro_cases - simp_all - -theorem pure_left' {P} {H : SLP p} : (P → (⟦⟧ ⊢ H)) → (P ⊢ H) := by - unfold entails lift - tauto - -theorem pure_right: P → (H₁ ⊢ H₂) → (H₁ ⊢ P ⋆ H₂) := by - unfold star entails lift - intros - repeat apply Exists.intro - simp_all - apply And.intro ?_ - apply And.intro ?_ - apply And.intro rfl - apply_assumption - assumption - simp - simp [Finmap.disjoint_empty] - -theorem entails_self : H ⊢ H := by tauto - -theorem entails_top : H ⊢ ⊤ := by tauto - -@[simp] -theorem forall_unused {α : Type u} [Inhabited α] {P : SLP p} : (∀∀(_:α), P) = P := by - funext - unfold forall' - rw [eq_iff_iff] - apply Iff.intro - · intro - apply_assumption - apply Inhabited.default - · intros - apply_assumption - -end basic - -section star - -variable {p : Prime} - -theorem star_comm {G H:SLP p} : (G ⋆ H) = (H ⋆ G) := by - funext - unfold star - rw [eq_iff_iff] - apply Iff.intro <;> { - intro_cases - repeat apply Exists.intro - rw [Finmap.Disjoint.symm_iff] - apply And.intro (by assumption) - rw [Finmap.union_comm_of_disjoint (by rw [Finmap.Disjoint.symm_iff]; assumption)] - tauto - } - -@[simp] -theorem true_star {H:SLP p} : (⟦⟧ ⋆ H) = H := by - funext - rw [eq_iff_iff] - unfold lift star - apply Iff.intro - · simp_all - · intro - exists ∅, ?_ - simp_all [Finmap.disjoint_empty] - -@[simp] -theorem star_true {H:SLP p} : (H ⋆ ⟦⟧) = H := by rw [star_comm]; simp - -@[simp] -theorem star_assoc {F G H:SLP p} : ((F ⋆ G) ⋆ H) = (F ⋆ G ⋆ H) := by - funext - rw [eq_iff_iff] - unfold star - apply Iff.intro - · intro_cases - subst_vars - rw [Finmap.union_assoc] - simp only [Finmap.disjoint_union_left] at * - cases_type And - repeat apply Exists.intro - apply And.intro ?_ - apply And.intro rfl - apply And.intro (by assumption) - repeat apply Exists.intro - apply And.intro ?_ - apply And.intro rfl - simp_all - assumption - simp_all [Finmap.disjoint_union_right] - · intro_cases - subst_vars - rw [←Finmap.union_assoc] - simp only [Finmap.disjoint_union_right] at * - cases_type And - repeat apply Exists.intro - apply And.intro ?_ - apply And.intro rfl - apply And.intro ?_ (by assumption) - repeat apply Exists.intro - apply And.intro ?_ - apply And.intro rfl - simp_all - assumption - simp_all [Finmap.disjoint_union_left] - -@[simp] -theorem ent_star_top {H : SLP p} : H ⊢ H ⋆ ⊤ := by - intro _ _ - exists ?_, ∅ - rw [Finmap.Disjoint.symm_iff] - simp_all [Finmap.disjoint_empty] - -theorem star_mono_r : (P ⊢ Q) → (P ⋆ R ⊢ Q ⋆ R) := by - unfold star entails - tauto - -theorem star_mono_l : (P ⊢ Q) → (R ⋆ P ⊢ R ⋆ Q) := by - unfold star entails - tauto - -theorem star_mono_l' : (⟦⟧ ⊢ Q) → (P ⊢ P ⋆ Q) := by - unfold star entails lift - intros - simp_all - repeat apply Exists.intro - apply And.intro ?_ - apply And.intro ?_ - tauto - simp - tauto - -theorem star_mono : (H₁ ⊢ H₂) → (Q₁ ⊢ Q₂) → (H₁ ⋆ Q₁ ⊢ H₂ ⋆ Q₂) := by - unfold star entails - tauto - -theorem forall_star {P : α → SLP p} : (∀∀x, P x) ⋆ Q ⊢ ∀∀x, P x ⋆ Q := by - unfold star forall' - tauto - -theorem star_forall {P : α → SLP p} : Q ⋆ (∀∀x, P x) ⊢ ∀∀x, Q ⋆ P x := by - unfold star forall' - tauto - -@[simp] -theorem top_star_top : (top (p := p) ⋆ ⊤) = ⊤ := by - unfold top star - funext x - simp - exists ∅, x - simp [Finmap.disjoint_empty] - -end star - -section wand - -variable {p : Prime} - -@[simp] -theorem wand_self_star {H:SLP p}: (H -⋆ H ⋆ top) = top := by - funext - unfold wand star - apply eq_iff_iff.mpr - apply Iff.intro - · intro - simp [lift] - · intros - repeat apply Exists.intro - apply And.intro ?_ - apply And.intro ?_ - apply And.intro (by assumption) - simp - rotate_left - rotate_left - rw [Finmap.union_comm_of_disjoint (by assumption)] - rw [Finmap.Disjoint.symm_iff] - assumption - - -theorem wand_intro {A B C : SLP p} : (A ⋆ B ⊢ C) → (A ⊢ B -⋆ C) := by - unfold wand star entails - intros - intros - apply_assumption - tauto - -theorem wand_cancel : (P ⋆ (P -⋆ Q)) ⊢ Q := by - unfold star wand entails - intro_cases - subst_vars - rename_i h - rw [Finmap.union_comm_of_disjoint (by assumption)] - apply_assumption - tauto - tauto - -end wand - -end SLP - -end Lampe diff --git a/Lampe/Lampe/SeparationLogic/SLH.lean b/Lampe/Lampe/SeparationLogic/SLH.lean new file mode 100644 index 0000000..66d99a1 --- /dev/null +++ b/Lampe/Lampe/SeparationLogic/SLH.lean @@ -0,0 +1,58 @@ +import Lampe.Tp + +namespace Lampe + +class SLH (α : Type _) where + union : α → α → α + disjoint : α → α → Prop + empty : α + + union_empty {a : α} : union a empty = a + union_assoc {s₁ s₂ s₃ : α} : union (union s₁ s₂) s₃ = union s₁ (union s₂ s₃) + disjoint_symm_iff {a b : α} : disjoint a b ↔ disjoint b a + union_comm_of_disjoint {s₁ s₂ : α} : disjoint s₁ s₂ → union s₁ s₂ = union s₂ s₁ + disjoint_union_left (x y z : α) : disjoint (union x y) z ↔ disjoint x z ∧ disjoint y z + disjoint_union_right (x y z : α) : disjoint x (union y z) ↔ disjoint x y ∧ disjoint x z + disjoint_empty (x : α) : disjoint empty x + +instance [SLH α] : Union α := ⟨SLH.union⟩ +instance [SLH α] : EmptyCollection α := ⟨SLH.empty⟩ + +theorem SLH_union_empty_commutative [SLH α] {a : α} : + a ∪ ∅ = ∅ ∪ a := by + have h := SLH.disjoint_empty a + simp [Union.union, EmptyCollection.emptyCollection] + rw [SLH.union_comm_of_disjoint h] + +theorem SLH_union_comm_of_disjoint [SLH α] {s₁ s₂ : α}: + SLH.disjoint s₁ s₂ → s₁ ∪ s₂ = s₂ ∪ s₁ := by + simp [Union.union] + exact SLH.union_comm_of_disjoint + +@[simp] +theorem SLH_union_empty [SLH α] {a : α} : + a ∪ ∅ = a := by + apply SLH.union_empty + +@[simp] +theorem SLH_empty_union [SLH α] {a : α} : + ∅ ∪ a = a := by + rw [←SLH_union_empty_commutative] + apply SLH.union_empty + +theorem SLH_union_assoc [SLH α] {s₁ s₂ s₃ : α} : + s₁ ∪ s₂ ∪ s₃ = s₁ ∪ (s₂ ∪ s₃) := by + simp [Union.union] + apply SLH.union_assoc + +theorem SLH_disjoint_union_left [SLH α] (x y z : α) : + SLH.disjoint (x ∪ y) z ↔ SLH.disjoint x z ∧ SLH.disjoint y z := by + simp [Union.union] + apply SLH.disjoint_union_left + +theorem SLH_disjoint_union_right [SLH α] (x y z : α) : + SLH.disjoint x (y ∪ z) ↔ SLH.disjoint x y ∧ SLH.disjoint x z := by + simp [Union.union] + apply SLH.disjoint_union_right + +end Lampe diff --git a/Lampe/Lampe/SeparationLogic/SLP.lean b/Lampe/Lampe/SeparationLogic/SLP.lean new file mode 100644 index 0000000..4e5f5ad --- /dev/null +++ b/Lampe/Lampe/SeparationLogic/SLP.lean @@ -0,0 +1,263 @@ +import Lampe.Tactic.IntroCases +import Lampe.SeparationLogic.SLH + +namespace Lampe + +def SLP (α) [SLH α] := α → Prop + +namespace SLP + +def star [SLH α] (lhs rhs : SLP α) := fun st => + ∃ st₁ st₂, SLH.disjoint st₁ st₂ ∧ st = st₁ ∪ st₂ ∧ lhs st₁ ∧ rhs st₂ + +def lift [SLH α] (pr : Prop) : SLP α := fun st => pr ∧ st = ∅ + +def wand [SLH α] (lhs rhs : SLP α) : SLP α := + fun st => ∀st', SLH.disjoint st st' → lhs st' → rhs (st ∪ st') + +def top [SLH α] : SLP α := fun _ => True + +def entails [SLH α] (a b : SLP α) := ∀st, a st → b st + +def forall' [SLH α] (f : β → SLP α) : SLP α := fun st => ∀v, f v st +def exists' [SLH α] (f : β → SLP α) : SLP α := fun st => ∃v, f v st + +instance [SLH α]: Coe Prop (SLP α) := ⟨lift⟩ + +notation:max "⊤" => top + +notation:max "⟦" x "⟧" => lift x + +notation:max "⟦⟧" => ⟦True⟧ + +infixr:35 " ⋆ " => star + +infixr:30 " -⋆ " => wand + +infix:10 " ⊢ " => entails + +open Lean.TSyntax.Compat in +macro "∃∃" xs:Lean.explicitBinders ", " b:term : term => Lean.expandExplicitBinders ``exists' xs b + +open Lean.TSyntax.Compat in +macro "∀∀" xs:Lean.explicitBinders ", " b:term : term => Lean.expandExplicitBinders ``forall' xs b + +theorem entails_trans [SLH α] {P Q R : SLP α}: (P ⊢ Q) → (Q ⊢ R) → (P ⊢ R) := by tauto + +section basic + +@[simp] +theorem apply_top [SLH α] {st : α} : ⊤ st := by trivial + +theorem forall_left [SLH β] {P : α → SLP β} : (P a ⊢ Q) → ((∀∀(a : α), P a) ⊢ Q) := by + unfold forall' + tauto + +theorem forall_right [SLH β] {H : SLP β} {H' : α → SLP β}: (∀x, H ⊢ H' x) → (H ⊢ ∀∀x, H' x) := by + unfold forall' entails + tauto + +theorem pure_left [SLH β] {H H' : SLP β} : (P → (H ⊢ H')) → (P ⋆ H ⊢ H') := by + unfold star entails lift + intro_cases + simp_all + +theorem pure_left' [SLH α] {H : SLP α} : (P → (⟦⟧ ⊢ H)) → (P ⊢ H) := by + unfold entails lift + tauto + +theorem pure_right [SLH α] {H₁ H₂ : SLP α} : P → (H₁ ⊢ H₂) → (H₁ ⊢ P ⋆ H₂) := by + unfold star entails lift + intros + repeat apply Exists.intro + simp_all + apply And.intro ?_ + apply And.intro ?_ + apply And.intro rfl + apply_assumption + assumption + . simp only [SLH_empty_union] + . apply SLH.disjoint_empty + +theorem entails_self [SLH α] {H : SLP α} : H ⊢ H := by tauto + +theorem entails_top [SLH α] {H : SLP α} : H ⊢ ⊤ := by tauto + +@[simp] +theorem forall_unused [SLH β] {α : Type u} [Inhabited α] {P : SLP β} : (∀∀ (_ : α), P) = P := by + funext + unfold forall' + rw [eq_iff_iff] + apply Iff.intro + · intro + apply_assumption + apply Inhabited.default + · intros + apply_assumption + +end basic + +section star + +theorem star_comm [SLH α] {G H : SLP α} : (G ⋆ H) = (H ⋆ G) := by + funext + unfold star + rw [eq_iff_iff] + apply Iff.intro <;> { + intro_cases + repeat apply Exists.intro + rw [SLH.disjoint_symm_iff] + apply And.intro (by assumption) + rw [SLH_union_comm_of_disjoint (by rw [SLH.disjoint_symm_iff]; assumption)] + tauto + } + +@[simp] +theorem true_star [SLH α] {H : SLP α} : (⟦⟧ ⋆ H) = H := by + funext + rw [eq_iff_iff] + unfold lift star + apply Iff.intro + · simp_all + · intro + exists ∅, ?_ + simp_all [SLH.disjoint_empty] + apply SLH.disjoint_empty + +@[simp] +theorem star_true [SLH α] {H : SLP α} : (H ⋆ ⟦⟧) = H := by rw [star_comm]; simp + +@[simp] +theorem star_assoc [SLH α] {F G H : SLP α} : ((F ⋆ G) ⋆ H) = (F ⋆ G ⋆ H) := by + funext + rw [eq_iff_iff] + unfold star + apply Iff.intro + · intro_cases + subst_vars + rw [SLH_union_assoc] + simp only [SLH_disjoint_union_left] at * + cases_type And + repeat apply Exists.intro + apply And.intro ?_ + apply And.intro rfl + apply And.intro (by assumption) + repeat apply Exists.intro + apply And.intro ?_ + apply And.intro rfl + simp_all + assumption + simp_all [SLH_disjoint_union_right] + · intro_cases + subst_vars + rw [←SLH_union_assoc] + simp only [SLH_disjoint_union_right] at * + cases_type And + repeat apply Exists.intro + apply And.intro ?_ + apply And.intro rfl + apply And.intro ?_ (by assumption) + repeat apply Exists.intro + apply And.intro ?_ + apply And.intro rfl + simp_all + assumption + simp_all [SLH_disjoint_union_left] + +@[simp] +theorem ent_star_top [SLH α] {H : SLP α} : H ⊢ H ⋆ ⊤ := by + intro _ _ + exists ?_, ∅ + rw [SLH.disjoint_symm_iff] + simp_all [SLH.disjoint_empty] + apply SLH.disjoint_empty + +theorem star_mono_r [SLH α] {P Q R : SLP α} : (P ⊢ Q) → (P ⋆ R ⊢ Q ⋆ R) := by + unfold star entails + tauto + +theorem star_mono_l [SLH α] {P Q R : SLP α} : (P ⊢ Q) → (R ⋆ P ⊢ R ⋆ Q) := by + unfold star entails + tauto + +theorem star_mono_l' [SLH α] {P Q : SLP α} : (⟦⟧ ⊢ Q) → (P ⊢ P ⋆ Q) := by + unfold star entails lift + intros + simp_all + repeat apply Exists.intro + apply And.intro ?_ + apply And.intro ?_ + tauto + simp + rw [SLH.disjoint_symm_iff] + apply SLH.disjoint_empty + +theorem star_mono [SLH α] {H₁ H₂ Q₁ Q₂ : SLP α} : (H₁ ⊢ H₂) → (Q₁ ⊢ Q₂) → (H₁ ⋆ Q₁ ⊢ H₂ ⋆ Q₂) := by + unfold star entails + tauto + +theorem forall_star [SLH α] {P : α → SLP α} : (∀∀x, P x) ⋆ Q ⊢ ∀∀x, P x ⋆ Q := by + unfold star forall' + tauto + +theorem star_forall [SLH β] {P : α → SLP β} {Q : SLP β} : Q ⋆ (∀∀x, P x) ⊢ ∀∀x, Q ⋆ P x := by + unfold star forall' + tauto + +@[simp] +theorem top_star_top [SLH α] : (top ⋆ (⊤ : SLP α)) = (⊤ : SLP α) := by + unfold top star + funext x + simp + exists ∅, x + simp [SLH.disjoint_empty] + apply SLH.disjoint_empty + +end star + +section wand + +variable {p : Prime} + +@[simp] +theorem wand_self_star [SLH α] {H : SLP α}: (H -⋆ H ⋆ top) = top := by + funext + unfold wand star + apply eq_iff_iff.mpr + apply Iff.intro + · intro + simp [lift] + · intros + repeat apply Exists.intro + apply And.intro ?_ + apply And.intro ?_ + apply And.intro (by assumption) + simp + rotate_left + rotate_left + rw [SLH_union_comm_of_disjoint (by assumption)] + rw [SLH.disjoint_symm_iff] + assumption + + +theorem wand_intro [SLH α] {A B C : SLP α} : (A ⋆ B ⊢ C) → (A ⊢ B -⋆ C) := by + unfold wand star entails + intros + intros + apply_assumption + tauto + +theorem wand_cancel [SLH α] {P Q : SLP α} : (P ⋆ (P -⋆ Q)) ⊢ Q := by + unfold star wand entails + intro_cases + subst_vars + rename_i h + rw [SLH_union_comm_of_disjoint (by assumption)] + apply_assumption + rw [SLH.disjoint_symm_iff] + tauto + tauto + +end wand + +end Lampe.SLP diff --git a/Lampe/Lampe/SeparationLogic/State.lean b/Lampe/Lampe/SeparationLogic/State.lean new file mode 100644 index 0000000..3171c7a --- /dev/null +++ b/Lampe/Lampe/SeparationLogic/State.lean @@ -0,0 +1,103 @@ +import Lampe.SeparationLogic.SLH +import Lampe.SeparationLogic.ValHeap +import Lampe.Ast + +namespace Lampe + +abbrev Closures := Finmap fun _ : Ref => Lampe.Function + +structure State (p : Prime) where + vals : ValHeap p + closures : Closures + +instance : Membership Ref (State p) where + mem := fun a e => e ∈ a.vals + +instance : Coe (State p) (ValHeap p) where + coe := fun s => s.vals + +/-- Maps a condition on `State`s to a condition on `ValHeap`s by keeping the closures fixed -/ +@[reducible] +def mapToValHeapCondition + (closures : Finmap fun _ : Ref => Function) + (Q : Option (State p × T) → Prop) : Option (ValHeap p × T) → Prop := + fun v => Q (v.map (fun (vals, t) => ⟨⟨vals, closures⟩, t⟩)) + +def State.insertVal (self : State p) (r : Ref) (v : AnyValue p) : State p := + ⟨self.vals.insert r v, self.closures⟩ + +theorem State.eq_parts : + v = v' → c = c' → State.mk v c = State.mk v' c' := by + intros + subst_vars + rfl + +instance : SLH (State p) where + union := fun a b => ⟨a.vals ∪ b.vals, a.closures ∪ b.closures⟩ + disjoint := fun a b => a.vals.Disjoint b.vals ∧ a.closures.Disjoint b.closures + empty := ⟨∅, ∅⟩ + union_empty := by + intros + simp only [Finmap.union_empty] + union_assoc := by + intros + simp only [Union.union] + apply State.eq_parts + . apply Finmap.union_assoc + . apply Finmap.union_assoc + disjoint_symm_iff := by tauto + union_comm_of_disjoint := by + intros + simp only [Union.union] + apply State.eq_parts + . apply Finmap.union_comm_of_disjoint + tauto + . apply Finmap.union_comm_of_disjoint + tauto + disjoint_empty := by tauto + disjoint_union_left := by + intros x y z + have h1 := (Finmap.disjoint_union_left x.vals y.vals z.vals) + have h2 := (Finmap.disjoint_union_left x.closures y.closures z.closures) + constructor + intro h3 + simp only [Union.union] at h3 + constructor + constructor + cases h3 + tauto + tauto + tauto + intro h3 + simp only [Union.union] at h3 + constructor + tauto + tauto + disjoint_union_right := by + intros x y z + have h1 := (Finmap.disjoint_union_right x.vals y.vals z.vals) + have h2 := (Finmap.disjoint_union_right x.closures y.closures z.closures) + constructor + intro h3 + simp only [Union.union] at h3 + constructor + constructor + cases h3 + tauto + tauto + tauto + intro h3 + simp only [Union.union] at h3 + constructor + tauto + tauto + +def State.singleton (r : Ref) (v : AnyValue p) : SLP (State p) := fun st => st.vals = Finmap.singleton r v + +notation:max "[" l " ↦ " r "]" => State.singleton l r + +theorem State.union_parts : + (State.mk v c ∪ st₂ = State.mk (v ∪ st₂.vals) (c ∪ st₂.closures)) := by + aesop + +end Lampe diff --git a/Lampe/Lampe/SeparationLogic/ValHeap.lean b/Lampe/Lampe/SeparationLogic/ValHeap.lean new file mode 100644 index 0000000..5497621 --- /dev/null +++ b/Lampe/Lampe/SeparationLogic/ValHeap.lean @@ -0,0 +1,34 @@ +import Lampe.SeparationLogic.SLH +import Lampe.SeparationLogic.SLP + +lemma Finmap.insert_eq_singleton_union [DecidableEq α] {ref : α}: + m.insert ref v = Finmap.singleton ref v ∪ m := by rfl + +@[simp] +lemma Finmap.singleton_disjoint_of_not_mem (hp : ref ∉ s): + Finmap.Disjoint (Finmap.singleton ref v) s := by + simp_all [Finmap.Disjoint] + +namespace Lampe + +def AnyValue (p : Prime) := (tp : Tp) × tp.denote p + +abbrev ValHeap (p : Prime) := Finmap (fun (_ : Ref) => AnyValue p) + +instance : SLH (ValHeap p) where + union := fun a b => a ∪ b + disjoint := fun a b => a.Disjoint b + empty := ∅ + union_empty := Finmap.union_empty + union_assoc := Finmap.union_assoc + disjoint_symm_iff := by tauto + union_comm_of_disjoint := Finmap.union_comm_of_disjoint + disjoint_union_left := Finmap.disjoint_union_left + disjoint_union_right := Finmap.disjoint_union_right + disjoint_empty := Finmap.disjoint_empty + +def ValHeap.singleton (r : Ref) (v : AnyValue p) : SLP (ValHeap p) := fun st => st = Finmap.singleton r v + +notation:max "[" l " ↦ " r "]" => ValHeap.singleton l r + +end Lampe diff --git a/Lampe/Lampe/State.lean b/Lampe/Lampe/State.lean deleted file mode 100644 index df41b1d..0000000 --- a/Lampe/Lampe/State.lean +++ /dev/null @@ -1,19 +0,0 @@ -import Mathlib -import Lampe.ValHeap -import Lampe.Ast - -namespace Lampe - -abbrev Closures := List Function - -structure State (p : Prime) where - vals : ValHeap p - cls : Closures - -@[reducible] -def mapToValHeapCondition - (Q : Option (State p × T) → Prop) - (closures : Closures) : Option (ValHeap p × T) → Prop := - fun vv => Q (vv.map fun (vals, v) => ⟨⟨vals, closures⟩, v⟩) - -end Lampe diff --git a/Lampe/Lampe/Tactic/SeparationLogic.lean b/Lampe/Lampe/Tactic/SeparationLogic.lean index 996ebfe..5b20564 100644 --- a/Lampe/Lampe/Tactic/SeparationLogic.lean +++ b/Lampe/Lampe/Tactic/SeparationLogic.lean @@ -1,4 +1,4 @@ -import Lampe.SeparationLogic +import Lampe.SeparationLogic.SLP import Lampe.Hoare.SepTotal import Lampe.Hoare.Builtins import Lampe.Syntax @@ -43,13 +43,13 @@ instance : ToString SLTerm := ⟨SLTerm.toString⟩ instance : Inhabited SLTerm := ⟨SLTerm.top⟩ -theorem star_exists {Q : α → SLP p} : ((∃∃x, Q x) ⋆ P) = (∃∃x, Q x ⋆ P) := by +theorem star_exists [SLH β] {Q : α → SLP β} : ((∃∃x, Q x) ⋆ P) = (∃∃x, Q x ⋆ P) := by unfold SLP.exists' SLP.star funext st simp tauto -theorem exists_star {Q : α → SLP p} : ((∃∃x, Q x) ⋆ P) = (∃∃x, P ⋆ Q x) := by +theorem exists_star [SLH β] {Q : α → SLP β} : ((∃∃x, Q x) ⋆ P) = (∃∃x, P ⋆ Q x) := by rw [star_exists] simp [SLP.star_comm] @@ -74,7 +74,7 @@ theorem ref_intro' {p} {x : Tp.denote p tp} {Γ P}: apply SLP.ent_star_top -theorem Lampe.SLP.skip_fst : (R₁ ⊢ Q ⋆ X) → ([a ↦ b] ⋆ X ⊢ R₂) → ([a ↦ b] ⋆ R₁ ⊢ Q ⋆ R₂) := by +theorem Lampe.SLP.skip_fst [SLH β] : (R₁ ⊢ Q ⋆ X) → ([a ↦ b] ⋆ X ⊢ R₂) → ([a ↦ b] ⋆ R₁ ⊢ Q ⋆ R₂) := by intro h₁ h₂ apply entails_trans rotate_left @@ -95,10 +95,10 @@ theorem Lampe.SLP.skip_fst' : (⟦⟧ ⊢ Q ⋆ X) → ([a ↦ b] ⋆ X ⊢ R₂ assumption assumption -theorem Lampe.SLP.entails_star_true : H ⊢ H ⋆ ⟦⟧ := by +theorem Lampe.SLP.entails_star_true [SLH β] {H : SLP β} : H ⊢ H ⋆ ⟦⟧ := by simp [SLP.entails_self] -theorem SLP.eq_of_iff : (P ⊢ Q) → (Q ⊢ P) → P = Q := by +theorem SLP.eq_of_iff [SLH β] {P Q : SLP β} : (P ⊢ Q) → (Q ⊢ P) → P = Q := by intros apply funext intro @@ -129,7 +129,7 @@ macro "h_norm" : tactic => `(tactic|( subst_vars; )) -theorem SLP.pure_leftX : (P → (H ⊢ Q ⋆ R)) → (P ⋆ H ⊢ Q ⋆ P ⋆ R) := by +theorem SLP.pure_leftX [SLH β] (H Q R : SLP β) : (P → (H ⊢ Q ⋆ R)) → (P ⋆ H ⊢ Q ⋆ P ⋆ R) := by intro apply SLP.pure_left intro @@ -141,19 +141,20 @@ theorem SLP.pure_leftX : (P → (H ⊢ Q ⋆ R)) → (P ⋆ H ⊢ Q ⋆ P ⋆ R) tauto /-- only finisher, will waste mvars for top! -/ -theorem SLP.pure_ent_star_top : (P → Q) → ((P : SLP p) ⊢ Q ⋆ ⊤) := by +theorem SLP.pure_ent_star_top [SLH β] : (P → Q) → ((P : SLP β) ⊢ Q ⋆ ⊤) := by intro h st hp rcases hp with ⟨_, rfl, hp⟩ use ∅, ∅ - simp_all [Finmap.disjoint_empty, SLP.lift] + simp_all [SLH.disjoint_empty, SLP.lift] + apply SLH.disjoint_empty -theorem star_mono_l_sing : (P ⊢ Q) → (v₁ = v₂) → ([r ↦ v₁] ⋆ P ⊢ [r ↦ v₂] ⋆ Q) := by +theorem star_mono_l_sing [SLH β] {P Q : SLP β} : (P ⊢ Q) → (v₁ = v₂) → ([r ↦ v₁] ⋆ P ⊢ [r ↦ v₂] ⋆ Q) := by intro h₁ h₂ rw [h₂] apply SLP.star_mono_l apply h₁ -theorem star_mono_l_sing' : (⟦⟧ ⊢ Q) → (v₁ = v₂) → ([r ↦ v₁] ⊢ [r ↦ v₂] ⋆ Q) := by +theorem star_mono_l_sing' [SLH β] {Q : SLP β} : (⟦⟧ ⊢ Q) → (v₁ = v₂) → ([r ↦ v₁] ⊢ [r ↦ v₂] ⋆ Q) := by intro h₁ h₂ rw [h₂] apply SLP.star_mono_l' @@ -216,25 +217,25 @@ partial def parseEntailment (e: Expr): TacticM (SLTerm × SLTerm) := do return (pre, post) else throwError "not an entailment" -theorem star_top_of_star_mvar : (H ⊢ Q ⋆ R) → (H ⊢ Q ⋆ ⊤) := by +theorem star_top_of_star_mvar [SLH β] {H Q R : SLP β} : (H ⊢ Q ⋆ R) → (H ⊢ Q ⋆ ⊤) := by intro h apply SLP.entails_trans assumption apply SLP.star_mono_l apply SLP.entails_top -theorem solve_left_with_leftovers : (H ⊢ Q ⋆ R) → (R ⊢ P) → (H ⊢ Q ⋆ P) := by +theorem solve_left_with_leftovers [SLH β] {H Q R : SLP β} : (H ⊢ Q ⋆ R) → (R ⊢ P) → (H ⊢ Q ⋆ P) := by intros apply SLP.entails_trans assumption apply SLP.star_mono_l assumption -theorem solve_with_true : (H ⊢ Q) → (H ⊢ Q ⋆ ⟦⟧) := by +theorem solve_with_true [SLH β] {H Q : SLP β} : (H ⊢ Q) → (H ⊢ Q ⋆ ⟦⟧) := by aesop -- partial def solveNonMVarEntailment (goal : MVarId) (lhs : SLTerm) (rhs : SLTerm): TacticM (List MVarId × SLTerm) := do -theorem pure_ent_pure_star_mv: (P → Q) → ((P : SLP p) ⊢ Q ⋆ ⟦⟧) := by +theorem pure_ent_pure_star_mv [SLH β] : (P → Q) → ((P : SLP β) ⊢ Q ⋆ ⟦⟧) := by intro h apply SLP.pure_left' intro @@ -242,7 +243,7 @@ theorem pure_ent_pure_star_mv: (P → Q) → ((P : SLP p) ⊢ Q ⋆ ⟦⟧) := b tauto tauto -theorem pure_star_H_ent_pure_star_mv: (P → (H ⊢ Q ⋆ R)) → (P ⋆ H ⊢ Q ⋆ P ⋆ R) := by +theorem pure_star_H_ent_pure_star_mv [SLH β] {H Q : SLP β} : (P → (H ⊢ Q ⋆ R)) → (P ⋆ H ⊢ Q ⋆ P ⋆ R) := by intro apply SLP.pure_left intro @@ -252,7 +253,7 @@ theorem pure_star_H_ent_pure_star_mv: (P → (H ⊢ Q ⋆ R)) → (P ⋆ H ⊢ Q rw [SLP.star_comm] tauto -theorem skip_left_ent_star_mv : (R ⊢ P ⋆ H) → (L ⋆ R ⊢ P ⋆ L ⋆ H) := by +theorem skip_left_ent_star_mv [SLH β] {R P H L : SLP β} : (R ⊢ P ⋆ H) → (L ⋆ R ⊢ P ⋆ L ⋆ H) := by intro h apply SLP.entails_trans apply SLP.star_mono_l @@ -262,19 +263,19 @@ theorem skip_left_ent_star_mv : (R ⊢ P ⋆ H) → (L ⋆ R ⊢ P ⋆ L ⋆ H) rw [SLP.star_comm] apply SLP.entails_self -theorem skip_evidence_pure : Q → (H ⊢ Q ⋆ H) := by +theorem skip_evidence_pure [SLH β] {H : SLP β} : Q → (H ⊢ Q ⋆ H) := by intro apply SLP.pure_right tauto tauto -theorem SLP.exists_intro { Q : α → SLP p } {a} : (H ⊢ Q a) → (H ⊢ ∃∃a, Q a) := by +theorem SLP.exists_intro [SLH β] {Q : α → SLP β} {a} : (H ⊢ Q a) → (H ⊢ ∃∃a, Q a) := by intro h st H unfold SLP.exists' exists a tauto -theorem exi_prop {Q : P → SLP p} : (H ⊢ P ⋆ ⊤) → (∀(p:P), H ⊢ Q p) → (H ⊢ ∃∃p, Q p) := by +theorem exi_prop [SLH β] {Q : P → SLP β} : (H ⊢ P ⋆ ⊤) → (∀(p:P), H ⊢ Q p) → (H ⊢ ∃∃p, Q p) := by intro h₁ h₂ unfold SLP.entails at * intro st hH @@ -285,7 +286,7 @@ theorem exi_prop {Q : P → SLP p} : (H ⊢ P ⋆ ⊤) → (∀(p:P), H ⊢ Q p) apply_assumption assumption -theorem exi_prop_l {H : P → SLP p} : ((x:P) → ((P ⋆ H x) ⊢ Q)) → ((∃∃x, H x) ⊢ Q) := by +theorem exi_prop_l [SLH β] {H : P → SLP β} : ((x:P) → ((P ⋆ H x) ⊢ Q)) → ((∃∃x, H x) ⊢ Q) := by intro h st unfold SLP.entails SLP.exists' at * rintro ⟨v, hH⟩ @@ -294,7 +295,7 @@ theorem exi_prop_l {H : P → SLP p} : ((x:P) → ((P ⋆ H x) ⊢ Q)) → ((∃ simp_all [Finmap.disjoint_empty, SLP.lift] simp_all -theorem use_right : (R ⊢ G ⋆ H) → (L ⋆ R ⊢ G ⋆ L ⋆ H) := by +theorem use_right [SLH β] {R G H L : SLP β} : (R ⊢ G ⋆ H) → (L ⋆ R ⊢ G ⋆ L ⋆ H) := by intro apply SLP.entails_trans apply SLP.star_mono_l diff --git a/Lampe/Lampe/ValHeap.lean b/Lampe/Lampe/ValHeap.lean deleted file mode 100644 index 14cdf1f..0000000 --- a/Lampe/Lampe/ValHeap.lean +++ /dev/null @@ -1,18 +0,0 @@ -import Mathlib -import Lampe.Tp - -lemma Finmap.insert_eq_singleton_union [DecidableEq α] {ref : α}: - m.insert ref v = Finmap.singleton ref v ∪ m := by rfl - -@[simp] -lemma Finmap.singleton_disjoint_of_not_mem (hp : ref ∉ s): - Finmap.Disjoint (Finmap.singleton ref v) s := by - simp_all [Finmap.Disjoint] - -namespace Lampe - -def AnyValue (p : Prime) := (tp : Tp) × tp.denote p - -abbrev ValHeap (p : Prime) := Finmap (fun (_ : Ref) => AnyValue p) - -end Lampe From 5d1cc684bacbfe50323fae5ded39fd4b232043d0 Mon Sep 17 00:00:00 2001 From: utkn Date: Fri, 22 Nov 2024 12:49:52 +0300 Subject: [PATCH 11/39] Steps tactic fixed to use SLP (State _). Broken proof steps replaced with `sorry` to be completed later --- Lampe/Lampe.lean | 4 +- Lampe/Lampe/Basic.lean | 2 +- Lampe/Lampe/Hoare/Builtins.lean | 39 ++++++------ Lampe/Lampe/Hoare/Total.lean | 8 +-- Lampe/Lampe/Semantics.lean | 45 ++++++-------- Lampe/Lampe/SeparationLogic/State.lean | 8 ++- Lampe/Lampe/SeparationLogic/ValHeap.lean | 4 ++ Lampe/Lampe/Tactic/SeparationLogic.lean | 79 ++++++++++++++---------- 8 files changed, 100 insertions(+), 89 deletions(-) diff --git a/Lampe/Lampe.lean b/Lampe/Lampe.lean index 8058bb3..97150ad 100644 --- a/Lampe/Lampe.lean +++ b/Lampe/Lampe.lean @@ -11,7 +11,7 @@ nr_def simple_muts<>(x : Field) -> Field { example : STHoare p Γ ⟦⟧ (simple_muts.fn.body _ h![] h![x]) fun v => v = x := by simp only [simple_muts] - steps + steps <;> try (exact _) simp_all nr_def weirdEq(x : I, y : I) -> Unit { @@ -89,7 +89,7 @@ example {p Γ x y}: STHoare p Γ ⟦⟧ (simple_if_else.fn.body (Tp.denote p) h! . simp only [decide_True, exists_const] sl contradiction - . simp_all + . exact _ nr_def simple_lambda<>(x : Field) -> Field { let foo = |a|: Field -> Field { a }; diff --git a/Lampe/Lampe/Basic.lean b/Lampe/Lampe/Basic.lean index 149b4f4..98c0364 100644 --- a/Lampe/Lampe/Basic.lean +++ b/Lampe/Lampe/Basic.lean @@ -4,4 +4,4 @@ import Lampe.Ast import Lampe.Tp import Lampe.Hoare.Total import Lampe.Hoare.SepTotal --- import Lampe.Tactic.SeparationLogic +import Lampe.Tactic.SeparationLogic diff --git a/Lampe/Lampe/Hoare/Builtins.lean b/Lampe/Lampe/Hoare/Builtins.lean index f4a4191..59f2ab6 100644 --- a/Lampe/Lampe/Hoare/Builtins.lean +++ b/Lampe/Lampe/Hoare/Builtins.lean @@ -20,7 +20,6 @@ theorem pureBuiltin_intro {A : Type} {a : A} {sgn desc args} : . simp_all only [mapToValHeapCondition, SLP.true_star, exists_const] apply SLP.ent_star_top simp_all only [SLP.true_star, exists_const] - tauto . tauto . apply Builtin.genericPureOmni.err . tauto @@ -383,15 +382,15 @@ theorem readRef_intro: intro st rintro ⟨_, _, _, _, hs, _⟩ subst_vars - all_goals sorry - -- apply And.intro (by simp; rfl) - -- simp only [SLP.true_star, SLP.star_assoc] - -- exists (Finmap.singleton r ⟨tp, v⟩), ?_ - -- apply And.intro (by assumption) - -- apply And.intro rfl - -- apply And.intro (by simp [SLP.singleton]) - -- apply SLP.ent_star_top - -- assumption + apply And.intro (by sorry) + simp only [SLP.true_star, SLP.star_assoc] + rename_i st _ _ _ + exists st, ?_ + apply And.intro (by assumption) + apply And.intro rfl + apply And.intro (by sorry) + sorry + all_goals assumption theorem writeRef_intro: STHoare p Γ @@ -405,16 +404,16 @@ theorem writeRef_intro: rintro ⟨_, _, _, _, hs, _⟩ simp only [State.singleton] at hs subst_vars - sorry - -- apply And.intro (by simp) - -- simp only - -- simp only [Finmap.insert_eq_singleton_union, ←Finmap.union_assoc, Finmap.union_singleton, SLP.star_assoc] - -- use Finmap.singleton r ⟨tp, v'⟩, ?_ - -- apply And.intro (by assumption) - -- apply And.intro rfl - -- apply And.intro (by simp [SLP.singleton]) - -- apply SLP.ent_star_top - -- assumption + apply And.intro (by sorry) + simp only + simp only [Finmap.insert_eq_singleton_union, ←Finmap.union_assoc, Finmap.union_singleton, SLP.star_assoc] + rename_i st₁ st₂ _ _ + exists st₁, st₂ + apply And.intro (by assumption) + apply And.intro (by sorry) + apply And.intro (by sorry) + apply SLP.ent_star_top + assumption -- Misc diff --git a/Lampe/Lampe/Hoare/Total.lean b/Lampe/Lampe/Hoare/Total.lean index debd352..6fd1a27 100644 --- a/Lampe/Lampe/Hoare/Total.lean +++ b/Lampe/Lampe/Hoare/Total.lean @@ -67,9 +67,7 @@ theorem readRef_intro {ref}: unfold THoare intros constructor - constructor - tauto - tauto + constructor <;> tauto theorem writeRef_intro {ref v}: THoare p Γ @@ -79,9 +77,7 @@ theorem writeRef_intro {ref v}: unfold THoare intros constructor - constructor - tauto - tauto + constructor <;> tauto theorem fresh_intro: THoare p Γ diff --git a/Lampe/Lampe/Semantics.lean b/Lampe/Lampe/Semantics.lean index a10f366..7f63983 100644 --- a/Lampe/Lampe/Semantics.lean +++ b/Lampe/Lampe/Semantics.lean @@ -30,20 +30,20 @@ inductive Omni : (p : Prime) → Env → State p → Expr (Tp.denote p) tp → ( Omni p Γ st elseBranch Q → Omni p Γ st (Expr.ite false mainBranch elseBranch) Q | letIn : - Omni p Γ st e Q₁ → - (∀v st, Q₁ (some (st, v)) → Omni p Γ st (b v) Q) → - (Q₁ none → Q none) → - Omni p Γ st (.letIn e b) Q -| callBuiltin {Q} : - (b.omni p st.vals argTypes resType args (mapToValHeapCondition st.closures Q)) → - Omni p Γ st (Expr.call h![] argTypes resType (.builtin b) args) Q + Omni p Γ st e Q₁ → + (∀v st, Q₁ (some (st, v)) → Omni p Γ st (b v) Q) → + (Q₁ none → Q none) → + Omni p Γ st (.letIn e b) Q +| callBuiltin {st : State p} {Q : Option (State p × Tp.denote p resType) → Prop} : + (b.omni p st.vals argTypes resType args (mapToValHeapCondition st.closures Q)) → + Omni p Γ st (Expr.call h![] argTypes resType (.builtin b) args) Q | callDecl : - Γ fname = some fn → - (hkc : fn.generics = tyKinds) → - (htci : fn.inTps (hkc ▸ generics) = argTypes) → - (htco : fn.outTp (hkc ▸ generics) = res) → - Omni p Γ st (htco ▸ fn.body _ (hkc ▸ generics) (htci ▸ args)) Q → - Omni p Γ st (@Expr.call _ tyKinds generics argTypes res (.decl fname) args) Q + Γ fname = some fn → + (hkc : fn.generics = tyKinds) → + (htci : fn.inTps (hkc ▸ generics) = argTypes) → + (htco : fn.outTp (hkc ▸ generics) = res) → + Omni p Γ st (htco ▸ fn.body _ (hkc ▸ generics) (htci ▸ args)) Q → + Omni p Γ st (@Expr.call _ tyKinds generics argTypes res (.decl fname) args) Q | callLambda {cls : Closures} : cls.lookup lambdaRef = some fn → (hg : fn.generics = []) → @@ -52,12 +52,12 @@ inductive Omni : (p : Prime) → Env → State p → Expr (Tp.denote p) tp → ( Omni p Γ st (ho ▸ fn.body _ (hg ▸ h![]) (hi ▸ args)) Q → Omni p Γ st (Expr.call h![] argTps outTp (.lambda lambdaRef) args) Q | loopDone : - lo ≥ hi → - Omni p Γ st (.loop lo hi body) Q + lo ≥ hi → + Omni p Γ st (.loop lo hi body) Q | loopNext {s} {lo hi : U s} {body} : - lo < hi → - Omni p Γ st (.letIn (body lo) (fun _ => .loop (lo + 1) hi body)) Q → - Omni p Γ st (.loop lo hi body) Q + lo < hi → + Omni p Γ st (.letIn (body lo) (fun _ => .loop (lo + 1) hi body)) Q → + Omni p Γ st (.loop lo hi body) Q theorem Omni.consequence {p Γ st tp} {e : Expr (Tp.denote p) tp} {Q Q'}: Omni p Γ st e Q → @@ -113,15 +113,10 @@ theorem Omni.frame {p Γ tp} {st₁ st₂ : State p} {cls : Closures} {e : Expr assumption assumption · simp_all - | callBuiltin hq => + | callBuiltin => rename Builtin => b intros - constructor - have uni_vals {p' : Prime} {a b : State p'} : (a ∪ b).vals = (a.vals ∪ b.vals) := by rfl - have uni_cls {p' : Prime} {a b : State p'} : (a ∪ b).closures = (a.closures ∪ b.closures) := by rfl - rw [uni_vals, uni_cls] - unfold mapToValHeapCondition - simp_all + constructor <;> try tauto sorry | callDecl _ _ _ _ _ ih => intro diff --git a/Lampe/Lampe/SeparationLogic/State.lean b/Lampe/Lampe/SeparationLogic/State.lean index 3171c7a..3c97807 100644 --- a/Lampe/Lampe/SeparationLogic/State.lean +++ b/Lampe/Lampe/SeparationLogic/State.lean @@ -21,7 +21,13 @@ instance : Coe (State p) (ValHeap p) where def mapToValHeapCondition (closures : Finmap fun _ : Ref => Function) (Q : Option (State p × T) → Prop) : Option (ValHeap p × T) → Prop := - fun v => Q (v.map (fun (vals, t) => ⟨⟨vals, closures⟩, t⟩)) + fun vv => Q (vv.map (fun (vals, t) => ⟨⟨vals, closures⟩, t⟩)) + +/-- Maps a condition on `ValHeap`s to a condition on `State`s -/ +@[reducible] +def mapToStateCondition + (Q : Option (ValHeap p × T) → Prop) : Option (State p × T) → Prop := + fun stv => Q (stv.map (fun (st, t) => ⟨st.vals, t⟩)) def State.insertVal (self : State p) (r : Ref) (v : AnyValue p) : State p := ⟨self.vals.insert r v, self.closures⟩ diff --git a/Lampe/Lampe/SeparationLogic/ValHeap.lean b/Lampe/Lampe/SeparationLogic/ValHeap.lean index 5497621..e56ece7 100644 --- a/Lampe/Lampe/SeparationLogic/ValHeap.lean +++ b/Lampe/Lampe/SeparationLogic/ValHeap.lean @@ -9,6 +9,10 @@ lemma Finmap.singleton_disjoint_of_not_mem (hp : ref ∉ s): Finmap.Disjoint (Finmap.singleton ref v) s := by simp_all [Finmap.Disjoint] +theorem Finmap.union_self [DecidableEq α] {a : Finmap fun _: α => β} : + a ∪ a = a := by + sorry + namespace Lampe def AnyValue (p : Prime) := (tp : Tp) × tp.denote p diff --git a/Lampe/Lampe/Tactic/SeparationLogic.lean b/Lampe/Lampe/Tactic/SeparationLogic.lean index 5b20564..30e5ceb 100644 --- a/Lampe/Lampe/Tactic/SeparationLogic.lean +++ b/Lampe/Lampe/Tactic/SeparationLogic.lean @@ -1,4 +1,4 @@ -import Lampe.SeparationLogic.SLP +import Lampe.SeparationLogic.State import Lampe.Hoare.SepTotal import Lampe.Hoare.Builtins import Lampe.Syntax @@ -43,13 +43,13 @@ instance : ToString SLTerm := ⟨SLTerm.toString⟩ instance : Inhabited SLTerm := ⟨SLTerm.top⟩ -theorem star_exists [SLH β] {Q : α → SLP β} : ((∃∃x, Q x) ⋆ P) = (∃∃x, Q x ⋆ P) := by +theorem star_exists {Q : α → SLP (State p)} : ((∃∃x, Q x) ⋆ P) = (∃∃x, Q x ⋆ P) := by unfold SLP.exists' SLP.star funext st simp tauto -theorem exists_star [SLH β] {Q : α → SLP β} : ((∃∃x, Q x) ⋆ P) = (∃∃x, P ⋆ Q x) := by +theorem exists_star {Q : α → SLP (State p)} : ((∃∃x, Q x) ⋆ P) = (∃∃x, P ⋆ Q x) := by rw [star_exists] simp [SLP.star_comm] @@ -74,7 +74,8 @@ theorem ref_intro' {p} {x : Tp.denote p tp} {Γ P}: apply SLP.ent_star_top -theorem Lampe.SLP.skip_fst [SLH β] : (R₁ ⊢ Q ⋆ X) → ([a ↦ b] ⋆ X ⊢ R₂) → ([a ↦ b] ⋆ R₁ ⊢ Q ⋆ R₂) := by +theorem Lampe.SLP.skip_fst {Q : SLP (State p)} : + (R₁ ⊢ Q ⋆ X) → (([a ↦ b] : SLP (State p)) ⋆ X ⊢ R₂) → (([a ↦ b] : SLP (State p)) ⋆ R₁ ⊢ Q ⋆ R₂) := by intro h₁ h₂ apply entails_trans rotate_left @@ -88,30 +89,36 @@ theorem Lampe.SLP.skip_fst [SLH β] : (R₁ ⊢ Q ⋆ X) → ([a ↦ b] ⋆ X rw [SLP.star_comm] apply entails_self -theorem Lampe.SLP.skip_fst' : (⟦⟧ ⊢ Q ⋆ X) → ([a ↦ b] ⋆ X ⊢ R₂) → ([a ↦ b] ⊢ Q ⋆ R₂) := by +theorem Lampe.SLP.skip_fst' {Q : SLP (State p)} : + (⟦⟧ ⊢ Q ⋆ X) → (([a ↦ b] : SLP (State p)) ⋆ X ⊢ R₂) → (([a ↦ b] : SLP (State p)) ⊢ Q ⋆ R₂) := by intro h₁ h₂ - rw [←SLP.star_true (H:=[a ↦ b])] + rw [←SLP.star_true (H := ([a ↦ b] : SLP (State p)))] apply Lampe.SLP.skip_fst assumption assumption -theorem Lampe.SLP.entails_star_true [SLH β] {H : SLP β} : H ⊢ H ⋆ ⟦⟧ := by +theorem Lampe.SLP.entails_star_true {H : SLP (State p)} : H ⊢ H ⋆ ⟦⟧ := by simp [SLP.entails_self] -theorem SLP.eq_of_iff [SLH β] {P Q : SLP β} : (P ⊢ Q) → (Q ⊢ P) → P = Q := by +theorem SLP.eq_of_iff {Q : SLP (State p)} : (P ⊢ Q) → (Q ⊢ P) → P = Q := by intros apply funext intro apply eq_iff_iff.mpr apply Iff.intro <;> apply_assumption -theorem pluck_pure_l {P : Prop} : ([a ↦ b] ⋆ P) = (P ⋆ [a ↦ b]) := by simp [SLP.star_comm] -theorem pluck_pure_all_l {P : Prop} : (SLP.forall' f ⋆ P) = (P ⋆ SLP.forall' f) := by simp [SLP.star_comm] -theorem pluck_pure_l_assoc { P : Prop } {Q : SLP p} : ([a ↦ b] ⋆ P ⋆ Q) = (P ⋆ [a ↦ b] ⋆ Q) := by +theorem pluck_pure_l {P : Prop} : + (([a ↦ b] : SLP (State p)) ⋆ P) = (P ⋆ ([a ↦ b] : SLP (State p))) := by simp [SLP.star_comm] + +theorem pluck_pure_all_l : + ((SLP.forall' f ⋆ P) : SLP (State p)) = (P ⋆ SLP.forall' f) := by simp [SLP.star_comm] + +theorem pluck_pure_l_assoc { P : Prop } {Q : SLP (State p)} : + (([a ↦ b] : SLP (State p)) ⋆ P ⋆ Q) = (P ⋆ ([a ↦ b] : SLP (State p)) ⋆ Q) := by rw [SLP.star_comm, SLP.star_assoc] apply SLP.eq_of_iff <;> {apply SLP.star_mono_l; rw [SLP.star_comm]; apply SLP.entails_self} -theorem SLP.pure_star_pure {p} { P Q : Prop }: (P ⋆ Q) = (⟦P ∧ Q⟧ : SLP p) := by +theorem SLP.pure_star_pure {p} { P Q : Prop }: (P ⋆ Q) = (⟦P ∧ Q⟧ : SLP (State p)) := by unfold SLP.star SLP.lift funext st apply eq_iff_iff.mpr @@ -120,7 +127,8 @@ theorem SLP.pure_star_pure {p} { P Q : Prop }: (P ⋆ Q) = (⟦P ∧ Q⟧ : SLP simp_all · intro_cases use ∅, ∅ - simp_all [Finmap.disjoint_empty] + simp_all + apply SLH.disjoint_empty macro "h_norm" : tactic => `(tactic|( try simp only [SLP.star_assoc, pluck_pure_l, pluck_pure_l_assoc, pluck_pure_all_l, SLP.star_true, SLP.true_star, star_exists, exists_star]; @@ -129,7 +137,7 @@ macro "h_norm" : tactic => `(tactic|( subst_vars; )) -theorem SLP.pure_leftX [SLH β] (H Q R : SLP β) : (P → (H ⊢ Q ⋆ R)) → (P ⋆ H ⊢ Q ⋆ P ⋆ R) := by +theorem SLP.pure_leftX {H : SLP (State p)} : (P → (H ⊢ Q ⋆ R)) → (P ⋆ H ⊢ Q ⋆ P ⋆ R) := by intro apply SLP.pure_left intro @@ -141,20 +149,22 @@ theorem SLP.pure_leftX [SLH β] (H Q R : SLP β) : (P → (H ⊢ Q ⋆ R)) → ( tauto /-- only finisher, will waste mvars for top! -/ -theorem SLP.pure_ent_star_top [SLH β] : (P → Q) → ((P : SLP β) ⊢ Q ⋆ ⊤) := by +theorem SLP.pure_ent_star_top : (P → Q) → ((P : SLP (State p)) ⊢ Q ⋆ ⊤) := by intro h st hp rcases hp with ⟨_, rfl, hp⟩ use ∅, ∅ simp_all [SLH.disjoint_empty, SLP.lift] apply SLH.disjoint_empty -theorem star_mono_l_sing [SLH β] {P Q : SLP β} : (P ⊢ Q) → (v₁ = v₂) → ([r ↦ v₁] ⋆ P ⊢ [r ↦ v₂] ⋆ Q) := by +theorem star_mono_l_sing {Q : SLP (State p)} : + (P ⊢ Q) → (v₁ = v₂) → (([r ↦ v₁] : SLP (State p)) ⋆ P ⊢ [r ↦ v₂] ⋆ Q) := by intro h₁ h₂ rw [h₂] apply SLP.star_mono_l apply h₁ -theorem star_mono_l_sing' [SLH β] {Q : SLP β} : (⟦⟧ ⊢ Q) → (v₁ = v₂) → ([r ↦ v₁] ⊢ [r ↦ v₂] ⋆ Q) := by +theorem star_mono_l_sing' {Q : SLP (State p)} : + (⟦⟧ ⊢ Q) → (v₁ = v₂) → (([r ↦ v₁] : SLP (State p)) ⊢ [r ↦ v₂] ⋆ Q) := by intro h₁ h₂ rw [h₂] apply SLP.star_mono_l' @@ -175,7 +185,7 @@ partial def parseSLExpr (e: Expr): TacticM SLTerm := do let fst ← parseSLExpr (←args[1]?) let snd ← parseSLExpr (←args[2]?) return SLTerm.star e fst snd - if e.isAppOf ``SLP.singleton then + if e.isAppOf ``State.singleton then let args := e.getAppArgs let fst ←args[1]? let snd ← args[2]? @@ -217,25 +227,25 @@ partial def parseEntailment (e: Expr): TacticM (SLTerm × SLTerm) := do return (pre, post) else throwError "not an entailment" -theorem star_top_of_star_mvar [SLH β] {H Q R : SLP β} : (H ⊢ Q ⋆ R) → (H ⊢ Q ⋆ ⊤) := by +theorem star_top_of_star_mvar {Q : SLP (State p)} : (H ⊢ Q ⋆ R) → (H ⊢ Q ⋆ ⊤) := by intro h apply SLP.entails_trans assumption apply SLP.star_mono_l apply SLP.entails_top -theorem solve_left_with_leftovers [SLH β] {H Q R : SLP β} : (H ⊢ Q ⋆ R) → (R ⊢ P) → (H ⊢ Q ⋆ P) := by +theorem solve_left_with_leftovers {Q : SLP (State p)} : (H ⊢ Q ⋆ R) → (R ⊢ P) → (H ⊢ Q ⋆ P) := by intros apply SLP.entails_trans assumption apply SLP.star_mono_l assumption -theorem solve_with_true [SLH β] {H Q : SLP β} : (H ⊢ Q) → (H ⊢ Q ⋆ ⟦⟧) := by +theorem solve_with_true {Q : SLP (State p)} : (H ⊢ Q) → (H ⊢ Q ⋆ ⟦⟧) := by aesop -- partial def solveNonMVarEntailment (goal : MVarId) (lhs : SLTerm) (rhs : SLTerm): TacticM (List MVarId × SLTerm) := do -theorem pure_ent_pure_star_mv [SLH β] : (P → Q) → ((P : SLP β) ⊢ Q ⋆ ⟦⟧) := by +theorem pure_ent_pure_star_mv : (P → Q) → ((P : SLP (State p)) ⊢ Q ⋆ ⟦⟧) := by intro h apply SLP.pure_left' intro @@ -243,7 +253,7 @@ theorem pure_ent_pure_star_mv [SLH β] : (P → Q) → ((P : SLP β) ⊢ Q ⋆ tauto tauto -theorem pure_star_H_ent_pure_star_mv [SLH β] {H Q : SLP β} : (P → (H ⊢ Q ⋆ R)) → (P ⋆ H ⊢ Q ⋆ P ⋆ R) := by +theorem pure_star_H_ent_pure_star_mv {Q : SLP (State p)} : (P → (H ⊢ Q ⋆ R)) → (P ⋆ H ⊢ Q ⋆ P ⋆ R) := by intro apply SLP.pure_left intro @@ -253,7 +263,7 @@ theorem pure_star_H_ent_pure_star_mv [SLH β] {H Q : SLP β} : (P → (H ⊢ Q rw [SLP.star_comm] tauto -theorem skip_left_ent_star_mv [SLH β] {R P H L : SLP β} : (R ⊢ P ⋆ H) → (L ⋆ R ⊢ P ⋆ L ⋆ H) := by +theorem skip_left_ent_star_mv {R : SLP (State p)} : (R ⊢ P ⋆ H) → (L ⋆ R ⊢ P ⋆ L ⋆ H) := by intro h apply SLP.entails_trans apply SLP.star_mono_l @@ -263,19 +273,19 @@ theorem skip_left_ent_star_mv [SLH β] {R P H L : SLP β} : (R ⊢ P ⋆ H) → rw [SLP.star_comm] apply SLP.entails_self -theorem skip_evidence_pure [SLH β] {H : SLP β} : Q → (H ⊢ Q ⋆ H) := by +theorem skip_evidence_pure : Q → ((H : Prop) ⊢ Q ⋆ (H : SLP (State p))) := by intro apply SLP.pure_right tauto tauto -theorem SLP.exists_intro [SLH β] {Q : α → SLP β} {a} : (H ⊢ Q a) → (H ⊢ ∃∃a, Q a) := by +theorem SLP.exists_intro {Q : α → SLP (State p)} {a} : (H ⊢ Q a) → (H ⊢ ∃∃a, Q a) := by intro h st H unfold SLP.exists' exists a tauto -theorem exi_prop [SLH β] {Q : P → SLP β} : (H ⊢ P ⋆ ⊤) → (∀(p:P), H ⊢ Q p) → (H ⊢ ∃∃p, Q p) := by +theorem exi_prop {Q : P → SLP (State p)} : (H ⊢ (P : SLP (State p)) ⋆ ⊤) → (∀(p:P), H ⊢ Q p) → (H ⊢ ∃∃p, Q p) := by intro h₁ h₂ unfold SLP.entails at * intro st hH @@ -286,16 +296,17 @@ theorem exi_prop [SLH β] {Q : P → SLP β} : (H ⊢ P ⋆ ⊤) → (∀(p:P), apply_assumption assumption -theorem exi_prop_l [SLH β] {H : P → SLP β} : ((x:P) → ((P ⋆ H x) ⊢ Q)) → ((∃∃x, H x) ⊢ Q) := by +theorem exi_prop_l {H : P → SLP (State p)} : ((x : P) → ((P ⋆ H x) ⊢ Q)) → ((∃∃x, H x) ⊢ Q) := by intro h st unfold SLP.entails SLP.exists' at * rintro ⟨v, hH⟩ apply h use ∅, st - simp_all [Finmap.disjoint_empty, SLP.lift] - simp_all + simp_all [SLP.lift] + apply SLH.disjoint_empty + assumption -theorem use_right [SLH β] {R G H L : SLP β} : (R ⊢ G ⋆ H) → (L ⋆ R ⊢ G ⋆ L ⋆ H) := by +theorem use_right {L : SLP (State p)} : (R ⊢ G ⋆ H) → (L ⋆ R ⊢ G ⋆ L ⋆ H) := by intro apply SLP.entails_trans apply SLP.star_mono_l @@ -305,17 +316,17 @@ theorem use_right [SLH β] {R G H L : SLP β} : (R ⊢ G ⋆ H) → (L ⋆ R ⊢ rw [SLP.star_comm] apply SLP.entails_self -theorem singleton_congr {p} {r} {v₁ v₂ : AnyValue p} : (v₁ = v₂) → ([r ↦ v₁] ⊢ [r ↦ v₂]) := by +theorem singleton_congr {p} {r} {v₁ v₂ : AnyValue p} : (v₁ = v₂) → (([r ↦ v₁] : SLP (State p)) ⊢ [r ↦ v₂]) := by intro h rw [h] apply SLP.entails_self -theorem singleton_congr_mv {p} {r} {v₁ v₂ : AnyValue p} : (v₁ = v₂) → ([r ↦ v₁] ⊢ [r ↦ v₂] ⋆ ⟦⟧) := by +theorem singleton_congr_mv {p} {r} {v₁ v₂ : AnyValue p} : (v₁ = v₂) → (([r ↦ v₁] : SLP (State p)) ⊢ [r ↦ v₂] ⋆ ⟦⟧) := by rintro rfl simp apply SLP.entails_self -theorem singleton_star_congr {p} {r} {v₁ v₂ : AnyValue p} {R} : (v₁ = v₂) → ([r ↦ v₁] ⋆ R ⊢ [r ↦ v₂] ⋆ R) := by +theorem singleton_star_congr {p} {r} {v₁ v₂ : AnyValue p} {R} : (v₁ = v₂) → (([r ↦ v₁] : SLP (State p)) ⋆ R ⊢ [r ↦ v₂] ⋆ R) := by rintro rfl apply SLP.entails_self From 19e84de67540a9f7e651e8ce9c11ed179f086e5d Mon Sep 17 00:00:00 2001 From: utkn Date: Fri, 22 Nov 2024 17:36:58 +0300 Subject: [PATCH 12/39] Some proofs completed --- Lampe/Lampe/Hoare/Builtins.lean | 62 ++++++++++++++++-------- Lampe/Lampe/Hoare/Total.lean | 23 +++------ Lampe/Lampe/Semantics.lean | 3 ++ Lampe/Lampe/SeparationLogic/SLH.lean | 3 +- Lampe/Lampe/SeparationLogic/State.lean | 21 ++++++-- Lampe/Lampe/SeparationLogic/ValHeap.lean | 12 ++++- 6 files changed, 84 insertions(+), 40 deletions(-) diff --git a/Lampe/Lampe/Hoare/Builtins.lean b/Lampe/Lampe/Hoare/Builtins.lean index 59f2ab6..a0d26d7 100644 --- a/Lampe/Lampe/Hoare/Builtins.lean +++ b/Lampe/Lampe/Hoare/Builtins.lean @@ -46,7 +46,6 @@ lemma pureBuiltin_intro_consequence rintro ⟨_, _⟩ simp_all [SLP.entails_top] - -- Arithmetics theorem uAdd_intro : STHoarePureBuiltin p Γ Builtin.uAdd (by tauto) h![a, b] := by @@ -360,7 +359,7 @@ theorem ref_intro: apply And.intro (by rw [SLH.disjoint_symm_iff]; apply SLH.disjoint_empty) constructor . simp only [State.insertVal, Finmap.insert_eq_singleton_union, SLH_union_empty] - simp only [State.union_parts, Finmap.union_self] + simp only [State.union_parts_left, Finmap.union_self] . apply And.intro ?_ (by simp) exists (⟨Finmap.singleton r ⟨tp, v⟩, ∅⟩), st constructor @@ -368,7 +367,7 @@ theorem ref_intro: apply And.intro (by simp [Finmap.singleton_disjoint_of_not_mem hr]) (by tauto) . simp_all only apply And.intro _ (by trivial) - simp only [State.union_parts, Finmap.empty_union, Finmap.union_self] + simp only [State.union_parts_left, Finmap.empty_union, Finmap.union_self] theorem readRef_intro: STHoare p Γ @@ -382,15 +381,26 @@ theorem readRef_intro: intro st rintro ⟨_, _, _, _, hs, _⟩ subst_vars - apply And.intro (by sorry) + apply And.intro + . simp only [State.union_vals] + rename_i st₁ st₂ _ _ + have hsome : Finmap.lookup r st₁.vals = some ⟨tp, v⟩ := by + unfold State.valSingleton at hs + rw [hs] + apply Finmap.lookup_singleton_eq + rw [Finmap.lookup_union_left] + tauto + apply Finmap.lookup_isSome.mp + rw [hsome] + apply Option.isSome_some simp only [SLP.true_star, SLP.star_assoc] - rename_i st _ _ _ - exists st, ?_ - apply And.intro (by assumption) - apply And.intro rfl - apply And.intro (by sorry) - sorry - all_goals assumption + rename_i st₁ st₂ _ _ + exists (⟨Finmap.singleton r ⟨tp, v⟩, st₁.closures⟩), ?_ + unfold State.valSingleton at hs + rw [←hs] + repeat apply And.intro (by tauto) + apply SLP.ent_star_top + assumption theorem writeRef_intro: STHoare p Γ @@ -402,18 +412,30 @@ theorem writeRef_intro: apply THoare.consequence ?_ THoare.writeRef_intro (fun _ => SLP.entails_self) intro st rintro ⟨_, _, _, _, hs, _⟩ - simp only [State.singleton] at hs + simp only [State.valSingleton] at hs subst_vars - apply And.intro (by sorry) - simp only + apply And.intro + . rename_i st₁ st₂ _ _ + have _ : r ∈ st₁.vals := by rw [hs]; tauto + simp_all [State.membership_in_val] simp only [Finmap.insert_eq_singleton_union, ←Finmap.union_assoc, Finmap.union_singleton, SLP.star_assoc] rename_i st₁ st₂ _ _ - exists st₁, st₂ - apply And.intro (by assumption) - apply And.intro (by sorry) - apply And.intro (by sorry) - apply SLP.ent_star_top - assumption + exists (⟨Finmap.singleton r ⟨tp, v'⟩, st₁.closures⟩), ?_ + refine ⟨?_, ?_, ?_, (by apply SLP.ent_star_top; assumption)⟩ + . simp only [SLH.disjoint] at * + have _ : r ∉ st₂.vals := by + rename_i h _ + rw [hs] at h + apply Finmap.singleton_disjoint_iff_not_mem.mp <;> tauto + refine ⟨?_, by tauto⟩ + apply Finmap.singleton_disjoint_of_not_mem + assumption + . simp only [State.union_vals, State.union_closures, State.union_parts_left] + apply State.eq_parts <;> try rfl + rw [←Finmap.union_assoc, hs] + simp [Finmap.union_singleton] + . simp_all + -- Misc diff --git a/Lampe/Lampe/Hoare/Total.lean b/Lampe/Lampe/Hoare/Total.lean index 6fd1a27..bbbcebf 100644 --- a/Lampe/Lampe/Hoare/Total.lean +++ b/Lampe/Lampe/Hoare/Total.lean @@ -18,17 +18,10 @@ def THoare namespace THoare -variable - {p : Prime} - {H H₁ H₂ : SLP (State p)} - {Γ : Env} - {tp : Tp} - {e : Expr (Tp.denote p) tp} - -theorem consequence {Q₁ Q₂} +theorem consequence {Q₁ Q₂} {H₂ : SLP (State p)} (h_pre_conseq : H₂ ⊢ H₁) (h_hoare: THoare p Γ H₁ e Q₁) - (h_post_conseq : ∀ v, Q₁ v ⊢ Q₂ v): + (h_post_conseq : ∀ v, Q₁ v ⊢ Q₂ v) : THoare p Γ H₂ e Q₂ := by unfold THoare intros @@ -36,19 +29,19 @@ theorem consequence {Q₁ Q₂} any_goals tauto rintro (_|_) <;> tauto -theorem var_intro {v} {P : Tp.denote p tp → SLP (State p)}: +theorem var_intro {v} {P : Tp.denote p tp → SLP (State p)} : THoare p Γ (P v) (.var v) P := by tauto theorem letIn_intro {P Q} (h₁ : THoare p Γ H e₁ P) - (h₂ : ∀v, THoare p Γ (P v) (e₂ v) Q): + (h₂ : ∀v, THoare p Γ (P v) (e₂ v) Q) : THoare p Γ H (.letIn e₁ e₂) Q := by unfold THoare intros constructor <;> tauto -theorem ref_intro {v}: +theorem ref_intro {v P} : THoare p Γ (fun st => ∀r, r ∉ st → P r ⟨(st.vals.insert r ⟨tp, v⟩), st.closures⟩) (.call h![] [tp] (Tp.ref tp) (.builtin .ref) h![v]) @@ -59,7 +52,7 @@ theorem ref_intro {v}: constructor tauto -theorem readRef_intro {ref}: +theorem readRef_intro {ref} : THoare p Γ (fun st => st.vals.lookup ref = some ⟨tp, v⟩ ∧ P v st) (.call h![] [tp.ref] tp (.builtin .readRef) h![ref]) @@ -69,7 +62,7 @@ theorem readRef_intro {ref}: constructor constructor <;> tauto -theorem writeRef_intro {ref v}: +theorem writeRef_intro {ref v} : THoare p Γ (fun st => ref ∈ st ∧ P () ⟨(st.vals.insert ref ⟨tp, v⟩), st.closures⟩) (.call h![] [tp.ref, tp] .unit (.builtin .writeRef) h![ref, v]) @@ -79,7 +72,7 @@ theorem writeRef_intro {ref v}: constructor constructor <;> tauto -theorem fresh_intro: +theorem fresh_intro {P} : THoare p Γ (∀∀v, P v) (.call h![] [] tp (.builtin .fresh) h![]) diff --git a/Lampe/Lampe/Semantics.lean b/Lampe/Lampe/Semantics.lean index 7f63983..0f84aef 100644 --- a/Lampe/Lampe/Semantics.lean +++ b/Lampe/Lampe/Semantics.lean @@ -15,6 +15,8 @@ def Env.ofModule (m : Module) : Env := fun i => (m.decls.find? fun d => d.name = @[reducible] def Env.extend (Γ₁ : Env) (Γ₂ : Env) : Env := fun i => Γ₁ i <|> Γ₂ i +-- Q none <-> Q' none where Q' = mapToValHeapCond... () + inductive Omni : (p : Prime) → Env → State p → Expr (Tp.denote p) tp → (Option (State p × Tp.denote p tp) → Prop) → Prop where | litField {Q} : Q (some (st, n)) → Omni p Γ st (.lit .field n) Q | litFalse {Q} : Q (some (st, false)) → Omni p Γ st (.lit .bool 0) Q @@ -117,6 +119,7 @@ theorem Omni.frame {p Γ tp} {st₁ st₂ : State p} {cls : Closures} {e : Expr rename Builtin => b intros constructor <;> try tauto + -- apply b.frame sorry | callDecl _ _ _ _ _ ih => intro diff --git a/Lampe/Lampe/SeparationLogic/SLH.lean b/Lampe/Lampe/SeparationLogic/SLH.lean index 66d99a1..a092366 100644 --- a/Lampe/Lampe/SeparationLogic/SLH.lean +++ b/Lampe/Lampe/SeparationLogic/SLH.lean @@ -2,6 +2,7 @@ import Lampe.Tp namespace Lampe +-- [TODO] LawfulHeap class SLH (α : Type _) where union : α → α → α disjoint : α → α → Prop @@ -11,7 +12,7 @@ class SLH (α : Type _) where union_assoc {s₁ s₂ s₃ : α} : union (union s₁ s₂) s₃ = union s₁ (union s₂ s₃) disjoint_symm_iff {a b : α} : disjoint a b ↔ disjoint b a union_comm_of_disjoint {s₁ s₂ : α} : disjoint s₁ s₂ → union s₁ s₂ = union s₂ s₁ - disjoint_union_left (x y z : α) : disjoint (union x y) z ↔ disjoint x z ∧ disjoint y z + disjoint_union_left (x y z : α) : disjoint (union x y) z ↔ disjoint x z ∧ disjoint y z -- one of them should be enough disjoint_union_right (x y z : α) : disjoint x (union y z) ↔ disjoint x y ∧ disjoint x z disjoint_empty (x : α) : disjoint empty x diff --git a/Lampe/Lampe/SeparationLogic/State.lean b/Lampe/Lampe/SeparationLogic/State.lean index 3c97807..96945ed 100644 --- a/Lampe/Lampe/SeparationLogic/State.lean +++ b/Lampe/Lampe/SeparationLogic/State.lean @@ -13,6 +13,8 @@ structure State (p : Prime) where instance : Membership Ref (State p) where mem := fun a e => e ∈ a.vals +lemma State.membership_in_val {a : State p} : e ∈ a ↔ e ∈ a.vals := by rfl + instance : Coe (State p) (ValHeap p) where coe := fun s => s.vals @@ -98,12 +100,25 @@ instance : SLH (State p) where tauto tauto -def State.singleton (r : Ref) (v : AnyValue p) : SLP (State p) := fun st => st.vals = Finmap.singleton r v +@[reducible] +def State.valSingleton (r : Ref) (v : AnyValue p) : SLP (State p) := fun st => st.vals = Finmap.singleton r v -notation:max "[" l " ↦ " r "]" => State.singleton l r +notation:max "[" l " ↦ " r "]" => State.valSingleton l r -theorem State.union_parts : +lemma State.union_parts_left : (State.mk v c ∪ st₂ = State.mk (v ∪ st₂.vals) (c ∪ st₂.closures)) := by aesop +lemma State.union_parts : + st₁ ∪ st₂ = State.mk (st₁.vals ∪ st₂.vals) (st₁.closures ∪ st₂.closures) := by + rfl + +@[simp] +lemma State.union_vals {st₁ st₂ : State p} : + (st₁ ∪ st₂).vals = (st₁.vals ∪ st₂.vals) := by rfl + +@[simp] +lemma State.union_closures {st₁ st₂ : State p} : + (st₁ ∪ st₂).closures = (st₁.closures ∪ st₂.closures) := by rfl + end Lampe diff --git a/Lampe/Lampe/SeparationLogic/ValHeap.lean b/Lampe/Lampe/SeparationLogic/ValHeap.lean index e56ece7..18aab9d 100644 --- a/Lampe/Lampe/SeparationLogic/ValHeap.lean +++ b/Lampe/Lampe/SeparationLogic/ValHeap.lean @@ -9,9 +9,19 @@ lemma Finmap.singleton_disjoint_of_not_mem (hp : ref ∉ s): Finmap.Disjoint (Finmap.singleton ref v) s := by simp_all [Finmap.Disjoint] +lemma Finmap.singleton_disjoint_iff_not_mem : + Finmap.Disjoint (Finmap.singleton ref v) s ↔ (ref ∉ s) := by + simp_all [Finmap.Disjoint] + theorem Finmap.union_self [DecidableEq α] {a : Finmap fun _: α => β} : a ∪ a = a := by - sorry + apply Finmap.ext_lookup + intro x + cases em (x ∈ a) + . rw [Finmap.lookup_union_left] + assumption + . rw [Finmap.lookup_union_left_of_not_in] + assumption namespace Lampe From ba1211aa7ae91ed49c1e049694ca560dd78f00c8 Mon Sep 17 00:00:00 2001 From: utkn Date: Fri, 22 Nov 2024 22:19:58 +0300 Subject: [PATCH 13/39] Omni frame proven --- Lampe/Lampe/Semantics.lean | 58 ++++++++++++++++++++++---- Lampe/Lampe/SeparationLogic/State.lean | 10 ++++- 2 files changed, 59 insertions(+), 9 deletions(-) diff --git a/Lampe/Lampe/Semantics.lean b/Lampe/Lampe/Semantics.lean index 0f84aef..17b7a13 100644 --- a/Lampe/Lampe/Semantics.lean +++ b/Lampe/Lampe/Semantics.lean @@ -15,8 +15,6 @@ def Env.ofModule (m : Module) : Env := fun i => (m.decls.find? fun d => d.name = @[reducible] def Env.extend (Γ₁ : Env) (Γ₂ : Env) : Env := fun i => Γ₁ i <|> Γ₂ i --- Q none <-> Q' none where Q' = mapToValHeapCond... () - inductive Omni : (p : Prime) → Env → State p → Expr (Tp.denote p) tp → (Option (State p × Tp.denote p tp) → Prop) → Prop where | litField {Q} : Q (some (st, n)) → Omni p Γ st (.lit .field n) Q | litFalse {Q} : Q (some (st, false)) → Omni p Γ st (.lit .bool 0) Q @@ -81,7 +79,8 @@ theorem Omni.consequence {p Γ st tp} {e : Expr (Tp.denote p) tp} {Q Q'}: apply loopNext (by assumption) tauto -theorem Omni.frame {p Γ tp} {st₁ st₂ : State p} {cls : Closures} {e : Expr (Tp.denote p) tp} {Q}: + +theorem Omni.frame {p Γ tp} {st₁ st₂ : State p} {e : Expr (Tp.denote p) tp} {Q}: Omni p Γ st₁ e Q → SLH.disjoint st₁ st₂ → Omni p Γ (st₁ ∪ st₂) e (fun stv => match stv with @@ -115,12 +114,57 @@ theorem Omni.frame {p Γ tp} {st₁ st₂ : State p} {cls : Closures} {e : Expr assumption assumption · simp_all - | callBuiltin => + | callBuiltin hq => rename Builtin => b intros - constructor <;> try tauto - -- apply b.frame - sorry + have hf := b.frame hq (st₂ := st₂) + apply callBuiltin + simp only [State.union_closures, State.union_vals] + rename_i _ _ _ _ _ st₁ Q' hd + unfold mapToValHeapCondition at * + simp_all [Option.map, SLH.disjoint] + convert hf + funext + rename_i x _ + cases x + rfl + rename_i v + simp_all only [SLP.star, eq_iff_iff] + apply Iff.intro + . intros hin + obtain ⟨s₁, ⟨s₂, hin'⟩⟩ := hin + obtain ⟨hin₁, hin₂, hin₃, hin₄⟩ := hin' + exists s₁, s₂ + simp only [SLH.disjoint] at * + refine ⟨by tauto, ?_, ?_, ?_⟩ + . simp only [State.union_parts] at hin₂ + apply State.eq_parts_inv at hin₂ + tauto + . have hc : s₁.closures = st₁.closures := by + obtain ⟨hin₅, hin₆⟩ := hin₁ + simp only [State.union_parts] at hin₂ + apply State.eq_parts_inv at hin₂ + obtain ⟨hin₇, hin₈⟩ := hin₂ + apply State.eq_parts_inv at hin₄ + obtain ⟨hin₉, hin₀⟩ := hin₄ + rw [←hin₀] at hin₈ + obtain ⟨_, hl⟩ := hd + rw [←hin₀] at hl + rw [Finmap.union_cancel hl (hin₆)] at hin₈ + tauto + rw [←hc] + tauto + . tauto + . intros hin + obtain ⟨s₁, ⟨s₂, hin'⟩⟩ := hin + obtain ⟨hin₁, hin₂, hin₃, hin₄⟩ := hin' + exists ⟨s₁, st₁.closures⟩, ⟨s₂, st₂.closures⟩ + simp only [SLH.disjoint] at * + refine ⟨by tauto, ?_, ?_, ?_⟩ + . simp only [State.union_parts] + apply State.eq_parts <;> tauto + . tauto + . rw [hin₄] | callDecl _ _ _ _ _ ih => intro constructor diff --git a/Lampe/Lampe/SeparationLogic/State.lean b/Lampe/Lampe/SeparationLogic/State.lean index 96945ed..adf3a91 100644 --- a/Lampe/Lampe/SeparationLogic/State.lean +++ b/Lampe/Lampe/SeparationLogic/State.lean @@ -18,14 +18,14 @@ lemma State.membership_in_val {a : State p} : e ∈ a ↔ e ∈ a.vals := by rfl instance : Coe (State p) (ValHeap p) where coe := fun s => s.vals -/-- Maps a condition on `State`s to a condition on `ValHeap`s by keeping the closures fixed -/ +/-- Maps a post-condition on `State`s to a post-condition on `ValHeap`s by keeping the closures fixed -/ @[reducible] def mapToValHeapCondition (closures : Finmap fun _ : Ref => Function) (Q : Option (State p × T) → Prop) : Option (ValHeap p × T) → Prop := fun vv => Q (vv.map (fun (vals, t) => ⟨⟨vals, closures⟩, t⟩)) -/-- Maps a condition on `ValHeap`s to a condition on `State`s -/ +/-- Maps a post-condition on `ValHeap`s to a post-condition on `State`s -/ @[reducible] def mapToStateCondition (Q : Option (ValHeap p × T) → Prop) : Option (State p × T) → Prop := @@ -40,6 +40,12 @@ theorem State.eq_parts : subst_vars rfl +theorem State.eq_parts_inv : + State.mk v c = State.mk v' c' → v = v' ∧ c = c' := by + intro + sorry + + instance : SLH (State p) where union := fun a b => ⟨a.vals ∪ b.vals, a.closures ∪ b.closures⟩ disjoint := fun a b => a.vals.Disjoint b.vals ∧ a.closures.Disjoint b.closures From ea4166ced8bb7df71077ef3c7d1e1ca20fda187d Mon Sep 17 00:00:00 2001 From: utkn Date: Sat, 23 Nov 2024 00:14:26 +0300 Subject: [PATCH 14/39] omni frame proof simplified (a bit) --- Lampe/Lampe/Semantics.lean | 35 ++++++++++++++--------------------- 1 file changed, 14 insertions(+), 21 deletions(-) diff --git a/Lampe/Lampe/Semantics.lean b/Lampe/Lampe/Semantics.lean index 17b7a13..a8f5633 100644 --- a/Lampe/Lampe/Semantics.lean +++ b/Lampe/Lampe/Semantics.lean @@ -117,26 +117,24 @@ theorem Omni.frame {p Γ tp} {st₁ st₂ : State p} {e : Expr (Tp.denote p) tp} | callBuiltin hq => rename Builtin => b intros - have hf := b.frame hq (st₂ := st₂) - apply callBuiltin + constructor simp only [State.union_closures, State.union_vals] - rename_i _ _ _ _ _ st₁ Q' hd + rename_i _ _ _ _ _ st₁ _ hd + have hf := b.frame hq (st₂ := st₂) unfold mapToValHeapCondition at * - simp_all [Option.map, SLH.disjoint] + simp_all only [Option.map, SLH.disjoint, true_implies] convert hf funext - rename_i x _ - cases x - rfl - rename_i v + casesm Option (ValHeap _ × _) <;> try rfl simp_all only [SLP.star, eq_iff_iff] apply Iff.intro - . intros hin - obtain ⟨s₁, ⟨s₂, hin'⟩⟩ := hin - obtain ⟨hin₁, hin₂, hin₃, hin₄⟩ := hin' - exists s₁, s₂ + all_goals ( + intros hin + obtain ⟨s₁, ⟨s₂, ⟨hin₁, hin₂, hin₃, hin₄⟩⟩⟩ := hin + ) + . exists s₁, s₂ simp only [SLH.disjoint] at * - refine ⟨by tauto, ?_, ?_, ?_⟩ + refine ⟨by tauto, ?_, ?_, by tauto⟩ . simp only [State.union_parts] at hin₂ apply State.eq_parts_inv at hin₂ tauto @@ -150,20 +148,15 @@ theorem Omni.frame {p Γ tp} {st₁ st₂ : State p} {e : Expr (Tp.denote p) tp} rw [←hin₀] at hin₈ obtain ⟨_, hl⟩ := hd rw [←hin₀] at hl - rw [Finmap.union_cancel hl (hin₆)] at hin₈ + rw [Finmap.union_cancel hl hin₆] at hin₈ tauto rw [←hc] tauto - . tauto - . intros hin - obtain ⟨s₁, ⟨s₂, hin'⟩⟩ := hin - obtain ⟨hin₁, hin₂, hin₃, hin₄⟩ := hin' - exists ⟨s₁, st₁.closures⟩, ⟨s₂, st₂.closures⟩ + . exists ⟨s₁, st₁.closures⟩, ⟨s₂, st₂.closures⟩ simp only [SLH.disjoint] at * - refine ⟨by tauto, ?_, ?_, ?_⟩ + refine ⟨by tauto, ?_, by tauto, ?_⟩ . simp only [State.union_parts] apply State.eq_parts <;> tauto - . tauto . rw [hin₄] | callDecl _ _ _ _ _ ih => intro From 9abf997c3ad174e166112922fadbd1c19aeaff74 Mon Sep 17 00:00:00 2001 From: utkn Date: Sat, 23 Nov 2024 19:05:44 +0300 Subject: [PATCH 15/39] Proof simplifications. Tactics work over heaps --- Lampe/Lampe/Hoare/Builtins.lean | 3 +- Lampe/Lampe/Semantics.lean | 30 +++++---- Lampe/Lampe/SeparationLogic/State.lean | 38 ++++++++---- Lampe/Lampe/SeparationLogic/ValHeap.lean | 4 +- Lampe/Lampe/Tactic/SeparationLogic.lean | 79 ++++++++++++------------ 5 files changed, 81 insertions(+), 73 deletions(-) diff --git a/Lampe/Lampe/Hoare/Builtins.lean b/Lampe/Lampe/Hoare/Builtins.lean index a0d26d7..979b48b 100644 --- a/Lampe/Lampe/Hoare/Builtins.lean +++ b/Lampe/Lampe/Hoare/Builtins.lean @@ -430,8 +430,7 @@ theorem writeRef_intro: refine ⟨?_, by tauto⟩ apply Finmap.singleton_disjoint_of_not_mem assumption - . simp only [State.union_vals, State.union_closures, State.union_parts_left] - apply State.eq_parts <;> try rfl + . simp only [State.mk.injEq, State.union_vals, State.union_closures, State.union_parts_left] rw [←Finmap.union_assoc, hs] simp [Finmap.union_singleton] . simp_all diff --git a/Lampe/Lampe/Semantics.lean b/Lampe/Lampe/Semantics.lean index a8f5633..111754a 100644 --- a/Lampe/Lampe/Semantics.lean +++ b/Lampe/Lampe/Semantics.lean @@ -79,7 +79,6 @@ theorem Omni.consequence {p Γ st tp} {e : Expr (Tp.denote p) tp} {Q Q'}: apply loopNext (by assumption) tauto - theorem Omni.frame {p Γ tp} {st₁ st₂ : State p} {e : Expr (Tp.denote p) tp} {Q}: Omni p Γ st₁ e Q → SLH.disjoint st₁ st₂ → @@ -122,9 +121,10 @@ theorem Omni.frame {p Γ tp} {st₁ st₂ : State p} {e : Expr (Tp.denote p) tp} rename_i _ _ _ _ _ st₁ _ hd have hf := b.frame hq (st₂ := st₂) unfold mapToValHeapCondition at * - simp_all only [Option.map, SLH.disjoint, true_implies] + simp_all only [SLH.disjoint, true_implies] convert hf funext + rename Option _ → Prop => Q' casesm Option (ValHeap _ × _) <;> try rfl simp_all only [SLP.star, eq_iff_iff] apply Iff.intro @@ -136,27 +136,25 @@ theorem Omni.frame {p Γ tp} {st₁ st₂ : State p} {e : Expr (Tp.denote p) tp} simp only [SLH.disjoint] at * refine ⟨by tauto, ?_, ?_, by tauto⟩ . simp only [State.union_parts] at hin₂ - apply State.eq_parts_inv at hin₂ - tauto + injection hin₂ . have hc : s₁.closures = st₁.closures := by - obtain ⟨hin₅, hin₆⟩ := hin₁ - simp only [State.union_parts] at hin₂ - apply State.eq_parts_inv at hin₂ - obtain ⟨hin₇, hin₈⟩ := hin₂ - apply State.eq_parts_inv at hin₄ - obtain ⟨hin₉, hin₀⟩ := hin₄ - rw [←hin₀] at hin₈ - obtain ⟨_, hl⟩ := hd - rw [←hin₀] at hl - rw [Finmap.union_cancel hl hin₆] at hin₈ + obtain ⟨_, hd₂⟩ := hin₁ + rw [State.union_parts] at hin₂ + injection hin₂ with _ hu + rw [State.mk.injEq] at hin₄ + obtain ⟨_, hin₀⟩ := hin₄ + rw [←hin₀] at hu + obtain ⟨_, hd₁⟩ := hd + rw [←hin₀] at hd₁ + rw [Finmap.union_cancel hd₁ hd₂] at hu tauto rw [←hc] tauto . exists ⟨s₁, st₁.closures⟩, ⟨s₂, st₂.closures⟩ simp only [SLH.disjoint] at * refine ⟨by tauto, ?_, by tauto, ?_⟩ - . simp only [State.union_parts] - apply State.eq_parts <;> tauto + . simp only [State.union_parts, State.mk.injEq] + tauto . rw [hin₄] | callDecl _ _ _ _ _ ih => intro diff --git a/Lampe/Lampe/SeparationLogic/State.lean b/Lampe/Lampe/SeparationLogic/State.lean index adf3a91..08db9d1 100644 --- a/Lampe/Lampe/SeparationLogic/State.lean +++ b/Lampe/Lampe/SeparationLogic/State.lean @@ -34,16 +34,15 @@ def mapToStateCondition def State.insertVal (self : State p) (r : Ref) (v : AnyValue p) : State p := ⟨self.vals.insert r v, self.closures⟩ -theorem State.eq_parts : - v = v' → c = c' → State.mk v c = State.mk v' c' := by - intros - subst_vars +lemma State.eq_constructor {st₁ : State p} : + (st₁ = st₂) ↔ (State.mk st₁.vals st₁.closures = State.mk st₂.vals st₂.closures) := by rfl -theorem State.eq_parts_inv : - State.mk v c = State.mk v' c' → v = v' ∧ c = c' := by - intro - sorry +@[simp] +lemma State.eq_closures : + (State.mk v c = State.mk v' c') → (c = c') := by + intro h + injection h instance : SLH (State p) where @@ -55,15 +54,15 @@ instance : SLH (State p) where simp only [Finmap.union_empty] union_assoc := by intros - simp only [Union.union] - apply State.eq_parts + simp only [Union.union, State.mk.injEq] + apply And.intro . apply Finmap.union_assoc . apply Finmap.union_assoc disjoint_symm_iff := by tauto union_comm_of_disjoint := by intros - simp only [Union.union] - apply State.eq_parts + simp only [Union.union, State.mk.injEq] + apply And.intro . apply Finmap.union_comm_of_disjoint tauto . apply Finmap.union_comm_of_disjoint @@ -111,10 +110,23 @@ def State.valSingleton (r : Ref) (v : AnyValue p) : SLP (State p) := fun st => s notation:max "[" l " ↦ " r "]" => State.valSingleton l r +@[reducible] +def State.clsSingleton (r : Ref) (v : Function) : SLP (State p) := fun st => st.closures = Finmap.singleton r v + +notation:max "[" l " ↣ " r "]" => State.cls l r + + +@[simp] lemma State.union_parts_left : (State.mk v c ∪ st₂ = State.mk (v ∪ st₂.vals) (c ∪ st₂.closures)) := by - aesop + rfl +@[simp] +lemma State.union_parts_right : + (st₂ ∪ State.mk v c = State.mk (st₂.vals ∪ v) (st₂.closures ∪ c)) := by + rfl + +@[simp] lemma State.union_parts : st₁ ∪ st₂ = State.mk (st₁.vals ∪ st₂.vals) (st₁.closures ∪ st₂.closures) := by rfl diff --git a/Lampe/Lampe/SeparationLogic/ValHeap.lean b/Lampe/Lampe/SeparationLogic/ValHeap.lean index 18aab9d..4616161 100644 --- a/Lampe/Lampe/SeparationLogic/ValHeap.lean +++ b/Lampe/Lampe/SeparationLogic/ValHeap.lean @@ -41,8 +41,8 @@ instance : SLH (ValHeap p) where disjoint_union_right := Finmap.disjoint_union_right disjoint_empty := Finmap.disjoint_empty -def ValHeap.singleton (r : Ref) (v : AnyValue p) : SLP (ValHeap p) := fun st => st = Finmap.singleton r v +-- def ValHeap.singleton (r : Ref) (v : AnyValue p) : SLP (ValHeap p) := fun st => st = Finmap.singleton r v -notation:max "[" l " ↦ " r "]" => ValHeap.singleton l r +-- notation:max "[" l " ↦ " r "]" => ValHeap.singleton l r end Lampe diff --git a/Lampe/Lampe/Tactic/SeparationLogic.lean b/Lampe/Lampe/Tactic/SeparationLogic.lean index 30e5ceb..f265e54 100644 --- a/Lampe/Lampe/Tactic/SeparationLogic.lean +++ b/Lampe/Lampe/Tactic/SeparationLogic.lean @@ -43,13 +43,13 @@ instance : ToString SLTerm := ⟨SLTerm.toString⟩ instance : Inhabited SLTerm := ⟨SLTerm.top⟩ -theorem star_exists {Q : α → SLP (State p)} : ((∃∃x, Q x) ⋆ P) = (∃∃x, Q x ⋆ P) := by +theorem star_exists [SLH α] {Q : α → SLP α} : ((∃∃x, Q x) ⋆ P) = (∃∃x, Q x ⋆ P) := by unfold SLP.exists' SLP.star funext st simp tauto -theorem exists_star {Q : α → SLP (State p)} : ((∃∃x, Q x) ⋆ P) = (∃∃x, P ⋆ Q x) := by +theorem exists_star [SLH α] {Q : α → SLP α} : ((∃∃x, Q x) ⋆ P) = (∃∃x, P ⋆ Q x) := by rw [star_exists] simp [SLP.star_comm] @@ -74,8 +74,7 @@ theorem ref_intro' {p} {x : Tp.denote p tp} {Γ P}: apply SLP.ent_star_top -theorem Lampe.SLP.skip_fst {Q : SLP (State p)} : - (R₁ ⊢ Q ⋆ X) → (([a ↦ b] : SLP (State p)) ⋆ X ⊢ R₂) → (([a ↦ b] : SLP (State p)) ⋆ R₁ ⊢ Q ⋆ R₂) := by +theorem Lampe.SLP.skip_fst : (R₁ ⊢ Q ⋆ X) → ([a ↦ b] ⋆ X ⊢ R₂) → ([a ↦ b] ⋆ R₁ ⊢ Q ⋆ R₂) := by intro h₁ h₂ apply entails_trans rotate_left @@ -89,36 +88,34 @@ theorem Lampe.SLP.skip_fst {Q : SLP (State p)} : rw [SLP.star_comm] apply entails_self -theorem Lampe.SLP.skip_fst' {Q : SLP (State p)} : - (⟦⟧ ⊢ Q ⋆ X) → (([a ↦ b] : SLP (State p)) ⋆ X ⊢ R₂) → (([a ↦ b] : SLP (State p)) ⊢ Q ⋆ R₂) := by +theorem Lampe.SLP.skip_fst' : (⟦⟧ ⊢ Q ⋆ X) → ([a ↦ b] ⋆ X ⊢ R₂) → ([a ↦ b] ⊢ Q ⋆ R₂) := by intro h₁ h₂ - rw [←SLP.star_true (H := ([a ↦ b] : SLP (State p)))] + rw [←SLP.star_true (H:=[a ↦ b])] apply Lampe.SLP.skip_fst assumption assumption -theorem Lampe.SLP.entails_star_true {H : SLP (State p)} : H ⊢ H ⋆ ⟦⟧ := by +theorem Lampe.SLP.entails_star_true [SLH α] {H : SLP α} : H ⊢ H ⋆ ⟦⟧ := by simp [SLP.entails_self] -theorem SLP.eq_of_iff {Q : SLP (State p)} : (P ⊢ Q) → (Q ⊢ P) → P = Q := by +theorem SLP.eq_of_iff [SLH α] {P Q : SLP α} : (P ⊢ Q) → (Q ⊢ P) → P = Q := by intros apply funext intro apply eq_iff_iff.mpr apply Iff.intro <;> apply_assumption -theorem pluck_pure_l {P : Prop} : - (([a ↦ b] : SLP (State p)) ⋆ P) = (P ⋆ ([a ↦ b] : SLP (State p))) := by simp [SLP.star_comm] +theorem pluck_pure_l {P : Prop} : ([a ↦ b] ⋆ P) = (P ⋆ [a ↦ b]) := by + simp [SLP.star_comm] -theorem pluck_pure_all_l : - ((SLP.forall' f ⋆ P) : SLP (State p)) = (P ⋆ SLP.forall' f) := by simp [SLP.star_comm] +theorem pluck_pure_all_l [SLH α] {P : Prop} {f : Prop → SLP α} : (SLP.forall' f ⋆ P) = (P ⋆ SLP.forall' f) := by + simp [SLP.star_comm] -theorem pluck_pure_l_assoc { P : Prop } {Q : SLP (State p)} : - (([a ↦ b] : SLP (State p)) ⋆ P ⋆ Q) = (P ⋆ ([a ↦ b] : SLP (State p)) ⋆ Q) := by +theorem pluck_pure_l_assoc {P : Prop} {Q : SLP (State p)} : ([a ↦ b] ⋆ P ⋆ Q) = (P ⋆ [a ↦ b] ⋆ Q) := by rw [SLP.star_comm, SLP.star_assoc] apply SLP.eq_of_iff <;> {apply SLP.star_mono_l; rw [SLP.star_comm]; apply SLP.entails_self} -theorem SLP.pure_star_pure {p} { P Q : Prop }: (P ⋆ Q) = (⟦P ∧ Q⟧ : SLP (State p)) := by +theorem SLP.pure_star_pure [SLH α] {P Q : Prop} : (P ⋆ Q) = (⟦P ∧ Q⟧ : SLP α) := by unfold SLP.star SLP.lift funext st apply eq_iff_iff.mpr @@ -127,8 +124,9 @@ theorem SLP.pure_star_pure {p} { P Q : Prop }: (P ⋆ Q) = (⟦P ∧ Q⟧ : SLP simp_all · intro_cases use ∅, ∅ - simp_all + refine ⟨?_, ?_, ?_, ?_, ?_⟩ apply SLH.disjoint_empty + all_goals simp_all [SLH.disjoint_empty] macro "h_norm" : tactic => `(tactic|( try simp only [SLP.star_assoc, pluck_pure_l, pluck_pure_l_assoc, pluck_pure_all_l, SLP.star_true, SLP.true_star, star_exists, exists_star]; @@ -137,7 +135,7 @@ macro "h_norm" : tactic => `(tactic|( subst_vars; )) -theorem SLP.pure_leftX {H : SLP (State p)} : (P → (H ⊢ Q ⋆ R)) → (P ⋆ H ⊢ Q ⋆ P ⋆ R) := by +theorem SLP.pure_leftX [SLH α] {H Q R : SLP α} : (P → (H ⊢ Q ⋆ R)) → (P ⋆ H ⊢ Q ⋆ P ⋆ R) := by intro apply SLP.pure_left intro @@ -149,22 +147,21 @@ theorem SLP.pure_leftX {H : SLP (State p)} : (P → (H ⊢ Q ⋆ R)) → (P ⋆ tauto /-- only finisher, will waste mvars for top! -/ -theorem SLP.pure_ent_star_top : (P → Q) → ((P : SLP (State p)) ⊢ Q ⋆ ⊤) := by +theorem SLP.pure_ent_star_top [SLH α] : (P → Q) → ((P : SLP α) ⊢ Q ⋆ ⊤) := by intro h st hp rcases hp with ⟨_, rfl, hp⟩ use ∅, ∅ - simp_all [SLH.disjoint_empty, SLP.lift] + refine ⟨?_, ?_, ?_, ?_⟩ apply SLH.disjoint_empty + all_goals simp_all [SLH.disjoint_empty, SLP.lift] -theorem star_mono_l_sing {Q : SLP (State p)} : - (P ⊢ Q) → (v₁ = v₂) → (([r ↦ v₁] : SLP (State p)) ⋆ P ⊢ [r ↦ v₂] ⋆ Q) := by +theorem star_mono_l_sing : (P ⊢ Q) → (v₁ = v₂) → ([r ↦ v₁] ⋆ P ⊢ [r ↦ v₂] ⋆ Q) := by intro h₁ h₂ rw [h₂] apply SLP.star_mono_l apply h₁ -theorem star_mono_l_sing' {Q : SLP (State p)} : - (⟦⟧ ⊢ Q) → (v₁ = v₂) → (([r ↦ v₁] : SLP (State p)) ⊢ [r ↦ v₂] ⋆ Q) := by +theorem star_mono_l_sing' : (⟦⟧ ⊢ Q) → (v₁ = v₂) → ([r ↦ v₁] ⊢ [r ↦ v₂] ⋆ Q) := by intro h₁ h₂ rw [h₂] apply SLP.star_mono_l' @@ -185,7 +182,7 @@ partial def parseSLExpr (e: Expr): TacticM SLTerm := do let fst ← parseSLExpr (←args[1]?) let snd ← parseSLExpr (←args[2]?) return SLTerm.star e fst snd - if e.isAppOf ``State.singleton then + if e.isAppOf ``State.valSingleton then let args := e.getAppArgs let fst ←args[1]? let snd ← args[2]? @@ -227,25 +224,25 @@ partial def parseEntailment (e: Expr): TacticM (SLTerm × SLTerm) := do return (pre, post) else throwError "not an entailment" -theorem star_top_of_star_mvar {Q : SLP (State p)} : (H ⊢ Q ⋆ R) → (H ⊢ Q ⋆ ⊤) := by +theorem star_top_of_star_mvar [SLH α] {H Q R : SLP α} : (H ⊢ Q ⋆ R) → (H ⊢ Q ⋆ ⊤) := by intro h apply SLP.entails_trans assumption apply SLP.star_mono_l apply SLP.entails_top -theorem solve_left_with_leftovers {Q : SLP (State p)} : (H ⊢ Q ⋆ R) → (R ⊢ P) → (H ⊢ Q ⋆ P) := by +theorem solve_left_with_leftovers [SLH α] {H Q R : SLP α} : (H ⊢ Q ⋆ R) → (R ⊢ P) → (H ⊢ Q ⋆ P) := by intros apply SLP.entails_trans assumption apply SLP.star_mono_l assumption -theorem solve_with_true {Q : SLP (State p)} : (H ⊢ Q) → (H ⊢ Q ⋆ ⟦⟧) := by +theorem solve_with_true [SLH α] {H Q : SLP α}: (H ⊢ Q) → (H ⊢ Q ⋆ ⟦⟧) := by aesop -- partial def solveNonMVarEntailment (goal : MVarId) (lhs : SLTerm) (rhs : SLTerm): TacticM (List MVarId × SLTerm) := do -theorem pure_ent_pure_star_mv : (P → Q) → ((P : SLP (State p)) ⊢ Q ⋆ ⟦⟧) := by +theorem pure_ent_pure_star_mv [SLH α] : (P → Q) → ((P : SLP α) ⊢ Q ⋆ ⟦⟧) := by intro h apply SLP.pure_left' intro @@ -253,7 +250,7 @@ theorem pure_ent_pure_star_mv : (P → Q) → ((P : SLP (State p)) ⊢ Q ⋆ ⟦ tauto tauto -theorem pure_star_H_ent_pure_star_mv {Q : SLP (State p)} : (P → (H ⊢ Q ⋆ R)) → (P ⋆ H ⊢ Q ⋆ P ⋆ R) := by +theorem pure_star_H_ent_pure_star_mv [SLH α] {H Q R : SLP α} : (P → (H ⊢ Q ⋆ R)) → (P ⋆ H ⊢ Q ⋆ P ⋆ R) := by intro apply SLP.pure_left intro @@ -263,7 +260,7 @@ theorem pure_star_H_ent_pure_star_mv {Q : SLP (State p)} : (P → (H ⊢ Q ⋆ R rw [SLP.star_comm] tauto -theorem skip_left_ent_star_mv {R : SLP (State p)} : (R ⊢ P ⋆ H) → (L ⋆ R ⊢ P ⋆ L ⋆ H) := by +theorem skip_left_ent_star_mv [SLH α] {R L P H : SLP α} : (R ⊢ P ⋆ H) → (L ⋆ R ⊢ P ⋆ L ⋆ H) := by intro h apply SLP.entails_trans apply SLP.star_mono_l @@ -273,19 +270,21 @@ theorem skip_left_ent_star_mv {R : SLP (State p)} : (R ⊢ P ⋆ H) → (L ⋆ R rw [SLP.star_comm] apply SLP.entails_self -theorem skip_evidence_pure : Q → ((H : Prop) ⊢ Q ⋆ (H : SLP (State p))) := by +theorem skip_evidence_pure [SLH α] {H : SLP α} : Q → (H ⊢ Q ⋆ H) := by intro apply SLP.pure_right tauto tauto +-- [TODO] make this generic over `LawfulHeap`s theorem SLP.exists_intro {Q : α → SLP (State p)} {a} : (H ⊢ Q a) → (H ⊢ ∃∃a, Q a) := by intro h st H unfold SLP.exists' exists a tauto -theorem exi_prop {Q : P → SLP (State p)} : (H ⊢ (P : SLP (State p)) ⋆ ⊤) → (∀(p:P), H ⊢ Q p) → (H ⊢ ∃∃p, Q p) := by +-- [TODO] make this generic over `LawfulHeap`s +theorem exi_prop {Q : P → SLP (State p)} : (H ⊢ (P : SLP (State p)) ⋆ ⊤) → (∀(p : P), H ⊢ Q p) → (H ⊢ ∃∃p, Q p) := by intro h₁ h₂ unfold SLP.entails at * intro st hH @@ -296,17 +295,17 @@ theorem exi_prop {Q : P → SLP (State p)} : (H ⊢ (P : SLP (State p)) ⋆ ⊤) apply_assumption assumption -theorem exi_prop_l {H : P → SLP (State p)} : ((x : P) → ((P ⋆ H x) ⊢ Q)) → ((∃∃x, H x) ⊢ Q) := by +theorem exi_prop_l [SLH α] {H : P → SLP α} {Q : SLP α}: ((x : P) → ((P ⋆ H x) ⊢ Q)) → ((∃∃x, H x) ⊢ Q) := by intro h st unfold SLP.entails SLP.exists' at * rintro ⟨v, hH⟩ apply h use ∅, st - simp_all [SLP.lift] + refine ⟨?_, ?_, ?_, ?_⟩ apply SLH.disjoint_empty - assumption + all_goals simp_all [SLH.disjoint_empty, SLP.lift] -theorem use_right {L : SLP (State p)} : (R ⊢ G ⋆ H) → (L ⋆ R ⊢ G ⋆ L ⋆ H) := by +theorem use_right [SLH α] {R L G H : SLP α} : (R ⊢ G ⋆ H) → (L ⋆ R ⊢ G ⋆ L ⋆ H) := by intro apply SLP.entails_trans apply SLP.star_mono_l @@ -316,17 +315,17 @@ theorem use_right {L : SLP (State p)} : (R ⊢ G ⋆ H) → (L ⋆ R ⊢ G ⋆ L rw [SLP.star_comm] apply SLP.entails_self -theorem singleton_congr {p} {r} {v₁ v₂ : AnyValue p} : (v₁ = v₂) → (([r ↦ v₁] : SLP (State p)) ⊢ [r ↦ v₂]) := by +theorem singleton_congr {p} {r} {v₁ v₂ : AnyValue p} : (v₁ = v₂) → ([r ↦ v₁] ⊢ [r ↦ v₂]) := by intro h rw [h] apply SLP.entails_self -theorem singleton_congr_mv {p} {r} {v₁ v₂ : AnyValue p} : (v₁ = v₂) → (([r ↦ v₁] : SLP (State p)) ⊢ [r ↦ v₂] ⋆ ⟦⟧) := by +theorem singleton_congr_mv {p} {r} {v₁ v₂ : AnyValue p} : (v₁ = v₂) → ([r ↦ v₁] ⊢ [r ↦ v₂] ⋆ ⟦⟧) := by rintro rfl simp apply SLP.entails_self -theorem singleton_star_congr {p} {r} {v₁ v₂ : AnyValue p} {R} : (v₁ = v₂) → (([r ↦ v₁] : SLP (State p)) ⋆ R ⊢ [r ↦ v₂] ⋆ R) := by +theorem singleton_star_congr {p} {r} {v₁ v₂ : AnyValue p} {R} : (v₁ = v₂) → ([r ↦ v₁] ⋆ R ⊢ [r ↦ v₂] ⋆ R) := by rintro rfl apply SLP.entails_self From b90ad5404115268e4f3886bd048657f6fbc80c8e Mon Sep 17 00:00:00 2001 From: utkn Date: Sat, 23 Nov 2024 19:17:17 +0300 Subject: [PATCH 16/39] Renamed `SLH` to `LawfulHeap` --- Lampe/Lampe/Builtin/Memory.lean | 2 +- Lampe/Lampe/Hoare/Builtins.lean | 8 +- Lampe/Lampe/Hoare/Total.lean | 2 +- Lampe/Lampe/Semantics.lean | 8 +- Lampe/Lampe/SeparationLogic/LawfulHeap.lean | 58 ++++++++++ Lampe/Lampe/SeparationLogic/SLH.lean | 59 ---------- Lampe/Lampe/SeparationLogic/SLP.lean | 120 ++++++++++---------- Lampe/Lampe/SeparationLogic/State.lean | 4 +- Lampe/Lampe/SeparationLogic/ValHeap.lean | 4 +- Lampe/Lampe/Tactic/SeparationLogic.lean | 46 ++++---- 10 files changed, 155 insertions(+), 156 deletions(-) create mode 100644 Lampe/Lampe/SeparationLogic/LawfulHeap.lean delete mode 100644 Lampe/Lampe/SeparationLogic/SLH.lean diff --git a/Lampe/Lampe/Builtin/Memory.lean b/Lampe/Lampe/Builtin/Memory.lean index a55ced4..4710768 100644 --- a/Lampe/Lampe/Builtin/Memory.lean +++ b/Lampe/Lampe/Builtin/Memory.lean @@ -27,7 +27,7 @@ def ref : Builtin := { simp_all rfl simp [Finmap.insert_union] - simp_all [Finmap.insert_eq_singleton_union, SLH.disjoint, Finmap.disjoint_union_left] + simp_all [Finmap.insert_eq_singleton_union, LawfulHeap.disjoint, Finmap.disjoint_union_left] } inductive readRefOmni : Omni where diff --git a/Lampe/Lampe/Hoare/Builtins.lean b/Lampe/Lampe/Hoare/Builtins.lean index 979b48b..b8c568d 100644 --- a/Lampe/Lampe/Hoare/Builtins.lean +++ b/Lampe/Lampe/Hoare/Builtins.lean @@ -356,14 +356,14 @@ theorem ref_intro: simp only [SLP.true_star] intro st hH r hr exists (⟨Finmap.singleton r ⟨tp, v⟩, st.closures⟩ ∪ st), ∅ - apply And.intro (by rw [SLH.disjoint_symm_iff]; apply SLH.disjoint_empty) + apply And.intro (by rw [LawfulHeap.disjoint_symm_iff]; apply LawfulHeap.disjoint_empty) constructor - . simp only [State.insertVal, Finmap.insert_eq_singleton_union, SLH_union_empty] + . simp only [State.insertVal, Finmap.insert_eq_singleton_union, LawfulHeap_union_empty] simp only [State.union_parts_left, Finmap.union_self] . apply And.intro ?_ (by simp) exists (⟨Finmap.singleton r ⟨tp, v⟩, ∅⟩), st constructor - . simp only [SLH.disjoint] + . simp only [LawfulHeap.disjoint] apply And.intro (by simp [Finmap.singleton_disjoint_of_not_mem hr]) (by tauto) . simp_all only apply And.intro _ (by trivial) @@ -422,7 +422,7 @@ theorem writeRef_intro: rename_i st₁ st₂ _ _ exists (⟨Finmap.singleton r ⟨tp, v'⟩, st₁.closures⟩), ?_ refine ⟨?_, ?_, ?_, (by apply SLP.ent_star_top; assumption)⟩ - . simp only [SLH.disjoint] at * + . simp only [LawfulHeap.disjoint] at * have _ : r ∉ st₂.vals := by rename_i h _ rw [hs] at h diff --git a/Lampe/Lampe/Hoare/Total.lean b/Lampe/Lampe/Hoare/Total.lean index bbbcebf..ebddab3 100644 --- a/Lampe/Lampe/Hoare/Total.lean +++ b/Lampe/Lampe/Hoare/Total.lean @@ -1,7 +1,7 @@ import Lampe.Ast import Lampe.Tp import Lampe.Semantics -import Lampe.SeparationLogic.SLH +import Lampe.SeparationLogic.LawfulHeap namespace Lampe diff --git a/Lampe/Lampe/Semantics.lean b/Lampe/Lampe/Semantics.lean index 111754a..0241109 100644 --- a/Lampe/Lampe/Semantics.lean +++ b/Lampe/Lampe/Semantics.lean @@ -81,7 +81,7 @@ theorem Omni.consequence {p Γ st tp} {e : Expr (Tp.denote p) tp} {Q Q'}: theorem Omni.frame {p Γ tp} {st₁ st₂ : State p} {e : Expr (Tp.denote p) tp} {Q}: Omni p Γ st₁ e Q → - SLH.disjoint st₁ st₂ → + LawfulHeap.disjoint st₁ st₂ → Omni p Γ (st₁ ∪ st₂) e (fun stv => match stv with | some (st', v) => ((fun st => Q (some (st, v))) ⋆ (fun st => st = st₂)) st' | none => Q none @@ -121,7 +121,7 @@ theorem Omni.frame {p Γ tp} {st₁ st₂ : State p} {e : Expr (Tp.denote p) tp} rename_i _ _ _ _ _ st₁ _ hd have hf := b.frame hq (st₂ := st₂) unfold mapToValHeapCondition at * - simp_all only [SLH.disjoint, true_implies] + simp_all only [LawfulHeap.disjoint, true_implies] convert hf funext rename Option _ → Prop => Q' @@ -133,7 +133,7 @@ theorem Omni.frame {p Γ tp} {st₁ st₂ : State p} {e : Expr (Tp.denote p) tp} obtain ⟨s₁, ⟨s₂, ⟨hin₁, hin₂, hin₃, hin₄⟩⟩⟩ := hin ) . exists s₁, s₂ - simp only [SLH.disjoint] at * + simp only [LawfulHeap.disjoint] at * refine ⟨by tauto, ?_, ?_, by tauto⟩ . simp only [State.union_parts] at hin₂ injection hin₂ @@ -151,7 +151,7 @@ theorem Omni.frame {p Γ tp} {st₁ st₂ : State p} {e : Expr (Tp.denote p) tp} rw [←hc] tauto . exists ⟨s₁, st₁.closures⟩, ⟨s₂, st₂.closures⟩ - simp only [SLH.disjoint] at * + simp only [LawfulHeap.disjoint] at * refine ⟨by tauto, ?_, by tauto, ?_⟩ . simp only [State.union_parts, State.mk.injEq] tauto diff --git a/Lampe/Lampe/SeparationLogic/LawfulHeap.lean b/Lampe/Lampe/SeparationLogic/LawfulHeap.lean new file mode 100644 index 0000000..d1035dc --- /dev/null +++ b/Lampe/Lampe/SeparationLogic/LawfulHeap.lean @@ -0,0 +1,58 @@ +import Lampe.Tp + +namespace Lampe + +class LawfulHeap (α : Type _) where + union : α → α → α + disjoint : α → α → Prop + empty : α + + union_empty {a : α} : union a empty = a + union_assoc {s₁ s₂ s₃ : α} : union (union s₁ s₂) s₃ = union s₁ (union s₂ s₃) + disjoint_symm_iff {a b : α} : disjoint a b ↔ disjoint b a + union_comm_of_disjoint {s₁ s₂ : α} : disjoint s₁ s₂ → union s₁ s₂ = union s₂ s₁ + disjoint_union_left (x y z : α) : disjoint (union x y) z ↔ disjoint x z ∧ disjoint y z -- one of them should be enough + disjoint_union_right (x y z : α) : disjoint x (union y z) ↔ disjoint x y ∧ disjoint x z + disjoint_empty (x : α) : disjoint empty x + +instance [LawfulHeap α] : Union α := ⟨LawfulHeap.union⟩ +instance [LawfulHeap α] : EmptyCollection α := ⟨LawfulHeap.empty⟩ + +theorem LawfulHeap_union_empty_commutative [LawfulHeap α] {a : α} : + a ∪ ∅ = ∅ ∪ a := by + have h := LawfulHeap.disjoint_empty a + simp [Union.union, EmptyCollection.emptyCollection] + rw [LawfulHeap.union_comm_of_disjoint h] + +theorem LawfulHeap_union_comm_of_disjoint [LawfulHeap α] {s₁ s₂ : α}: + LawfulHeap.disjoint s₁ s₂ → s₁ ∪ s₂ = s₂ ∪ s₁ := by + simp [Union.union] + exact LawfulHeap.union_comm_of_disjoint + +@[simp] +theorem LawfulHeap_union_empty [LawfulHeap α] {a : α} : + a ∪ ∅ = a := by + apply LawfulHeap.union_empty + +@[simp] +theorem LawfulHeap_empty_union [LawfulHeap α] {a : α} : + ∅ ∪ a = a := by + rw [←LawfulHeap_union_empty_commutative] + apply LawfulHeap.union_empty + +theorem LawfulHeap_union_assoc [LawfulHeap α] {s₁ s₂ s₃ : α} : + s₁ ∪ s₂ ∪ s₃ = s₁ ∪ (s₂ ∪ s₃) := by + simp [Union.union] + apply LawfulHeap.union_assoc + +theorem LawfulHeap_disjoint_union_left [LawfulHeap α] (x y z : α) : + LawfulHeap.disjoint (x ∪ y) z ↔ LawfulHeap.disjoint x z ∧ LawfulHeap.disjoint y z := by + simp [Union.union] + apply LawfulHeap.disjoint_union_left + +theorem LawfulHeap_disjoint_union_right [LawfulHeap α] (x y z : α) : + LawfulHeap.disjoint x (y ∪ z) ↔ LawfulHeap.disjoint x y ∧ LawfulHeap.disjoint x z := by + simp [Union.union] + apply LawfulHeap.disjoint_union_right + +end Lampe diff --git a/Lampe/Lampe/SeparationLogic/SLH.lean b/Lampe/Lampe/SeparationLogic/SLH.lean deleted file mode 100644 index a092366..0000000 --- a/Lampe/Lampe/SeparationLogic/SLH.lean +++ /dev/null @@ -1,59 +0,0 @@ -import Lampe.Tp - -namespace Lampe - --- [TODO] LawfulHeap -class SLH (α : Type _) where - union : α → α → α - disjoint : α → α → Prop - empty : α - - union_empty {a : α} : union a empty = a - union_assoc {s₁ s₂ s₃ : α} : union (union s₁ s₂) s₃ = union s₁ (union s₂ s₃) - disjoint_symm_iff {a b : α} : disjoint a b ↔ disjoint b a - union_comm_of_disjoint {s₁ s₂ : α} : disjoint s₁ s₂ → union s₁ s₂ = union s₂ s₁ - disjoint_union_left (x y z : α) : disjoint (union x y) z ↔ disjoint x z ∧ disjoint y z -- one of them should be enough - disjoint_union_right (x y z : α) : disjoint x (union y z) ↔ disjoint x y ∧ disjoint x z - disjoint_empty (x : α) : disjoint empty x - -instance [SLH α] : Union α := ⟨SLH.union⟩ -instance [SLH α] : EmptyCollection α := ⟨SLH.empty⟩ - -theorem SLH_union_empty_commutative [SLH α] {a : α} : - a ∪ ∅ = ∅ ∪ a := by - have h := SLH.disjoint_empty a - simp [Union.union, EmptyCollection.emptyCollection] - rw [SLH.union_comm_of_disjoint h] - -theorem SLH_union_comm_of_disjoint [SLH α] {s₁ s₂ : α}: - SLH.disjoint s₁ s₂ → s₁ ∪ s₂ = s₂ ∪ s₁ := by - simp [Union.union] - exact SLH.union_comm_of_disjoint - -@[simp] -theorem SLH_union_empty [SLH α] {a : α} : - a ∪ ∅ = a := by - apply SLH.union_empty - -@[simp] -theorem SLH_empty_union [SLH α] {a : α} : - ∅ ∪ a = a := by - rw [←SLH_union_empty_commutative] - apply SLH.union_empty - -theorem SLH_union_assoc [SLH α] {s₁ s₂ s₃ : α} : - s₁ ∪ s₂ ∪ s₃ = s₁ ∪ (s₂ ∪ s₃) := by - simp [Union.union] - apply SLH.union_assoc - -theorem SLH_disjoint_union_left [SLH α] (x y z : α) : - SLH.disjoint (x ∪ y) z ↔ SLH.disjoint x z ∧ SLH.disjoint y z := by - simp [Union.union] - apply SLH.disjoint_union_left - -theorem SLH_disjoint_union_right [SLH α] (x y z : α) : - SLH.disjoint x (y ∪ z) ↔ SLH.disjoint x y ∧ SLH.disjoint x z := by - simp [Union.union] - apply SLH.disjoint_union_right - -end Lampe diff --git a/Lampe/Lampe/SeparationLogic/SLP.lean b/Lampe/Lampe/SeparationLogic/SLP.lean index 4e5f5ad..dfc266a 100644 --- a/Lampe/Lampe/SeparationLogic/SLP.lean +++ b/Lampe/Lampe/SeparationLogic/SLP.lean @@ -1,28 +1,28 @@ import Lampe.Tactic.IntroCases -import Lampe.SeparationLogic.SLH +import Lampe.SeparationLogic.LawfulHeap namespace Lampe -def SLP (α) [SLH α] := α → Prop +def SLP (α) [LawfulHeap α] := α → Prop namespace SLP -def star [SLH α] (lhs rhs : SLP α) := fun st => - ∃ st₁ st₂, SLH.disjoint st₁ st₂ ∧ st = st₁ ∪ st₂ ∧ lhs st₁ ∧ rhs st₂ +def star [LawfulHeap α] (lhs rhs : SLP α) := fun st => + ∃ st₁ st₂, LawfulHeap.disjoint st₁ st₂ ∧ st = st₁ ∪ st₂ ∧ lhs st₁ ∧ rhs st₂ -def lift [SLH α] (pr : Prop) : SLP α := fun st => pr ∧ st = ∅ +def lift [LawfulHeap α] (pr : Prop) : SLP α := fun st => pr ∧ st = ∅ -def wand [SLH α] (lhs rhs : SLP α) : SLP α := - fun st => ∀st', SLH.disjoint st st' → lhs st' → rhs (st ∪ st') +def wand [LawfulHeap α] (lhs rhs : SLP α) : SLP α := + fun st => ∀st', LawfulHeap.disjoint st st' → lhs st' → rhs (st ∪ st') -def top [SLH α] : SLP α := fun _ => True +def top [LawfulHeap α] : SLP α := fun _ => True -def entails [SLH α] (a b : SLP α) := ∀st, a st → b st +def entails [LawfulHeap α] (a b : SLP α) := ∀st, a st → b st -def forall' [SLH α] (f : β → SLP α) : SLP α := fun st => ∀v, f v st -def exists' [SLH α] (f : β → SLP α) : SLP α := fun st => ∃v, f v st +def forall' [LawfulHeap α] (f : β → SLP α) : SLP α := fun st => ∀v, f v st +def exists' [LawfulHeap α] (f : β → SLP α) : SLP α := fun st => ∃v, f v st -instance [SLH α]: Coe Prop (SLP α) := ⟨lift⟩ +instance [LawfulHeap α]: Coe Prop (SLP α) := ⟨lift⟩ notation:max "⊤" => top @@ -42,31 +42,31 @@ macro "∃∃" xs:Lean.explicitBinders ", " b:term : term => Lean.expandExplicit open Lean.TSyntax.Compat in macro "∀∀" xs:Lean.explicitBinders ", " b:term : term => Lean.expandExplicitBinders ``forall' xs b -theorem entails_trans [SLH α] {P Q R : SLP α}: (P ⊢ Q) → (Q ⊢ R) → (P ⊢ R) := by tauto +theorem entails_trans [LawfulHeap α] {P Q R : SLP α}: (P ⊢ Q) → (Q ⊢ R) → (P ⊢ R) := by tauto section basic @[simp] -theorem apply_top [SLH α] {st : α} : ⊤ st := by trivial +theorem apply_top [LawfulHeap α] {st : α} : ⊤ st := by trivial -theorem forall_left [SLH β] {P : α → SLP β} : (P a ⊢ Q) → ((∀∀(a : α), P a) ⊢ Q) := by +theorem forall_left [LawfulHeap β] {P : α → SLP β} : (P a ⊢ Q) → ((∀∀(a : α), P a) ⊢ Q) := by unfold forall' tauto -theorem forall_right [SLH β] {H : SLP β} {H' : α → SLP β}: (∀x, H ⊢ H' x) → (H ⊢ ∀∀x, H' x) := by +theorem forall_right [LawfulHeap β] {H : SLP β} {H' : α → SLP β}: (∀x, H ⊢ H' x) → (H ⊢ ∀∀x, H' x) := by unfold forall' entails tauto -theorem pure_left [SLH β] {H H' : SLP β} : (P → (H ⊢ H')) → (P ⋆ H ⊢ H') := by +theorem pure_left [LawfulHeap β] {H H' : SLP β} : (P → (H ⊢ H')) → (P ⋆ H ⊢ H') := by unfold star entails lift intro_cases simp_all -theorem pure_left' [SLH α] {H : SLP α} : (P → (⟦⟧ ⊢ H)) → (P ⊢ H) := by +theorem pure_left' [LawfulHeap α] {H : SLP α} : (P → (⟦⟧ ⊢ H)) → (P ⊢ H) := by unfold entails lift tauto -theorem pure_right [SLH α] {H₁ H₂ : SLP α} : P → (H₁ ⊢ H₂) → (H₁ ⊢ P ⋆ H₂) := by +theorem pure_right [LawfulHeap α] {H₁ H₂ : SLP α} : P → (H₁ ⊢ H₂) → (H₁ ⊢ P ⋆ H₂) := by unfold star entails lift intros repeat apply Exists.intro @@ -76,15 +76,15 @@ theorem pure_right [SLH α] {H₁ H₂ : SLP α} : P → (H₁ ⊢ H₂) → (H apply And.intro rfl apply_assumption assumption - . simp only [SLH_empty_union] - . apply SLH.disjoint_empty + . simp only [LawfulHeap_empty_union] + . apply LawfulHeap.disjoint_empty -theorem entails_self [SLH α] {H : SLP α} : H ⊢ H := by tauto +theorem entails_self [LawfulHeap α] {H : SLP α} : H ⊢ H := by tauto -theorem entails_top [SLH α] {H : SLP α} : H ⊢ ⊤ := by tauto +theorem entails_top [LawfulHeap α] {H : SLP α} : H ⊢ ⊤ := by tauto @[simp] -theorem forall_unused [SLH β] {α : Type u} [Inhabited α] {P : SLP β} : (∀∀ (_ : α), P) = P := by +theorem forall_unused [LawfulHeap β] {α : Type u} [Inhabited α] {P : SLP β} : (∀∀ (_ : α), P) = P := by funext unfold forall' rw [eq_iff_iff] @@ -99,21 +99,21 @@ end basic section star -theorem star_comm [SLH α] {G H : SLP α} : (G ⋆ H) = (H ⋆ G) := by +theorem star_comm [LawfulHeap α] {G H : SLP α} : (G ⋆ H) = (H ⋆ G) := by funext unfold star rw [eq_iff_iff] apply Iff.intro <;> { intro_cases repeat apply Exists.intro - rw [SLH.disjoint_symm_iff] + rw [LawfulHeap.disjoint_symm_iff] apply And.intro (by assumption) - rw [SLH_union_comm_of_disjoint (by rw [SLH.disjoint_symm_iff]; assumption)] + rw [LawfulHeap_union_comm_of_disjoint (by rw [LawfulHeap.disjoint_symm_iff]; assumption)] tauto } @[simp] -theorem true_star [SLH α] {H : SLP α} : (⟦⟧ ⋆ H) = H := by +theorem true_star [LawfulHeap α] {H : SLP α} : (⟦⟧ ⋆ H) = H := by funext rw [eq_iff_iff] unfold lift star @@ -121,22 +121,22 @@ theorem true_star [SLH α] {H : SLP α} : (⟦⟧ ⋆ H) = H := by · simp_all · intro exists ∅, ?_ - simp_all [SLH.disjoint_empty] - apply SLH.disjoint_empty + simp_all [LawfulHeap.disjoint_empty] + apply LawfulHeap.disjoint_empty @[simp] -theorem star_true [SLH α] {H : SLP α} : (H ⋆ ⟦⟧) = H := by rw [star_comm]; simp +theorem star_true [LawfulHeap α] {H : SLP α} : (H ⋆ ⟦⟧) = H := by rw [star_comm]; simp @[simp] -theorem star_assoc [SLH α] {F G H : SLP α} : ((F ⋆ G) ⋆ H) = (F ⋆ G ⋆ H) := by +theorem star_assoc [LawfulHeap α] {F G H : SLP α} : ((F ⋆ G) ⋆ H) = (F ⋆ G ⋆ H) := by funext rw [eq_iff_iff] unfold star apply Iff.intro · intro_cases subst_vars - rw [SLH_union_assoc] - simp only [SLH_disjoint_union_left] at * + rw [LawfulHeap_union_assoc] + simp only [LawfulHeap_disjoint_union_left] at * cases_type And repeat apply Exists.intro apply And.intro ?_ @@ -147,11 +147,11 @@ theorem star_assoc [SLH α] {F G H : SLP α} : ((F ⋆ G) ⋆ H) = (F ⋆ G ⋆ apply And.intro rfl simp_all assumption - simp_all [SLH_disjoint_union_right] + simp_all [LawfulHeap_disjoint_union_right] · intro_cases subst_vars - rw [←SLH_union_assoc] - simp only [SLH_disjoint_union_right] at * + rw [←LawfulHeap_union_assoc] + simp only [LawfulHeap_disjoint_union_right] at * cases_type And repeat apply Exists.intro apply And.intro ?_ @@ -162,25 +162,25 @@ theorem star_assoc [SLH α] {F G H : SLP α} : ((F ⋆ G) ⋆ H) = (F ⋆ G ⋆ apply And.intro rfl simp_all assumption - simp_all [SLH_disjoint_union_left] + simp_all [LawfulHeap_disjoint_union_left] @[simp] -theorem ent_star_top [SLH α] {H : SLP α} : H ⊢ H ⋆ ⊤ := by +theorem ent_star_top [LawfulHeap α] {H : SLP α} : H ⊢ H ⋆ ⊤ := by intro _ _ exists ?_, ∅ - rw [SLH.disjoint_symm_iff] - simp_all [SLH.disjoint_empty] - apply SLH.disjoint_empty + rw [LawfulHeap.disjoint_symm_iff] + simp_all [LawfulHeap.disjoint_empty] + apply LawfulHeap.disjoint_empty -theorem star_mono_r [SLH α] {P Q R : SLP α} : (P ⊢ Q) → (P ⋆ R ⊢ Q ⋆ R) := by +theorem star_mono_r [LawfulHeap α] {P Q R : SLP α} : (P ⊢ Q) → (P ⋆ R ⊢ Q ⋆ R) := by unfold star entails tauto -theorem star_mono_l [SLH α] {P Q R : SLP α} : (P ⊢ Q) → (R ⋆ P ⊢ R ⋆ Q) := by +theorem star_mono_l [LawfulHeap α] {P Q R : SLP α} : (P ⊢ Q) → (R ⋆ P ⊢ R ⋆ Q) := by unfold star entails tauto -theorem star_mono_l' [SLH α] {P Q : SLP α} : (⟦⟧ ⊢ Q) → (P ⊢ P ⋆ Q) := by +theorem star_mono_l' [LawfulHeap α] {P Q : SLP α} : (⟦⟧ ⊢ Q) → (P ⊢ P ⋆ Q) := by unfold star entails lift intros simp_all @@ -189,29 +189,29 @@ theorem star_mono_l' [SLH α] {P Q : SLP α} : (⟦⟧ ⊢ Q) → (P ⊢ P ⋆ Q apply And.intro ?_ tauto simp - rw [SLH.disjoint_symm_iff] - apply SLH.disjoint_empty + rw [LawfulHeap.disjoint_symm_iff] + apply LawfulHeap.disjoint_empty -theorem star_mono [SLH α] {H₁ H₂ Q₁ Q₂ : SLP α} : (H₁ ⊢ H₂) → (Q₁ ⊢ Q₂) → (H₁ ⋆ Q₁ ⊢ H₂ ⋆ Q₂) := by +theorem star_mono [LawfulHeap α] {H₁ H₂ Q₁ Q₂ : SLP α} : (H₁ ⊢ H₂) → (Q₁ ⊢ Q₂) → (H₁ ⋆ Q₁ ⊢ H₂ ⋆ Q₂) := by unfold star entails tauto -theorem forall_star [SLH α] {P : α → SLP α} : (∀∀x, P x) ⋆ Q ⊢ ∀∀x, P x ⋆ Q := by +theorem forall_star [LawfulHeap α] {P : α → SLP α} : (∀∀x, P x) ⋆ Q ⊢ ∀∀x, P x ⋆ Q := by unfold star forall' tauto -theorem star_forall [SLH β] {P : α → SLP β} {Q : SLP β} : Q ⋆ (∀∀x, P x) ⊢ ∀∀x, Q ⋆ P x := by +theorem star_forall [LawfulHeap β] {P : α → SLP β} {Q : SLP β} : Q ⋆ (∀∀x, P x) ⊢ ∀∀x, Q ⋆ P x := by unfold star forall' tauto @[simp] -theorem top_star_top [SLH α] : (top ⋆ (⊤ : SLP α)) = (⊤ : SLP α) := by +theorem top_star_top [LawfulHeap α] : (top ⋆ (⊤ : SLP α)) = (⊤ : SLP α) := by unfold top star funext x simp exists ∅, x - simp [SLH.disjoint_empty] - apply SLH.disjoint_empty + simp [LawfulHeap.disjoint_empty] + apply LawfulHeap.disjoint_empty end star @@ -220,7 +220,7 @@ section wand variable {p : Prime} @[simp] -theorem wand_self_star [SLH α] {H : SLP α}: (H -⋆ H ⋆ top) = top := by +theorem wand_self_star [LawfulHeap α] {H : SLP α}: (H -⋆ H ⋆ top) = top := by funext unfold wand star apply eq_iff_iff.mpr @@ -235,26 +235,26 @@ theorem wand_self_star [SLH α] {H : SLP α}: (H -⋆ H ⋆ top) = top := by simp rotate_left rotate_left - rw [SLH_union_comm_of_disjoint (by assumption)] - rw [SLH.disjoint_symm_iff] + rw [LawfulHeap_union_comm_of_disjoint (by assumption)] + rw [LawfulHeap.disjoint_symm_iff] assumption -theorem wand_intro [SLH α] {A B C : SLP α} : (A ⋆ B ⊢ C) → (A ⊢ B -⋆ C) := by +theorem wand_intro [LawfulHeap α] {A B C : SLP α} : (A ⋆ B ⊢ C) → (A ⊢ B -⋆ C) := by unfold wand star entails intros intros apply_assumption tauto -theorem wand_cancel [SLH α] {P Q : SLP α} : (P ⋆ (P -⋆ Q)) ⊢ Q := by +theorem wand_cancel [LawfulHeap α] {P Q : SLP α} : (P ⋆ (P -⋆ Q)) ⊢ Q := by unfold star wand entails intro_cases subst_vars rename_i h - rw [SLH_union_comm_of_disjoint (by assumption)] + rw [LawfulHeap_union_comm_of_disjoint (by assumption)] apply_assumption - rw [SLH.disjoint_symm_iff] + rw [LawfulHeap.disjoint_symm_iff] tauto tauto diff --git a/Lampe/Lampe/SeparationLogic/State.lean b/Lampe/Lampe/SeparationLogic/State.lean index 08db9d1..9bc7964 100644 --- a/Lampe/Lampe/SeparationLogic/State.lean +++ b/Lampe/Lampe/SeparationLogic/State.lean @@ -1,4 +1,4 @@ -import Lampe.SeparationLogic.SLH +import Lampe.SeparationLogic.LawfulHeap import Lampe.SeparationLogic.ValHeap import Lampe.Ast @@ -45,7 +45,7 @@ lemma State.eq_closures : injection h -instance : SLH (State p) where +instance : LawfulHeap (State p) where union := fun a b => ⟨a.vals ∪ b.vals, a.closures ∪ b.closures⟩ disjoint := fun a b => a.vals.Disjoint b.vals ∧ a.closures.Disjoint b.closures empty := ⟨∅, ∅⟩ diff --git a/Lampe/Lampe/SeparationLogic/ValHeap.lean b/Lampe/Lampe/SeparationLogic/ValHeap.lean index 4616161..b34f75d 100644 --- a/Lampe/Lampe/SeparationLogic/ValHeap.lean +++ b/Lampe/Lampe/SeparationLogic/ValHeap.lean @@ -1,4 +1,4 @@ -import Lampe.SeparationLogic.SLH +import Lampe.SeparationLogic.LawfulHeap import Lampe.SeparationLogic.SLP lemma Finmap.insert_eq_singleton_union [DecidableEq α] {ref : α}: @@ -29,7 +29,7 @@ def AnyValue (p : Prime) := (tp : Tp) × tp.denote p abbrev ValHeap (p : Prime) := Finmap (fun (_ : Ref) => AnyValue p) -instance : SLH (ValHeap p) where +instance : LawfulHeap (ValHeap p) where union := fun a b => a ∪ b disjoint := fun a b => a.Disjoint b empty := ∅ diff --git a/Lampe/Lampe/Tactic/SeparationLogic.lean b/Lampe/Lampe/Tactic/SeparationLogic.lean index f265e54..223a751 100644 --- a/Lampe/Lampe/Tactic/SeparationLogic.lean +++ b/Lampe/Lampe/Tactic/SeparationLogic.lean @@ -43,13 +43,13 @@ instance : ToString SLTerm := ⟨SLTerm.toString⟩ instance : Inhabited SLTerm := ⟨SLTerm.top⟩ -theorem star_exists [SLH α] {Q : α → SLP α} : ((∃∃x, Q x) ⋆ P) = (∃∃x, Q x ⋆ P) := by +theorem star_exists [LawfulHeap α] {Q : α → SLP α} : ((∃∃x, Q x) ⋆ P) = (∃∃x, Q x ⋆ P) := by unfold SLP.exists' SLP.star funext st simp tauto -theorem exists_star [SLH α] {Q : α → SLP α} : ((∃∃x, Q x) ⋆ P) = (∃∃x, P ⋆ Q x) := by +theorem exists_star [LawfulHeap α] {Q : α → SLP α} : ((∃∃x, Q x) ⋆ P) = (∃∃x, P ⋆ Q x) := by rw [star_exists] simp [SLP.star_comm] @@ -95,10 +95,10 @@ theorem Lampe.SLP.skip_fst' : (⟦⟧ ⊢ Q ⋆ X) → ([a ↦ b] ⋆ X ⊢ R₂ assumption assumption -theorem Lampe.SLP.entails_star_true [SLH α] {H : SLP α} : H ⊢ H ⋆ ⟦⟧ := by +theorem Lampe.SLP.entails_star_true [LawfulHeap α] {H : SLP α} : H ⊢ H ⋆ ⟦⟧ := by simp [SLP.entails_self] -theorem SLP.eq_of_iff [SLH α] {P Q : SLP α} : (P ⊢ Q) → (Q ⊢ P) → P = Q := by +theorem SLP.eq_of_iff [LawfulHeap α] {P Q : SLP α} : (P ⊢ Q) → (Q ⊢ P) → P = Q := by intros apply funext intro @@ -108,14 +108,14 @@ theorem SLP.eq_of_iff [SLH α] {P Q : SLP α} : (P ⊢ Q) → (Q ⊢ P) → P = theorem pluck_pure_l {P : Prop} : ([a ↦ b] ⋆ P) = (P ⋆ [a ↦ b]) := by simp [SLP.star_comm] -theorem pluck_pure_all_l [SLH α] {P : Prop} {f : Prop → SLP α} : (SLP.forall' f ⋆ P) = (P ⋆ SLP.forall' f) := by +theorem pluck_pure_all_l [LawfulHeap α] {P : Prop} {f : Prop → SLP α} : (SLP.forall' f ⋆ P) = (P ⋆ SLP.forall' f) := by simp [SLP.star_comm] theorem pluck_pure_l_assoc {P : Prop} {Q : SLP (State p)} : ([a ↦ b] ⋆ P ⋆ Q) = (P ⋆ [a ↦ b] ⋆ Q) := by rw [SLP.star_comm, SLP.star_assoc] apply SLP.eq_of_iff <;> {apply SLP.star_mono_l; rw [SLP.star_comm]; apply SLP.entails_self} -theorem SLP.pure_star_pure [SLH α] {P Q : Prop} : (P ⋆ Q) = (⟦P ∧ Q⟧ : SLP α) := by +theorem SLP.pure_star_pure [LawfulHeap α] {P Q : Prop} : (P ⋆ Q) = (⟦P ∧ Q⟧ : SLP α) := by unfold SLP.star SLP.lift funext st apply eq_iff_iff.mpr @@ -125,8 +125,8 @@ theorem SLP.pure_star_pure [SLH α] {P Q : Prop} : (P ⋆ Q) = (⟦P ∧ Q⟧ : · intro_cases use ∅, ∅ refine ⟨?_, ?_, ?_, ?_, ?_⟩ - apply SLH.disjoint_empty - all_goals simp_all [SLH.disjoint_empty] + apply LawfulHeap.disjoint_empty + all_goals simp_all [LawfulHeap.disjoint_empty] macro "h_norm" : tactic => `(tactic|( try simp only [SLP.star_assoc, pluck_pure_l, pluck_pure_l_assoc, pluck_pure_all_l, SLP.star_true, SLP.true_star, star_exists, exists_star]; @@ -135,7 +135,7 @@ macro "h_norm" : tactic => `(tactic|( subst_vars; )) -theorem SLP.pure_leftX [SLH α] {H Q R : SLP α} : (P → (H ⊢ Q ⋆ R)) → (P ⋆ H ⊢ Q ⋆ P ⋆ R) := by +theorem SLP.pure_leftX [LawfulHeap α] {H Q R : SLP α} : (P → (H ⊢ Q ⋆ R)) → (P ⋆ H ⊢ Q ⋆ P ⋆ R) := by intro apply SLP.pure_left intro @@ -147,13 +147,13 @@ theorem SLP.pure_leftX [SLH α] {H Q R : SLP α} : (P → (H ⊢ Q ⋆ R)) → ( tauto /-- only finisher, will waste mvars for top! -/ -theorem SLP.pure_ent_star_top [SLH α] : (P → Q) → ((P : SLP α) ⊢ Q ⋆ ⊤) := by +theorem SLP.pure_ent_star_top [LawfulHeap α] : (P → Q) → ((P : SLP α) ⊢ Q ⋆ ⊤) := by intro h st hp rcases hp with ⟨_, rfl, hp⟩ use ∅, ∅ refine ⟨?_, ?_, ?_, ?_⟩ - apply SLH.disjoint_empty - all_goals simp_all [SLH.disjoint_empty, SLP.lift] + apply LawfulHeap.disjoint_empty + all_goals simp_all [LawfulHeap.disjoint_empty, SLP.lift] theorem star_mono_l_sing : (P ⊢ Q) → (v₁ = v₂) → ([r ↦ v₁] ⋆ P ⊢ [r ↦ v₂] ⋆ Q) := by intro h₁ h₂ @@ -224,25 +224,25 @@ partial def parseEntailment (e: Expr): TacticM (SLTerm × SLTerm) := do return (pre, post) else throwError "not an entailment" -theorem star_top_of_star_mvar [SLH α] {H Q R : SLP α} : (H ⊢ Q ⋆ R) → (H ⊢ Q ⋆ ⊤) := by +theorem star_top_of_star_mvar [LawfulHeap α] {H Q R : SLP α} : (H ⊢ Q ⋆ R) → (H ⊢ Q ⋆ ⊤) := by intro h apply SLP.entails_trans assumption apply SLP.star_mono_l apply SLP.entails_top -theorem solve_left_with_leftovers [SLH α] {H Q R : SLP α} : (H ⊢ Q ⋆ R) → (R ⊢ P) → (H ⊢ Q ⋆ P) := by +theorem solve_left_with_leftovers [LawfulHeap α] {H Q R : SLP α} : (H ⊢ Q ⋆ R) → (R ⊢ P) → (H ⊢ Q ⋆ P) := by intros apply SLP.entails_trans assumption apply SLP.star_mono_l assumption -theorem solve_with_true [SLH α] {H Q : SLP α}: (H ⊢ Q) → (H ⊢ Q ⋆ ⟦⟧) := by +theorem solve_with_true [LawfulHeap α] {H Q : SLP α}: (H ⊢ Q) → (H ⊢ Q ⋆ ⟦⟧) := by aesop -- partial def solveNonMVarEntailment (goal : MVarId) (lhs : SLTerm) (rhs : SLTerm): TacticM (List MVarId × SLTerm) := do -theorem pure_ent_pure_star_mv [SLH α] : (P → Q) → ((P : SLP α) ⊢ Q ⋆ ⟦⟧) := by +theorem pure_ent_pure_star_mv [LawfulHeap α] : (P → Q) → ((P : SLP α) ⊢ Q ⋆ ⟦⟧) := by intro h apply SLP.pure_left' intro @@ -250,7 +250,7 @@ theorem pure_ent_pure_star_mv [SLH α] : (P → Q) → ((P : SLP α) ⊢ Q ⋆ tauto tauto -theorem pure_star_H_ent_pure_star_mv [SLH α] {H Q R : SLP α} : (P → (H ⊢ Q ⋆ R)) → (P ⋆ H ⊢ Q ⋆ P ⋆ R) := by +theorem pure_star_H_ent_pure_star_mv [LawfulHeap α] {H Q R : SLP α} : (P → (H ⊢ Q ⋆ R)) → (P ⋆ H ⊢ Q ⋆ P ⋆ R) := by intro apply SLP.pure_left intro @@ -260,7 +260,7 @@ theorem pure_star_H_ent_pure_star_mv [SLH α] {H Q R : SLP α} : (P → (H ⊢ Q rw [SLP.star_comm] tauto -theorem skip_left_ent_star_mv [SLH α] {R L P H : SLP α} : (R ⊢ P ⋆ H) → (L ⋆ R ⊢ P ⋆ L ⋆ H) := by +theorem skip_left_ent_star_mv [LawfulHeap α] {R L P H : SLP α} : (R ⊢ P ⋆ H) → (L ⋆ R ⊢ P ⋆ L ⋆ H) := by intro h apply SLP.entails_trans apply SLP.star_mono_l @@ -270,7 +270,7 @@ theorem skip_left_ent_star_mv [SLH α] {R L P H : SLP α} : (R ⊢ P ⋆ H) → rw [SLP.star_comm] apply SLP.entails_self -theorem skip_evidence_pure [SLH α] {H : SLP α} : Q → (H ⊢ Q ⋆ H) := by +theorem skip_evidence_pure [LawfulHeap α] {H : SLP α} : Q → (H ⊢ Q ⋆ H) := by intro apply SLP.pure_right tauto @@ -295,17 +295,17 @@ theorem exi_prop {Q : P → SLP (State p)} : (H ⊢ (P : SLP (State p)) ⋆ ⊤) apply_assumption assumption -theorem exi_prop_l [SLH α] {H : P → SLP α} {Q : SLP α}: ((x : P) → ((P ⋆ H x) ⊢ Q)) → ((∃∃x, H x) ⊢ Q) := by +theorem exi_prop_l [LawfulHeap α] {H : P → SLP α} {Q : SLP α}: ((x : P) → ((P ⋆ H x) ⊢ Q)) → ((∃∃x, H x) ⊢ Q) := by intro h st unfold SLP.entails SLP.exists' at * rintro ⟨v, hH⟩ apply h use ∅, st refine ⟨?_, ?_, ?_, ?_⟩ - apply SLH.disjoint_empty - all_goals simp_all [SLH.disjoint_empty, SLP.lift] + apply LawfulHeap.disjoint_empty + all_goals simp_all [LawfulHeap.disjoint_empty, SLP.lift] -theorem use_right [SLH α] {R L G H : SLP α} : (R ⊢ G ⋆ H) → (L ⋆ R ⊢ G ⋆ L ⋆ H) := by +theorem use_right [LawfulHeap α] {R L G H : SLP α} : (R ⊢ G ⋆ H) → (L ⋆ R ⊢ G ⋆ L ⋆ H) := by intro apply SLP.entails_trans apply SLP.star_mono_l From 5aa0fdd0b6566ad91a988c059912baecda3763d9 Mon Sep 17 00:00:00 2001 From: utkn Date: Sat, 23 Nov 2024 19:24:12 +0300 Subject: [PATCH 17/39] Disjoint_union_right is now proven automatically with disjoint_union_left --- Lampe/Lampe/SeparationLogic/LawfulHeap.lean | 14 ++++++++++---- Lampe/Lampe/SeparationLogic/State.lean | 18 ------------------ Lampe/Lampe/SeparationLogic/ValHeap.lean | 1 - 3 files changed, 10 insertions(+), 23 deletions(-) diff --git a/Lampe/Lampe/SeparationLogic/LawfulHeap.lean b/Lampe/Lampe/SeparationLogic/LawfulHeap.lean index d1035dc..0acb084 100644 --- a/Lampe/Lampe/SeparationLogic/LawfulHeap.lean +++ b/Lampe/Lampe/SeparationLogic/LawfulHeap.lean @@ -11,8 +11,7 @@ class LawfulHeap (α : Type _) where union_assoc {s₁ s₂ s₃ : α} : union (union s₁ s₂) s₃ = union s₁ (union s₂ s₃) disjoint_symm_iff {a b : α} : disjoint a b ↔ disjoint b a union_comm_of_disjoint {s₁ s₂ : α} : disjoint s₁ s₂ → union s₁ s₂ = union s₂ s₁ - disjoint_union_left (x y z : α) : disjoint (union x y) z ↔ disjoint x z ∧ disjoint y z -- one of them should be enough - disjoint_union_right (x y z : α) : disjoint x (union y z) ↔ disjoint x y ∧ disjoint x z + disjoint_union_left (x y z : α) : disjoint (union x y) z ↔ disjoint x z ∧ disjoint y z disjoint_empty (x : α) : disjoint empty x instance [LawfulHeap α] : Union α := ⟨LawfulHeap.union⟩ @@ -52,7 +51,14 @@ theorem LawfulHeap_disjoint_union_left [LawfulHeap α] (x y z : α) : theorem LawfulHeap_disjoint_union_right [LawfulHeap α] (x y z : α) : LawfulHeap.disjoint x (y ∪ z) ↔ LawfulHeap.disjoint x y ∧ LawfulHeap.disjoint x z := by - simp [Union.union] - apply LawfulHeap.disjoint_union_right + conv => + lhs + rw [LawfulHeap.disjoint_symm_iff] + conv => + rhs + rw [LawfulHeap.disjoint_symm_iff] + rhs + rw [LawfulHeap.disjoint_symm_iff] + apply LawfulHeap.disjoint_union_left end Lampe diff --git a/Lampe/Lampe/SeparationLogic/State.lean b/Lampe/Lampe/SeparationLogic/State.lean index 9bc7964..4ff78f1 100644 --- a/Lampe/Lampe/SeparationLogic/State.lean +++ b/Lampe/Lampe/SeparationLogic/State.lean @@ -86,24 +86,6 @@ instance : LawfulHeap (State p) where constructor tauto tauto - disjoint_union_right := by - intros x y z - have h1 := (Finmap.disjoint_union_right x.vals y.vals z.vals) - have h2 := (Finmap.disjoint_union_right x.closures y.closures z.closures) - constructor - intro h3 - simp only [Union.union] at h3 - constructor - constructor - cases h3 - tauto - tauto - tauto - intro h3 - simp only [Union.union] at h3 - constructor - tauto - tauto @[reducible] def State.valSingleton (r : Ref) (v : AnyValue p) : SLP (State p) := fun st => st.vals = Finmap.singleton r v diff --git a/Lampe/Lampe/SeparationLogic/ValHeap.lean b/Lampe/Lampe/SeparationLogic/ValHeap.lean index b34f75d..7a59596 100644 --- a/Lampe/Lampe/SeparationLogic/ValHeap.lean +++ b/Lampe/Lampe/SeparationLogic/ValHeap.lean @@ -38,7 +38,6 @@ instance : LawfulHeap (ValHeap p) where disjoint_symm_iff := by tauto union_comm_of_disjoint := Finmap.union_comm_of_disjoint disjoint_union_left := Finmap.disjoint_union_left - disjoint_union_right := Finmap.disjoint_union_right disjoint_empty := Finmap.disjoint_empty -- def ValHeap.singleton (r : Ref) (v : AnyValue p) : SLP (ValHeap p) := fun st => st = Finmap.singleton r v From a5b6342b11b213db9760b88f06880c85f18ecdc6 Mon Sep 17 00:00:00 2001 From: utkn Date: Sat, 23 Nov 2024 22:49:40 +0300 Subject: [PATCH 18/39] some tactic fixes --- Lampe/Lampe.lean | 5 +++-- Lampe/Lampe/Hoare/SepTotal.lean | 6 +++--- Lampe/Lampe/SeparationLogic/State.lean | 15 +++----------- Lampe/Lampe/Tactic/SeparationLogic.lean | 26 +++++++++++++------------ 4 files changed, 23 insertions(+), 29 deletions(-) diff --git a/Lampe/Lampe.lean b/Lampe/Lampe.lean index 97150ad..5372a18 100644 --- a/Lampe/Lampe.lean +++ b/Lampe/Lampe.lean @@ -11,7 +11,8 @@ nr_def simple_muts<>(x : Field) -> Field { example : STHoare p Γ ⟦⟧ (simple_muts.fn.body _ h![] h![x]) fun v => v = x := by simp only [simple_muts] - steps <;> try (exact _) + steps + try (exact _) simp_all nr_def weirdEq(x : I, y : I) -> Unit { @@ -89,7 +90,7 @@ example {p Γ x y}: STHoare p Γ ⟦⟧ (simple_if_else.fn.body (Tp.denote p) h! . simp only [decide_True, exists_const] sl contradiction - . exact _ + . aesop nr_def simple_lambda<>(x : Field) -> Field { let foo = |a|: Field -> Field { a }; diff --git a/Lampe/Lampe/Hoare/SepTotal.lean b/Lampe/Lampe/Hoare/SepTotal.lean index 00e6211..faff898 100644 --- a/Lampe/Lampe/Hoare/SepTotal.lean +++ b/Lampe/Lampe/Hoare/SepTotal.lean @@ -6,12 +6,12 @@ import Lampe.Hoare.Total namespace Lampe /-- -A Hoare triple `{P} e {λv. Q v}` has the following meaning: +A Hoare triple `{P} e {λv ↦ Q v}` has the following meaning: if the state of the program satisfies the pre-condition `P`, then the expression `e` terminates and evaluates to `v` such that the post-condition `Q v` holds. Note that even if `e` terminates, it may evaluate to an error, e.g., division-by-zero. -Accordingly, we interpret `{P} e {λv. Q v}` as follows: +Accordingly, we interpret `{P} e {λv ↦ Q v}` as follows: if `P` holds, then `e` terminates and it either (1) fails or (2) *successfully* evaluates to `v` such that `Q v` holds. This is logically equivalent to saying that if `P` holds, then `e` terminates and (`Q v` holds if `e` succeeds and evaluates to `v`). @@ -19,7 +19,7 @@ Hence, the triples are *partial* with respect to failure and *total* with respec An intuitive way of looking at this is thinking in terms of "knowledge discovery". For example, if the operation `a + b` succeeds, then we know that it evaluates to `v = a + b` **and** `a + b < 2^32`, i.e., no overflow has happened. -Then, we would define the post-condition such that `Q v = (v = a + b) ∧ (a + b < 2^32)`. +Then, we would define the post-condition such that `Q = λv ↦ (v = a + b) ∧ (a + b < 2^32)`. -/ def STHoare p Γ P e (Q : Tp.denote p tp → SLP (State p)) := ∀H, THoare p Γ (P ⋆ H) e (fun v => ((Q v) ⋆ H) ⋆ ⊤) diff --git a/Lampe/Lampe/SeparationLogic/State.lean b/Lampe/Lampe/SeparationLogic/State.lean index 4ff78f1..0c81dd7 100644 --- a/Lampe/Lampe/SeparationLogic/State.lean +++ b/Lampe/Lampe/SeparationLogic/State.lean @@ -13,6 +13,7 @@ structure State (p : Prime) where instance : Membership Ref (State p) where mem := fun a e => e ∈ a.vals +@[simp] lemma State.membership_in_val {a : State p} : e ∈ a ↔ e ∈ a.vals := by rfl instance : Coe (State p) (ValHeap p) where @@ -44,7 +45,6 @@ lemma State.eq_closures : intro h injection h - instance : LawfulHeap (State p) where union := fun a b => ⟨a.vals ∪ b.vals, a.closures ∪ b.closures⟩ disjoint := fun a b => a.vals.Disjoint b.vals ∧ a.closures.Disjoint b.closures @@ -75,17 +75,8 @@ instance : LawfulHeap (State p) where constructor intro h3 simp only [Union.union] at h3 - constructor - constructor - cases h3 - tauto - tauto - tauto - intro h3 - simp only [Union.union] at h3 - constructor - tauto - tauto + constructor <;> constructor + all_goals tauto @[reducible] def State.valSingleton (r : Ref) (v : AnyValue p) : SLP (State p) := fun st => st.vals = Finmap.singleton r v diff --git a/Lampe/Lampe/Tactic/SeparationLogic.lean b/Lampe/Lampe/Tactic/SeparationLogic.lean index 223a751..421b48d 100644 --- a/Lampe/Lampe/Tactic/SeparationLogic.lean +++ b/Lampe/Lampe/Tactic/SeparationLogic.lean @@ -179,8 +179,8 @@ def isLetIn (e: Expr): Bool := partial def parseSLExpr (e: Expr): TacticM SLTerm := do if e.isAppOf ``SLP.star then let args := e.getAppArgs - let fst ← parseSLExpr (←args[1]?) - let snd ← parseSLExpr (←args[2]?) + let fst ← parseSLExpr (←args[2]?) + let snd ← parseSLExpr (←args[3]?) return SLTerm.star e fst snd if e.isAppOf ``State.valSingleton then let args := e.getAppArgs @@ -191,19 +191,19 @@ partial def parseSLExpr (e: Expr): TacticM SLTerm := do return SLTerm.top else if e.isAppOf ``SLP.lift then let args := e.getAppArgs - return SLTerm.lift (←args[1]?) + return SLTerm.lift (←args[2]?) else if e.getAppFn.isMVar then return SLTerm.mvar e else if e.isAppOf ``SLP.forall' then let args := e.getAppArgs - return SLTerm.all (←args[2]?) + return SLTerm.all (←args[3]?) else if e.isAppOf ``SLP.exists' then let args := e.getAppArgs - return SLTerm.exi (←args[2]?) + return SLTerm.exi (←args[4]?) else if e.isAppOf ``SLP.wand then let args := e.getAppArgs - let lhs ← parseSLExpr (←args[1]?) - let rhs ← parseSLExpr (←args[2]?) + let lhs ← parseSLExpr (←args[2]?) + let rhs ← parseSLExpr (←args[3]?) return SLTerm.wand lhs rhs -- else if e.isAppOf ``SLTerm.lift then -- let args := e.getAppArgs @@ -219,10 +219,10 @@ partial def parseSLExpr (e: Expr): TacticM SLTerm := do partial def parseEntailment (e: Expr): TacticM (SLTerm × SLTerm) := do if e.isAppOf ``SLP.entails then let args := e.getAppArgs - let pre ← parseSLExpr (←args[1]?) - let post ← parseSLExpr (←args[2]?) + let pre ← parseSLExpr (←args[2]?) + let post ← parseSLExpr (←args[3]?) return (pre, post) - else throwError "not an entailment" + else throwError "not an entailment {e}" theorem star_top_of_star_mvar [LawfulHeap α] {H Q R : SLP α} : (H ⊢ Q ⋆ R) → (H ⊢ Q ⋆ ⊤) := by intro h @@ -284,7 +284,8 @@ theorem SLP.exists_intro {Q : α → SLP (State p)} {a} : (H ⊢ Q a) → (H ⊢ tauto -- [TODO] make this generic over `LawfulHeap`s -theorem exi_prop {Q : P → SLP (State p)} : (H ⊢ (P : SLP (State p)) ⋆ ⊤) → (∀(p : P), H ⊢ Q p) → (H ⊢ ∃∃p, Q p) := by +theorem exi_prop {Q : P → SLP (State p)} : + (H ⊢ (P : SLP (State p)) ⋆ ⊤) → (∀(p : P), H ⊢ Q p) → (H ⊢ ∃∃p, Q p) := by intro h₁ h₂ unfold SLP.entails at * intro st hH @@ -295,7 +296,8 @@ theorem exi_prop {Q : P → SLP (State p)} : (H ⊢ (P : SLP (State p)) ⋆ ⊤) apply_assumption assumption -theorem exi_prop_l [LawfulHeap α] {H : P → SLP α} {Q : SLP α}: ((x : P) → ((P ⋆ H x) ⊢ Q)) → ((∃∃x, H x) ⊢ Q) := by +theorem exi_prop_l [LawfulHeap α] {H : P → SLP α} {Q : SLP α} : + ((x : P) → ((P ⋆ H x) ⊢ Q)) → ((∃∃x, H x) ⊢ Q) := by intro h st unfold SLP.entails SLP.exists' at * rintro ⟨v, hH⟩ From e3b2604ca5252580d3ed4d77a5cdeabf721df815 Mon Sep 17 00:00:00 2001 From: utkn Date: Sun, 24 Nov 2024 17:34:40 +0300 Subject: [PATCH 19/39] Tactics fixed. Build passing --- Lampe/Lampe/SeparationLogic/SLP.lean | 2 +- Lampe/Lampe/Tactic/SeparationLogic.lean | 19 +++++++++---------- 2 files changed, 10 insertions(+), 11 deletions(-) diff --git a/Lampe/Lampe/SeparationLogic/SLP.lean b/Lampe/Lampe/SeparationLogic/SLP.lean index dfc266a..552243b 100644 --- a/Lampe/Lampe/SeparationLogic/SLP.lean +++ b/Lampe/Lampe/SeparationLogic/SLP.lean @@ -7,7 +7,7 @@ def SLP (α) [LawfulHeap α] := α → Prop namespace SLP -def star [LawfulHeap α] (lhs rhs : SLP α) := fun st => +def star [LawfulHeap α] (lhs rhs : SLP α) : SLP α := fun st => ∃ st₁ st₂, LawfulHeap.disjoint st₁ st₂ ∧ st = st₁ ∪ st₂ ∧ lhs st₁ ∧ rhs st₂ def lift [LawfulHeap α] (pr : Prop) : SLP α := fun st => pr ∧ st = ∅ diff --git a/Lampe/Lampe/Tactic/SeparationLogic.lean b/Lampe/Lampe/Tactic/SeparationLogic.lean index 421b48d..6c2da1d 100644 --- a/Lampe/Lampe/Tactic/SeparationLogic.lean +++ b/Lampe/Lampe/Tactic/SeparationLogic.lean @@ -43,13 +43,13 @@ instance : ToString SLTerm := ⟨SLTerm.toString⟩ instance : Inhabited SLTerm := ⟨SLTerm.top⟩ -theorem star_exists [LawfulHeap α] {Q : α → SLP α} : ((∃∃x, Q x) ⋆ P) = (∃∃x, Q x ⋆ P) := by +theorem star_exists [LawfulHeap α] {P : SLP α} {Q : β → SLP α} : ((∃∃x, Q x) ⋆ P) = (∃∃x, Q x ⋆ P) := by unfold SLP.exists' SLP.star funext st simp tauto -theorem exists_star [LawfulHeap α] {Q : α → SLP α} : ((∃∃x, Q x) ⋆ P) = (∃∃x, P ⋆ Q x) := by +theorem exists_star [LawfulHeap α] {P : SLP α} {Q : β → SLP α} : ((∃∃x, Q x) ⋆ P) = (∃∃x, P ⋆ Q x) := by rw [star_exists] simp [SLP.star_comm] @@ -199,7 +199,7 @@ partial def parseSLExpr (e: Expr): TacticM SLTerm := do return SLTerm.all (←args[3]?) else if e.isAppOf ``SLP.exists' then let args := e.getAppArgs - return SLTerm.exi (←args[4]?) + return SLTerm.exi (←args[3]?) else if e.isAppOf ``SLP.wand then let args := e.getAppArgs let lhs ← parseSLExpr (←args[2]?) @@ -250,7 +250,8 @@ theorem pure_ent_pure_star_mv [LawfulHeap α] : (P → Q) → ((P : SLP α) ⊢ tauto tauto -theorem pure_star_H_ent_pure_star_mv [LawfulHeap α] {H Q R : SLP α} : (P → (H ⊢ Q ⋆ R)) → (P ⋆ H ⊢ Q ⋆ P ⋆ R) := by +theorem pure_star_H_ent_pure_star_mv [LawfulHeap α] {H Q R : SLP α} : + (P → (H ⊢ Q ⋆ R)) → (P ⋆ H ⊢ Q ⋆ P ⋆ R) := by intro apply SLP.pure_left intro @@ -276,16 +277,14 @@ theorem skip_evidence_pure [LawfulHeap α] {H : SLP α} : Q → (H ⊢ Q ⋆ H) tauto tauto --- [TODO] make this generic over `LawfulHeap`s -theorem SLP.exists_intro {Q : α → SLP (State p)} {a} : (H ⊢ Q a) → (H ⊢ ∃∃a, Q a) := by +theorem SLP.exists_intro [LawfulHeap α] {H : SLP α} {Q : β → SLP α} {a} : (H ⊢ Q a) → (H ⊢ ∃∃a, Q a) := by intro h st H unfold SLP.exists' exists a tauto --- [TODO] make this generic over `LawfulHeap`s -theorem exi_prop {Q : P → SLP (State p)} : - (H ⊢ (P : SLP (State p)) ⋆ ⊤) → (∀(p : P), H ⊢ Q p) → (H ⊢ ∃∃p, Q p) := by +theorem exi_prop [LawfulHeap α] {P : Prop} {H : SLP α} {Q : P → SLP α} : + (H ⊢ P ⋆ ⊤) → (∀(p : P), H ⊢ Q p) → (H ⊢ ∃∃p, Q p) := by intro h₁ h₂ unfold SLP.entails at * intro st hH @@ -296,7 +295,7 @@ theorem exi_prop {Q : P → SLP (State p)} : apply_assumption assumption -theorem exi_prop_l [LawfulHeap α] {H : P → SLP α} {Q : SLP α} : +theorem exi_prop_l [LawfulHeap α] {P : Prop} {H : P → SLP α} {Q : SLP α} : ((x : P) → ((P ⋆ H x) ⊢ Q)) → ((∃∃x, H x) ⊢ Q) := by intro h st unfold SLP.entails SLP.exists' at * From 944def95f7e434c8884a01a82fcaaf14aff58b3c Mon Sep 17 00:00:00 2001 From: utkn Date: Mon, 25 Nov 2024 11:20:24 +0300 Subject: [PATCH 20/39] callLambda intro --- Lampe/Lampe/Hoare/Builtins.lean | 6 ---- Lampe/Lampe/Hoare/SepTotal.lean | 41 ++++++++++++++++++++++++++ Lampe/Lampe/SeparationLogic/State.lean | 3 +- 3 files changed, 42 insertions(+), 8 deletions(-) diff --git a/Lampe/Lampe/Hoare/Builtins.lean b/Lampe/Lampe/Hoare/Builtins.lean index b8c568d..d234f41 100644 --- a/Lampe/Lampe/Hoare/Builtins.lean +++ b/Lampe/Lampe/Hoare/Builtins.lean @@ -442,10 +442,4 @@ theorem assert_intro : STHoarePureBuiltin p Γ Builtin.assert (by tauto) h![a] ( apply pureBuiltin_intro_consequence <;> tauto tauto -theorem newLambda_intro {cl : Closures} : - STHoare p Γ P (.lambda argTps outTp body) Q := by - unfold STHoare - intro H - sorry - end Lampe.STHoare diff --git a/Lampe/Lampe/Hoare/SepTotal.lean b/Lampe/Lampe/Hoare/SepTotal.lean index faff898..8d6d131 100644 --- a/Lampe/Lampe/Hoare/SepTotal.lean +++ b/Lampe/Lampe/Hoare/SepTotal.lean @@ -306,4 +306,45 @@ theorem skip_intro : . apply SLP.ent_star_top tauto +theorem callLambda_intro {fn : Function} : + (hg : fn.generics = []) → + (hi : fn.inTps (hg ▸ h![]) = argTps) → + (ho : fn.outTp (hg ▸ h![]) = outTp) → + STHoare p Γ ⟦⟧ (ho ▸ fn.body _ (hg ▸ h![]) (hi ▸ args)) (fun v => v = v') → + STHoare p Γ [lambdaRef ↣ fn] (.call h![] argTps outTp (.lambda lambdaRef) args) + (fun v => ⟦v = v'⟧ ⋆ [lambdaRef ↣ fn]) := by + intros + rename_i h + unfold STHoare THoare + intros + constructor <;> tauto + unfold SLP.star at * + rotate_right 1 + . rename_i st h + exact st.closures + . rename_i st h + obtain ⟨st₁, ⟨st₂, ⟨_, h₂, h₃, _⟩⟩⟩ := h + simp only [State.union_parts, h₃] at h₂ + simp only [h₂] + simp_all + . apply consequence + <;> tauto + apply consequence_frame_left + rotate_left 2 + exact ⟦⟧ + exact h + simp only [SLP.true_star, SLP.entails_self] + + +theorem newLambda_intro : + STHoare p Γ P (.lambda argTps outTp body) Q := by + unfold STHoare + intro H + sorry + + + + + + end Lampe.STHoare diff --git a/Lampe/Lampe/SeparationLogic/State.lean b/Lampe/Lampe/SeparationLogic/State.lean index 0c81dd7..d7d4970 100644 --- a/Lampe/Lampe/SeparationLogic/State.lean +++ b/Lampe/Lampe/SeparationLogic/State.lean @@ -86,8 +86,7 @@ notation:max "[" l " ↦ " r "]" => State.valSingleton l r @[reducible] def State.clsSingleton (r : Ref) (v : Function) : SLP (State p) := fun st => st.closures = Finmap.singleton r v -notation:max "[" l " ↣ " r "]" => State.cls l r - +notation:max "[" l " ↣ " r "]" => State.clsSingleton l r @[simp] lemma State.union_parts_left : From cf80adb0b8af1cb4d00ecab5ff1bb3dabec5de25 Mon Sep 17 00:00:00 2001 From: utkn Date: Mon, 25 Nov 2024 18:04:54 +0300 Subject: [PATCH 21/39] newLambda Omni frame proven, callLambda intro rule added --- Lampe/Lampe/Ast.lean | 14 ++---- Lampe/Lampe/Hoare/SepTotal.lean | 45 ++++++++++++------- Lampe/Lampe/Semantics.lean | 56 ++++++++++++++++++++---- Lampe/Lampe/SeparationLogic/State.lean | 10 +++-- Lampe/Lampe/SeparationLogic/ValHeap.lean | 1 + Lampe/Lampe/Tactic/SeparationLogic.lean | 3 ++ 6 files changed, 90 insertions(+), 39 deletions(-) diff --git a/Lampe/Lampe/Ast.lean b/Lampe/Lampe/Ast.lean index a9c7ad2..b74d1f9 100644 --- a/Lampe/Lampe/Ast.lean +++ b/Lampe/Lampe/Ast.lean @@ -29,7 +29,7 @@ inductive Expr (rep : Tp → Type) : Tp → Type where | ite : rep .bool → Expr rep a → Expr rep a → Expr rep a | skip : Expr rep .unit | loop : rep (.u s) → rep (.u s) → (rep (.u s) → Expr rep r) → Expr rep .unit -| lambda : (argTps : List Tp) → (outTp : Tp) → (HList rep argTps → Expr rep outTp) → Expr rep .lambdaRef +| lambda : ((argTps : List Tp) → (outTp : Tp) → (HList rep argTps → Expr rep outTp)) → Expr rep .lambdaRef structure Function : Type _ where generics : List Kind @@ -47,16 +47,8 @@ example : Function := { body := fun _ h![_] h![x] => .var x } -abbrev Lambda (argTps outTp) := - (rep : Tp → Type) → HList Kind.denote [] → HList rep argTps → Expr rep outTp - -def newLambda (argTps : List Tp) (outTp : Tp) -(body : Lambda argTps outTp) : Function := { - generics := [] - inTps := fun _ => argTps - outTp := fun _ => outTp - body := body -} +def Lambda : Type _ := + (rep : Tp → Type) → (argTps : List Tp) → (outTp : Tp) → HList rep argTps → Expr rep outTp structure FunctionDecl where name : Ident diff --git a/Lampe/Lampe/Hoare/SepTotal.lean b/Lampe/Lampe/Hoare/SepTotal.lean index 8d6d131..48e7eb0 100644 --- a/Lampe/Lampe/Hoare/SepTotal.lean +++ b/Lampe/Lampe/Hoare/SepTotal.lean @@ -306,22 +306,16 @@ theorem skip_intro : . apply SLP.ent_star_top tauto -theorem callLambda_intro {fn : Function} : - (hg : fn.generics = []) → - (hi : fn.inTps (hg ▸ h![]) = argTps) → - (ho : fn.outTp (hg ▸ h![]) = outTp) → - STHoare p Γ ⟦⟧ (ho ▸ fn.body _ (hg ▸ h![]) (hi ▸ args)) (fun v => v = v') → - STHoare p Γ [lambdaRef ↣ fn] (.call h![] argTps outTp (.lambda lambdaRef) args) - (fun v => ⟦v = v'⟧ ⋆ [lambdaRef ↣ fn]) := by +theorem callLambda_intro : + @STHoare outTp p Γ ⟦⟧ (lambdaBody (Tp.denote p) argTps outTp args) (fun v => v = v') → + STHoare p Γ [ref ↣ lambdaBody] (.call h![] argTps outTp (.lambda ref) args) + (fun v => ⟦v = v'⟧ ⋆ [ref ↣ lambdaBody]) := by intros rename_i h unfold STHoare THoare intros constructor <;> tauto unfold SLP.star at * - rotate_right 1 - . rename_i st h - exact st.closures . rename_i st h obtain ⟨st₁, ⟨st₂, ⟨_, h₂, h₃, _⟩⟩⟩ := h simp only [State.union_parts, h₃] at h₂ @@ -336,11 +330,32 @@ theorem callLambda_intro {fn : Function} : simp only [SLP.true_star, SLP.entails_self] -theorem newLambda_intro : - STHoare p Γ P (.lambda argTps outTp body) Q := by - unfold STHoare - intro H - sorry +theorem newLambda_intro {r : Ref} : + STHoare p Γ ⟦⟧ (.lambda (lambda _)) + fun v => ⟦v = r⟧ ⋆ [r ↣ lambda] := by + unfold STHoare THoare + intros H st h + obtain ⟨s₁, ⟨s₂, ⟨h₁, h₂, h₃, h₄⟩⟩⟩ := h + simp only [LawfulHeap.disjoint] at h₁ + simp only [State.union_parts, State.eq_constructor, State.mk.injEq] at h₂ + simp only at h₂ + injection h₂ + rename_i hi₁ hi₂ + obtain ⟨hd₁, hd₂⟩ := h₁ + constructor <;> tauto + all_goals sorry + + + + + + + + + + + + diff --git a/Lampe/Lampe/Semantics.lean b/Lampe/Lampe/Semantics.lean index 0241109..f92bd13 100644 --- a/Lampe/Lampe/Semantics.lean +++ b/Lampe/Lampe/Semantics.lean @@ -45,12 +45,12 @@ inductive Omni : (p : Prime) → Env → State p → Expr (Tp.denote p) tp → ( Omni p Γ st (htco ▸ fn.body _ (hkc ▸ generics) (htci ▸ args)) Q → Omni p Γ st (@Expr.call _ tyKinds generics argTypes res (.decl fname) args) Q | callLambda {cls : Closures} : - cls.lookup lambdaRef = some fn → - (hg : fn.generics = []) → - (hi : fn.inTps (hg ▸ h![]) = argTps) → - (ho : fn.outTp (hg ▸ h![]) = outTp) → - Omni p Γ st (ho ▸ fn.body _ (hg ▸ h![]) (hi ▸ args)) Q → - Omni p Γ st (Expr.call h![] argTps outTp (.lambda lambdaRef) args) Q + cls.lookup ref = some lambda → + Omni p Γ ⟨vh, cls⟩ (lambda (Tp.denote p) argTps outTp args) Q → + Omni p Γ ⟨vh, cls⟩ (Expr.call h![] argTps outTp (.lambda ref) args) Q +| newLambda {Q} : + (∀ ref, ref ∉ cls → Q (some (⟨vh, cls.insert ref lambda⟩, ref))) → + Omni p Γ ⟨vh, cls⟩ (@Expr.lambda (Tp.denote p) (lambda _)) Q | loopDone : lo ≥ hi → Omni p Γ st (.loop lo hi body) Q @@ -156,7 +156,7 @@ theorem Omni.frame {p Γ tp} {st₁ st₂ : State p} {e : Expr (Tp.denote p) tp} . simp only [State.union_parts, State.mk.injEq] tauto . rw [hin₄] - | callDecl _ _ _ _ _ ih => + | callDecl => intro constructor all_goals (try assumption) @@ -175,8 +175,46 @@ theorem Omni.frame {p Γ tp} {st₁ st₂ : State p} {e : Expr (Tp.denote p) tp} constructor apply ih tauto - | callLambda => - intro + | callLambda h _ _ => + intro hd constructor <;> try tauto + simp_all + simp only [LawfulHeap.disjoint] at hd + simp only [Finmap.lookup_union_left (Finmap.mem_of_lookup_eq_some h)] + tauto + | newLambda => + intros h + simp only [LawfulHeap.disjoint, State.union_parts_left] at * + obtain ⟨_, hd⟩ := h + constructor + intros + simp only + rename Closures => cls + rename ValHeap _ => vh + rename Ref => r + rename Lambda => lambda + have hi : r ∉ cls ∧ r ∉ st₂.closures := by aesop + have hd₁ : Finmap.Disjoint cls (Finmap.singleton r lambda) := by + apply Finmap.Disjoint.symm + apply Finmap.singleton_disjoint_iff_not_mem.mpr + tauto + have hd₂ : Finmap.Disjoint (Finmap.singleton r lambda) st₂.closures := by + apply Finmap.singleton_disjoint_iff_not_mem.mpr + tauto + exists ⟨vh, cls ∪ Finmap.singleton r lambda⟩, st₂ + refine ⟨?_, ?_, ?_, by tauto⟩ + . simp only [LawfulHeap.disjoint] + refine ⟨by tauto, ?_⟩ + simp only [Finmap.disjoint_union_left] + tauto + . simp only [State.union_parts_left, State.mk.injEq] + refine ⟨by tauto, ?_⟩ + simp only [Finmap.insert_eq_singleton_union] + simp only [Finmap.union_comm_of_disjoint hd₁, Finmap.union_assoc] + . simp [Finmap.union_comm_of_disjoint, Finmap.insert_eq_singleton_union] + rename (∀ref ∉ cls, _) => hQ + simp only [Finmap.insert_eq_singleton_union] at hQ + rw [Finmap.union_comm_of_disjoint hd₁] + tauto end Lampe diff --git a/Lampe/Lampe/SeparationLogic/State.lean b/Lampe/Lampe/SeparationLogic/State.lean index d7d4970..8eff141 100644 --- a/Lampe/Lampe/SeparationLogic/State.lean +++ b/Lampe/Lampe/SeparationLogic/State.lean @@ -4,7 +4,7 @@ import Lampe.Ast namespace Lampe -abbrev Closures := Finmap fun _ : Ref => Lampe.Function +abbrev Closures := Finmap fun _ : Ref => Lambda structure State (p : Prime) where vals : ValHeap p @@ -22,7 +22,7 @@ instance : Coe (State p) (ValHeap p) where /-- Maps a post-condition on `State`s to a post-condition on `ValHeap`s by keeping the closures fixed -/ @[reducible] def mapToValHeapCondition - (closures : Finmap fun _ : Ref => Function) + (closures : Closures) (Q : Option (State p × T) → Prop) : Option (ValHeap p × T) → Prop := fun vv => Q (vv.map (fun (vals, t) => ⟨⟨vals, closures⟩, t⟩)) @@ -79,12 +79,14 @@ instance : LawfulHeap (State p) where all_goals tauto @[reducible] -def State.valSingleton (r : Ref) (v : AnyValue p) : SLP (State p) := fun st => st.vals = Finmap.singleton r v +def State.valSingleton (r : Ref) (v : AnyValue p) : SLP (State p) := + fun st => st.vals = Finmap.singleton r v notation:max "[" l " ↦ " r "]" => State.valSingleton l r @[reducible] -def State.clsSingleton (r : Ref) (v : Function) : SLP (State p) := fun st => st.closures = Finmap.singleton r v +def State.clsSingleton (r : Ref) (v : Lambda) : SLP (State p) := + fun st => st.closures = Finmap.singleton r v notation:max "[" l " ↣ " r "]" => State.clsSingleton l r diff --git a/Lampe/Lampe/SeparationLogic/ValHeap.lean b/Lampe/Lampe/SeparationLogic/ValHeap.lean index 7a59596..c4057b4 100644 --- a/Lampe/Lampe/SeparationLogic/ValHeap.lean +++ b/Lampe/Lampe/SeparationLogic/ValHeap.lean @@ -9,6 +9,7 @@ lemma Finmap.singleton_disjoint_of_not_mem (hp : ref ∉ s): Finmap.Disjoint (Finmap.singleton ref v) s := by simp_all [Finmap.Disjoint] + lemma Finmap.singleton_disjoint_iff_not_mem : Finmap.Disjoint (Finmap.singleton ref v) s ↔ (ref ∉ s) := by simp_all [Finmap.Disjoint] diff --git a/Lampe/Lampe/Tactic/SeparationLogic.lean b/Lampe/Lampe/Tactic/SeparationLogic.lean index 6c2da1d..2d3e9c0 100644 --- a/Lampe/Lampe/Tactic/SeparationLogic.lean +++ b/Lampe/Lampe/Tactic/SeparationLogic.lean @@ -464,6 +464,7 @@ macro "stephelper1" : tactic => `(tactic|( | apply fresh_intro | apply assert_intro | apply skip_intro + | apply callLambda_intro -- memory builtins | apply var_intro | apply ref_intro @@ -515,6 +516,7 @@ macro "stephelper2" : tactic => `(tactic|( | apply consequence_frame_left fresh_intro | apply consequence_frame_left Lampe.STHoare.litU_intro | apply consequence_frame_left assert_intro + | apply consequence_frame_left callLambda_intro -- | apply consequence_frame_left skip_intro -- memory builtins | apply consequence_frame_left var_intro @@ -569,6 +571,7 @@ macro "stephelper3" : tactic => `(tactic|( | apply ramified_frame_top Lampe.STHoare.litU_intro | apply ramified_frame_top assert_intro | apply ramified_frame_top skip_intro + | apply ramified_frame_top callLambda_intro -- memory builtins | apply ramified_frame_top var_intro | apply ramified_frame_top ref_intro From 7bc25ac198fd008f3d3cce801604e9ce80f1ff26 Mon Sep 17 00:00:00 2001 From: utkn Date: Mon, 25 Nov 2024 18:31:04 +0300 Subject: [PATCH 22/39] callLambda intro rule --- Lampe/Lampe/Hoare/SepTotal.lean | 52 ++++++++++++++------------------- 1 file changed, 22 insertions(+), 30 deletions(-) diff --git a/Lampe/Lampe/Hoare/SepTotal.lean b/Lampe/Lampe/Hoare/SepTotal.lean index 48e7eb0..fa2baa1 100644 --- a/Lampe/Lampe/Hoare/SepTotal.lean +++ b/Lampe/Lampe/Hoare/SepTotal.lean @@ -329,37 +329,29 @@ theorem callLambda_intro : exact h simp only [SLP.true_star, SLP.entails_self] - -theorem newLambda_intro {r : Ref} : +theorem newLambda_intro : STHoare p Γ ⟦⟧ (.lambda (lambda _)) - fun v => ⟦v = r⟧ ⋆ [r ↣ lambda] := by + fun v => [v ↣ lambda] := by unfold STHoare THoare intros H st h - obtain ⟨s₁, ⟨s₂, ⟨h₁, h₂, h₃, h₄⟩⟩⟩ := h - simp only [LawfulHeap.disjoint] at h₁ - simp only [State.union_parts, State.eq_constructor, State.mk.injEq] at h₂ - simp only at h₂ - injection h₂ - rename_i hi₁ hi₂ - obtain ⟨hd₁, hd₂⟩ := h₁ - constructor <;> tauto - all_goals sorry - - - - - - - - - - - - - - - - - - + constructor + intros + simp_all only [SLP.true_star, SLP.star_assoc] + rename Ref => r + exists ⟨∅, Finmap.singleton r lambda⟩, st + refine ⟨?_, ?_, ?_, ?_⟩ + . simp only [LawfulHeap.disjoint] + refine ⟨?_, ?_⟩ + . apply Finmap.disjoint_empty + . apply Finmap.singleton_disjoint_of_not_mem (by tauto) + . simp only [State.union_parts, State.mk.injEq] + refine ⟨by simp_all, ?_⟩ + have hd : Finmap.Disjoint st.closures (Finmap.singleton r lambda) := by + rw [Finmap.Disjoint.symm_iff] + apply Finmap.singleton_disjoint_of_not_mem (by assumption) + simp only [Finmap.insert_eq_singleton_union, Finmap.union_comm_of_disjoint hd] + . unfold State.clsSingleton + tauto + . apply SLP.ent_star_top + tauto end Lampe.STHoare From 6305337baeef2cb61fc18dfab0c56fa8a2d455f1 Mon Sep 17 00:00:00 2001 From: utkn Date: Mon, 25 Nov 2024 20:45:35 +0300 Subject: [PATCH 23/39] Lambda structure created. Parsing fixed. `clsSingleton` SLTerm integrated into the tactics --- Lampe/Lampe.lean | 2 +- Lampe/Lampe/Ast.lean | 9 +++-- Lampe/Lampe/Hoare/SepTotal.lean | 15 ++++--- Lampe/Lampe/Semantics.lean | 11 +++--- Lampe/Lampe/SeparationLogic/ValHeap.lean | 1 - Lampe/Lampe/Tactic/SeparationLogic.lean | 50 +++++++++++++++++++++++- 6 files changed, 70 insertions(+), 18 deletions(-) diff --git a/Lampe/Lampe.lean b/Lampe/Lampe.lean index 5372a18..d3d7bd3 100644 --- a/Lampe/Lampe.lean +++ b/Lampe/Lampe.lean @@ -99,5 +99,5 @@ nr_def simple_lambda<>(x : Field) -> Field { example {p Γ x} : STHoare p Γ ⟦⟧ (simple_lambda.fn.body (Tp.denote p) h![] h![x]) fun v => v = x := by simp only [simple_lambda] - steps + steps <;> try tauto all_goals sorry diff --git a/Lampe/Lampe/Ast.lean b/Lampe/Lampe/Ast.lean index b74d1f9..f5f2c01 100644 --- a/Lampe/Lampe/Ast.lean +++ b/Lampe/Lampe/Ast.lean @@ -29,7 +29,7 @@ inductive Expr (rep : Tp → Type) : Tp → Type where | ite : rep .bool → Expr rep a → Expr rep a → Expr rep a | skip : Expr rep .unit | loop : rep (.u s) → rep (.u s) → (rep (.u s) → Expr rep r) → Expr rep .unit -| lambda : ((argTps : List Tp) → (outTp : Tp) → (HList rep argTps → Expr rep outTp)) → Expr rep .lambdaRef +| lambda : (argTps : List Tp) → (outTp : Tp) → (HList rep argTps → Expr rep outTp) → Expr rep .lambdaRef structure Function : Type _ where generics : List Kind @@ -47,8 +47,11 @@ example : Function := { body := fun _ h![_] h![x] => .var x } -def Lambda : Type _ := - (rep : Tp → Type) → (argTps : List Tp) → (outTp : Tp) → HList rep argTps → Expr rep outTp +structure Lambda where + rep : Tp → Type + argTps : List Tp + outTp : Tp + body : HList rep argTps → Expr rep outTp structure FunctionDecl where name : Ident diff --git a/Lampe/Lampe/Hoare/SepTotal.lean b/Lampe/Lampe/Hoare/SepTotal.lean index fa2baa1..32f8489 100644 --- a/Lampe/Lampe/Hoare/SepTotal.lean +++ b/Lampe/Lampe/Hoare/SepTotal.lean @@ -306,10 +306,11 @@ theorem skip_intro : . apply SLP.ent_star_top tauto -theorem callLambda_intro : - @STHoare outTp p Γ ⟦⟧ (lambdaBody (Tp.denote p) argTps outTp args) (fun v => v = v') → - STHoare p Γ [ref ↣ lambdaBody] (.call h![] argTps outTp (.lambda ref) args) - (fun v => ⟦v = v'⟧ ⋆ [ref ↣ lambdaBody]) := by +theorem callLambda_intro {lambdaBody} : + @STHoare outTp p Γ ⟦⟧ (lambdaBody args) (fun v => v = v') → + STHoare p Γ [ref ↣ ⟨_, argTps, outTp, lambdaBody⟩] + (Expr.call h![] argTps outTp (.lambda ref) args) + (fun v => ⟦v = v'⟧ ⋆ [ref ↣ ⟨_, argTps, outTp, lambdaBody⟩]) := by intros rename_i h unfold STHoare THoare @@ -330,14 +331,15 @@ theorem callLambda_intro : simp only [SLP.true_star, SLP.entails_self] theorem newLambda_intro : - STHoare p Γ ⟦⟧ (.lambda (lambda _)) - fun v => [v ↣ lambda] := by + STHoare p Γ ⟦⟧ (.lambda argTps outTp lambdaBody) + fun v => [v ↣ ⟨_, argTps, outTp, lambdaBody⟩] := by unfold STHoare THoare intros H st h constructor intros simp_all only [SLP.true_star, SLP.star_assoc] rename Ref => r + generalize (⟨_, _, _, _⟩ : Lambda) = lambda exists ⟨∅, Finmap.singleton r lambda⟩, st refine ⟨?_, ?_, ?_, ?_⟩ . simp only [LawfulHeap.disjoint] @@ -354,4 +356,5 @@ theorem newLambda_intro : tauto . apply SLP.ent_star_top tauto + end Lampe.STHoare diff --git a/Lampe/Lampe/Semantics.lean b/Lampe/Lampe/Semantics.lean index f92bd13..37e90bb 100644 --- a/Lampe/Lampe/Semantics.lean +++ b/Lampe/Lampe/Semantics.lean @@ -45,12 +45,12 @@ inductive Omni : (p : Prime) → Env → State p → Expr (Tp.denote p) tp → ( Omni p Γ st (htco ▸ fn.body _ (hkc ▸ generics) (htci ▸ args)) Q → Omni p Γ st (@Expr.call _ tyKinds generics argTypes res (.decl fname) args) Q | callLambda {cls : Closures} : - cls.lookup ref = some lambda → - Omni p Γ ⟨vh, cls⟩ (lambda (Tp.denote p) argTps outTp args) Q → + cls.lookup ref = some ⟨Tp.denote p, argTps, outTp, lambdaBody⟩ → + Omni p Γ ⟨vh, cls⟩ (lambdaBody args) Q → Omni p Γ ⟨vh, cls⟩ (Expr.call h![] argTps outTp (.lambda ref) args) Q | newLambda {Q} : - (∀ ref, ref ∉ cls → Q (some (⟨vh, cls.insert ref lambda⟩, ref))) → - Omni p Γ ⟨vh, cls⟩ (@Expr.lambda (Tp.denote p) (lambda _)) Q + (∀ ref, ref ∉ cls → Q (some (⟨vh, cls.insert ref ⟨_, argTps, outTp, lambdaBody⟩⟩, ref))) → + Omni p Γ ⟨vh, cls⟩ (.lambda argTps outTp lambdaBody) Q | loopDone : lo ≥ hi → Omni p Γ st (.loop lo hi body) Q @@ -192,7 +192,7 @@ theorem Omni.frame {p Γ tp} {st₁ st₂ : State p} {e : Expr (Tp.denote p) tp} rename Closures => cls rename ValHeap _ => vh rename Ref => r - rename Lambda => lambda + generalize hL : (⟨_, _, _, _⟩ : Lambda) = lambda have hi : r ∉ cls ∧ r ∉ st₂.closures := by aesop have hd₁ : Finmap.Disjoint cls (Finmap.singleton r lambda) := by apply Finmap.Disjoint.symm @@ -213,6 +213,7 @@ theorem Omni.frame {p Γ tp} {st₁ st₂ : State p} {e : Expr (Tp.denote p) tp} simp only [Finmap.union_comm_of_disjoint hd₁, Finmap.union_assoc] . simp [Finmap.union_comm_of_disjoint, Finmap.insert_eq_singleton_union] rename (∀ref ∉ cls, _) => hQ + rw [hL] at hQ simp only [Finmap.insert_eq_singleton_union] at hQ rw [Finmap.union_comm_of_disjoint hd₁] tauto diff --git a/Lampe/Lampe/SeparationLogic/ValHeap.lean b/Lampe/Lampe/SeparationLogic/ValHeap.lean index c4057b4..7a59596 100644 --- a/Lampe/Lampe/SeparationLogic/ValHeap.lean +++ b/Lampe/Lampe/SeparationLogic/ValHeap.lean @@ -9,7 +9,6 @@ lemma Finmap.singleton_disjoint_of_not_mem (hp : ref ∉ s): Finmap.Disjoint (Finmap.singleton ref v) s := by simp_all [Finmap.Disjoint] - lemma Finmap.singleton_disjoint_iff_not_mem : Finmap.Disjoint (Finmap.singleton ref v) s ↔ (ref ∉ s) := by simp_all [Finmap.Disjoint] diff --git a/Lampe/Lampe/Tactic/SeparationLogic.lean b/Lampe/Lampe/Tactic/SeparationLogic.lean index 2d3e9c0..960ac29 100644 --- a/Lampe/Lampe/Tactic/SeparationLogic.lean +++ b/Lampe/Lampe/Tactic/SeparationLogic.lean @@ -14,6 +14,7 @@ inductive SLTerm where | star : Expr → SLTerm → SLTerm → SLTerm | lift : Expr → SLTerm | singleton : Expr → Expr → SLTerm +| clsSingleton : Expr → Expr → SLTerm | mvar : Expr → SLTerm | all : Expr → SLTerm | exi : Expr → SLTerm @@ -28,6 +29,7 @@ def SLTerm.toString : SLTerm → String | star _ a b => s!"({a.toString} ⋆ {b.toString})" | lift e => s!"⟦{e.dbgToString}⟧" | singleton e₁ _ => s!"[{e₁.dbgToString} ↦ _]" +| clsSingleton e₁ _ => s!"[{e₁.dbgToString} ↣ _]" | mvar e => s!"MV{e.dbgToString}" | unrecognized e => s!"" @@ -184,9 +186,14 @@ partial def parseSLExpr (e: Expr): TacticM SLTerm := do return SLTerm.star e fst snd if e.isAppOf ``State.valSingleton then let args := e.getAppArgs - let fst ←args[1]? + let fst ← args[1]? let snd ← args[2]? return SLTerm.singleton fst snd + else if e.isAppOf ``State.clsSingleton then + let args := e.getAppArgs + let fst ← args[1]? + let snd ← args[2]? + return SLTerm.clsSingleton fst snd else if e.isAppOf ``SLP.top then return SLTerm.top else if e.isAppOf ``SLP.lift then @@ -326,10 +333,22 @@ theorem singleton_congr_mv {p} {r} {v₁ v₂ : AnyValue p} : (v₁ = v₂) → simp apply SLP.entails_self +theorem clsSingleton_congr_mv {p} {r} {l₁ l₂ : Lambda} : (l₁ = l₂) → (([r ↣ l₁] : SLP (State p)) ⊢ [r ↣ l₂] ⋆ ⟦⟧) := by + rintro rfl + simp + apply SLP.entails_self + + theorem singleton_star_congr {p} {r} {v₁ v₂ : AnyValue p} {R} : (v₁ = v₂) → ([r ↦ v₁] ⋆ R ⊢ [r ↦ v₂] ⋆ R) := by rintro rfl apply SLP.entails_self +theorem clsSingleton_star_congr {p} {r} {v₁ v₂ : Lambda} {R : SLP (State p)} : + (v₁ = v₂) → ([r ↣ v₁] ⋆ R ⊢ [r ↣ v₂] ⋆ R) := by + rintro rfl + apply SLP.entails_self + + def canSolveSingleton (lhs : SLTerm) (rhsV : Expr): Bool := match lhs with | SLTerm.singleton v _ => v == rhsV @@ -346,6 +365,14 @@ partial def solveSingletonStarMV (goal : MVarId) (lhs : SLTerm) (rhs : Expr): Ta catch _ => pure [newGoal] pure $ newGoal ++ newGoals else throwError "not equal" + | SLTerm.clsSingleton v _ => + if v == rhs then + let newGoals ← goal.apply (←mkConstWithFreshMVarLevels ``clsSingleton_congr_mv) + let newGoal ← newGoals[0]? + let newGoal ← try newGoal.refl; pure [] + catch _ => pure [newGoal] + pure $ newGoal ++ newGoals + else throwError "not equal" | SLTerm.star _ l r => match l with | SLTerm.singleton v _ => do @@ -361,6 +388,19 @@ partial def solveSingletonStarMV (goal : MVarId) (lhs : SLTerm) (rhs : Expr): Ta let newGoal ← newGoals[0]? let new' ← solveSingletonStarMV newGoal r rhs return new' ++ newGoals + | SLTerm.clsSingleton v _ => do + if v == rhs then + -- [TODO] This should use EQ, not ent_self as well + let newGoals ← goal.apply (←mkConstWithFreshMVarLevels ``clsSingleton_star_congr) + let newGoal ← newGoals[0]? + let newGoal ← try newGoal.refl; pure [] + catch _ => pure [newGoal] + pure $ newGoal ++ newGoals + else + let newGoals ← goal.apply (←mkConstWithFreshMVarLevels ``use_right) + let newGoal ← newGoals[0]? + let new' ← solveSingletonStarMV newGoal r rhs + return new' ++ newGoals | SLTerm.lift _ => let goals ← goal.apply (←mkConstWithFreshMVarLevels ``pure_star_H_ent_pure_star_mv) let g ← goals[0]? @@ -392,11 +432,14 @@ partial def solvePureStarMV (goal : MVarId) (lhs : SLTerm): TacticM (List MVarId return ng ++ goals | .singleton _ _ => goal.apply (←mkConstWithFreshMVarLevels ``skip_evidence_pure) + | .clsSingleton _ _ => + goal.apply (←mkConstWithFreshMVarLevels ``skip_evidence_pure) | _ => throwError "not a lift {lhs}" partial def solveStarMV (goal : MVarId) (lhs : SLTerm) (rhsNonMv : SLTerm): TacticM (List MVarId) := do match rhsNonMv with | .singleton v _ => solveSingletonStarMV goal lhs v + | .clsSingleton v _ => solveSingletonStarMV goal lhs v | .lift _ => solvePureStarMV goal lhs | _ => throwError "not a singleton srry {rhsNonMv}" @@ -465,6 +508,7 @@ macro "stephelper1" : tactic => `(tactic|( | apply assert_intro | apply skip_intro | apply callLambda_intro + | apply newLambda_intro -- memory builtins | apply var_intro | apply ref_intro @@ -516,8 +560,9 @@ macro "stephelper2" : tactic => `(tactic|( | apply consequence_frame_left fresh_intro | apply consequence_frame_left Lampe.STHoare.litU_intro | apply consequence_frame_left assert_intro - | apply consequence_frame_left callLambda_intro -- | apply consequence_frame_left skip_intro + | apply consequence_frame_left callLambda_intro + | apply consequence_frame_left newLambda_intro -- memory builtins | apply consequence_frame_left var_intro | apply consequence_frame_left ref_intro @@ -572,6 +617,7 @@ macro "stephelper3" : tactic => `(tactic|( | apply ramified_frame_top assert_intro | apply ramified_frame_top skip_intro | apply ramified_frame_top callLambda_intro + | apply ramified_frame_top newLambda_intro -- memory builtins | apply ramified_frame_top var_intro | apply ramified_frame_top ref_intro From 4d067d07d5d5b8efe743d79c75b9595745b547e6 Mon Sep 17 00:00:00 2001 From: utkn Date: Tue, 26 Nov 2024 11:54:27 +0300 Subject: [PATCH 24/39] "Closure" renamed to "lambda" across the project for precision --- Lampe/Lampe.lean | 6 +++-- Lampe/Lampe/Hoare/Builtins.lean | 6 ++--- Lampe/Lampe/Hoare/SepTotal.lean | 4 +-- Lampe/Lampe/Hoare/Total.lean | 4 +-- Lampe/Lampe/Semantics.lean | 30 +++++++++++----------- Lampe/Lampe/SeparationLogic/State.lean | 34 ++++++++++++------------- Lampe/Lampe/Tactic/SeparationLogic.lean | 24 ++++++++--------- 7 files changed, 55 insertions(+), 53 deletions(-) diff --git a/Lampe/Lampe.lean b/Lampe/Lampe.lean index d3d7bd3..b26d06c 100644 --- a/Lampe/Lampe.lean +++ b/Lampe/Lampe.lean @@ -97,7 +97,9 @@ nr_def simple_lambda<>(x : Field) -> Field { ^foo(x) : Field; } -example {p Γ x} : STHoare p Γ ⟦⟧ (simple_lambda.fn.body (Tp.denote p) h![] h![x]) fun v => v = x := by +example {p Γ x} : STHoare p Γ ⟦⟧ (simple_lambda.fn.body (Tp.denote p) h![] h![x]) + fun v => v = x := by simp only [simple_lambda] steps <;> try tauto - all_goals sorry + . apply STHoare.var_intro + . sorry -- unsolvable diff --git a/Lampe/Lampe/Hoare/Builtins.lean b/Lampe/Lampe/Hoare/Builtins.lean index d234f41..f4f937d 100644 --- a/Lampe/Lampe/Hoare/Builtins.lean +++ b/Lampe/Lampe/Hoare/Builtins.lean @@ -355,7 +355,7 @@ theorem ref_intro: apply THoare.consequence ?_ THoare.ref_intro (fun _ => SLP.entails_self) simp only [SLP.true_star] intro st hH r hr - exists (⟨Finmap.singleton r ⟨tp, v⟩, st.closures⟩ ∪ st), ∅ + exists (⟨Finmap.singleton r ⟨tp, v⟩, st.lambdas⟩ ∪ st), ∅ apply And.intro (by rw [LawfulHeap.disjoint_symm_iff]; apply LawfulHeap.disjoint_empty) constructor . simp only [State.insertVal, Finmap.insert_eq_singleton_union, LawfulHeap_union_empty] @@ -395,7 +395,7 @@ theorem readRef_intro: apply Option.isSome_some simp only [SLP.true_star, SLP.star_assoc] rename_i st₁ st₂ _ _ - exists (⟨Finmap.singleton r ⟨tp, v⟩, st₁.closures⟩), ?_ + exists (⟨Finmap.singleton r ⟨tp, v⟩, st₁.lambdas⟩), ?_ unfold State.valSingleton at hs rw [←hs] repeat apply And.intro (by tauto) @@ -420,7 +420,7 @@ theorem writeRef_intro: simp_all [State.membership_in_val] simp only [Finmap.insert_eq_singleton_union, ←Finmap.union_assoc, Finmap.union_singleton, SLP.star_assoc] rename_i st₁ st₂ _ _ - exists (⟨Finmap.singleton r ⟨tp, v'⟩, st₁.closures⟩), ?_ + exists (⟨Finmap.singleton r ⟨tp, v'⟩, st₁.lambdas⟩), ?_ refine ⟨?_, ?_, ?_, (by apply SLP.ent_star_top; assumption)⟩ . simp only [LawfulHeap.disjoint] at * have _ : r ∉ st₂.vals := by diff --git a/Lampe/Lampe/Hoare/SepTotal.lean b/Lampe/Lampe/Hoare/SepTotal.lean index 32f8489..44ab122 100644 --- a/Lampe/Lampe/Hoare/SepTotal.lean +++ b/Lampe/Lampe/Hoare/SepTotal.lean @@ -348,11 +348,11 @@ theorem newLambda_intro : . apply Finmap.singleton_disjoint_of_not_mem (by tauto) . simp only [State.union_parts, State.mk.injEq] refine ⟨by simp_all, ?_⟩ - have hd : Finmap.Disjoint st.closures (Finmap.singleton r lambda) := by + have hd : Finmap.Disjoint st.lambdas (Finmap.singleton r lambda) := by rw [Finmap.Disjoint.symm_iff] apply Finmap.singleton_disjoint_of_not_mem (by assumption) simp only [Finmap.insert_eq_singleton_union, Finmap.union_comm_of_disjoint hd] - . unfold State.clsSingleton + . unfold State.lmbSingleton tauto . apply SLP.ent_star_top tauto diff --git a/Lampe/Lampe/Hoare/Total.lean b/Lampe/Lampe/Hoare/Total.lean index ebddab3..23dcfb7 100644 --- a/Lampe/Lampe/Hoare/Total.lean +++ b/Lampe/Lampe/Hoare/Total.lean @@ -43,7 +43,7 @@ theorem letIn_intro {P Q} theorem ref_intro {v P} : THoare p Γ - (fun st => ∀r, r ∉ st → P r ⟨(st.vals.insert r ⟨tp, v⟩), st.closures⟩) + (fun st => ∀r, r ∉ st → P r ⟨(st.vals.insert r ⟨tp, v⟩), st.lambdas⟩) (.call h![] [tp] (Tp.ref tp) (.builtin .ref) h![v]) P := by unfold THoare @@ -64,7 +64,7 @@ theorem readRef_intro {ref} : theorem writeRef_intro {ref v} : THoare p Γ - (fun st => ref ∈ st ∧ P () ⟨(st.vals.insert ref ⟨tp, v⟩), st.closures⟩) + (fun st => ref ∈ st ∧ P () ⟨(st.vals.insert ref ⟨tp, v⟩), st.lambdas⟩) (.call h![] [tp.ref, tp] .unit (.builtin .writeRef) h![ref, v]) P := by unfold THoare diff --git a/Lampe/Lampe/Semantics.lean b/Lampe/Lampe/Semantics.lean index 37e90bb..2304b14 100644 --- a/Lampe/Lampe/Semantics.lean +++ b/Lampe/Lampe/Semantics.lean @@ -35,7 +35,7 @@ inductive Omni : (p : Prime) → Env → State p → Expr (Tp.denote p) tp → ( (Q₁ none → Q none) → Omni p Γ st (.letIn e b) Q | callBuiltin {st : State p} {Q : Option (State p × Tp.denote p resType) → Prop} : - (b.omni p st.vals argTypes resType args (mapToValHeapCondition st.closures Q)) → + (b.omni p st.vals argTypes resType args (mapToValHeapCondition st.lambdas Q)) → Omni p Γ st (Expr.call h![] argTypes resType (.builtin b) args) Q | callDecl : Γ fname = some fn → @@ -44,13 +44,13 @@ inductive Omni : (p : Prime) → Env → State p → Expr (Tp.denote p) tp → ( (htco : fn.outTp (hkc ▸ generics) = res) → Omni p Γ st (htco ▸ fn.body _ (hkc ▸ generics) (htci ▸ args)) Q → Omni p Γ st (@Expr.call _ tyKinds generics argTypes res (.decl fname) args) Q -| callLambda {cls : Closures} : - cls.lookup ref = some ⟨Tp.denote p, argTps, outTp, lambdaBody⟩ → - Omni p Γ ⟨vh, cls⟩ (lambdaBody args) Q → - Omni p Γ ⟨vh, cls⟩ (Expr.call h![] argTps outTp (.lambda ref) args) Q +| callLambda {lambdas : Lambdas} : + lambdas.lookup ref = some ⟨Tp.denote p, argTps, outTp, lambdaBody⟩ → + Omni p Γ ⟨vh, lambdas⟩ (lambdaBody args) Q → + Omni p Γ ⟨vh, lambdas⟩ (Expr.call h![] argTps outTp (.lambda ref) args) Q | newLambda {Q} : - (∀ ref, ref ∉ cls → Q (some (⟨vh, cls.insert ref ⟨_, argTps, outTp, lambdaBody⟩⟩, ref))) → - Omni p Γ ⟨vh, cls⟩ (.lambda argTps outTp lambdaBody) Q + (∀ ref, ref ∉ lambdas → Q (some (⟨vh, lambdas.insert ref ⟨_, argTps, outTp, lambdaBody⟩⟩, ref))) → + Omni p Γ ⟨vh, lambdas⟩ (Expr.lambda argTps outTp lambdaBody) Q | loopDone : lo ≥ hi → Omni p Γ st (.loop lo hi body) Q @@ -137,7 +137,7 @@ theorem Omni.frame {p Γ tp} {st₁ st₂ : State p} {e : Expr (Tp.denote p) tp} refine ⟨by tauto, ?_, ?_, by tauto⟩ . simp only [State.union_parts] at hin₂ injection hin₂ - . have hc : s₁.closures = st₁.closures := by + . have hc : s₁.lambdas = st₁.lambdas := by obtain ⟨_, hd₂⟩ := hin₁ rw [State.union_parts] at hin₂ injection hin₂ with _ hu @@ -150,7 +150,7 @@ theorem Omni.frame {p Γ tp} {st₁ st₂ : State p} {e : Expr (Tp.denote p) tp} tauto rw [←hc] tauto - . exists ⟨s₁, st₁.closures⟩, ⟨s₂, st₂.closures⟩ + . exists ⟨s₁, st₁.lambdas⟩, ⟨s₂, st₂.lambdas⟩ simp only [LawfulHeap.disjoint] at * refine ⟨by tauto, ?_, by tauto, ?_⟩ . simp only [State.union_parts, State.mk.injEq] @@ -189,19 +189,19 @@ theorem Omni.frame {p Γ tp} {st₁ st₂ : State p} {e : Expr (Tp.denote p) tp} constructor intros simp only - rename Closures => cls + rename Lambdas => lmbs rename ValHeap _ => vh rename Ref => r generalize hL : (⟨_, _, _, _⟩ : Lambda) = lambda - have hi : r ∉ cls ∧ r ∉ st₂.closures := by aesop - have hd₁ : Finmap.Disjoint cls (Finmap.singleton r lambda) := by + have hi : r ∉ lmbs ∧ r ∉ st₂.lambdas := by aesop + have hd₁ : Finmap.Disjoint lmbs (Finmap.singleton r lambda) := by apply Finmap.Disjoint.symm apply Finmap.singleton_disjoint_iff_not_mem.mpr tauto - have hd₂ : Finmap.Disjoint (Finmap.singleton r lambda) st₂.closures := by + have hd₂ : Finmap.Disjoint (Finmap.singleton r lambda) st₂.lambdas := by apply Finmap.singleton_disjoint_iff_not_mem.mpr tauto - exists ⟨vh, cls ∪ Finmap.singleton r lambda⟩, st₂ + exists ⟨vh, lmbs ∪ Finmap.singleton r lambda⟩, st₂ refine ⟨?_, ?_, ?_, by tauto⟩ . simp only [LawfulHeap.disjoint] refine ⟨by tauto, ?_⟩ @@ -212,7 +212,7 @@ theorem Omni.frame {p Γ tp} {st₁ st₂ : State p} {e : Expr (Tp.denote p) tp} simp only [Finmap.insert_eq_singleton_union] simp only [Finmap.union_comm_of_disjoint hd₁, Finmap.union_assoc] . simp [Finmap.union_comm_of_disjoint, Finmap.insert_eq_singleton_union] - rename (∀ref ∉ cls, _) => hQ + rename (∀ref ∉ lmbs, _) => hQ rw [hL] at hQ simp only [Finmap.insert_eq_singleton_union] at hQ rw [Finmap.union_comm_of_disjoint hd₁] diff --git a/Lampe/Lampe/SeparationLogic/State.lean b/Lampe/Lampe/SeparationLogic/State.lean index 8eff141..81ff226 100644 --- a/Lampe/Lampe/SeparationLogic/State.lean +++ b/Lampe/Lampe/SeparationLogic/State.lean @@ -4,11 +4,11 @@ import Lampe.Ast namespace Lampe -abbrev Closures := Finmap fun _ : Ref => Lambda +abbrev Lambdas := Finmap fun _ : Ref => Lambda structure State (p : Prime) where vals : ValHeap p - closures : Closures + lambdas : Lambdas instance : Membership Ref (State p) where mem := fun a e => e ∈ a.vals @@ -19,12 +19,12 @@ lemma State.membership_in_val {a : State p} : e ∈ a ↔ e ∈ a.vals := by rfl instance : Coe (State p) (ValHeap p) where coe := fun s => s.vals -/-- Maps a post-condition on `State`s to a post-condition on `ValHeap`s by keeping the closures fixed -/ +/-- Maps a post-condition on `State`s to a post-condition on `ValHeap`s by keeping the lambdas fixed -/ @[reducible] def mapToValHeapCondition - (closures : Closures) + (lambdas : Lambdas) (Q : Option (State p × T) → Prop) : Option (ValHeap p × T) → Prop := - fun vv => Q (vv.map (fun (vals, t) => ⟨⟨vals, closures⟩, t⟩)) + fun vv => Q (vv.map (fun (vals, t) => ⟨⟨vals, lambdas⟩, t⟩)) /-- Maps a post-condition on `ValHeap`s to a post-condition on `State`s -/ @[reducible] @@ -33,10 +33,10 @@ def mapToStateCondition fun stv => Q (stv.map (fun (st, t) => ⟨st.vals, t⟩)) def State.insertVal (self : State p) (r : Ref) (v : AnyValue p) : State p := - ⟨self.vals.insert r v, self.closures⟩ + ⟨self.vals.insert r v, self.lambdas⟩ lemma State.eq_constructor {st₁ : State p} : - (st₁ = st₂) ↔ (State.mk st₁.vals st₁.closures = State.mk st₂.vals st₂.closures) := by + (st₁ = st₂) ↔ (State.mk st₁.vals st₁.lambdas = State.mk st₂.vals st₂.lambdas) := by rfl @[simp] @@ -46,8 +46,8 @@ lemma State.eq_closures : injection h instance : LawfulHeap (State p) where - union := fun a b => ⟨a.vals ∪ b.vals, a.closures ∪ b.closures⟩ - disjoint := fun a b => a.vals.Disjoint b.vals ∧ a.closures.Disjoint b.closures + union := fun a b => ⟨a.vals ∪ b.vals, a.lambdas ∪ b.lambdas⟩ + disjoint := fun a b => a.vals.Disjoint b.vals ∧ a.lambdas.Disjoint b.lambdas empty := ⟨∅, ∅⟩ union_empty := by intros @@ -71,7 +71,7 @@ instance : LawfulHeap (State p) where disjoint_union_left := by intros x y z have h1 := (Finmap.disjoint_union_left x.vals y.vals z.vals) - have h2 := (Finmap.disjoint_union_left x.closures y.closures z.closures) + have h2 := (Finmap.disjoint_union_left x.lambdas y.lambdas z.lambdas) constructor intro h3 simp only [Union.union] at h3 @@ -85,24 +85,24 @@ def State.valSingleton (r : Ref) (v : AnyValue p) : SLP (State p) := notation:max "[" l " ↦ " r "]" => State.valSingleton l r @[reducible] -def State.clsSingleton (r : Ref) (v : Lambda) : SLP (State p) := - fun st => st.closures = Finmap.singleton r v +def State.lmbSingleton (r : Ref) (v : Lambda) : SLP (State p) := + fun st => st.lambdas = Finmap.singleton r v -notation:max "[" l " ↣ " r "]" => State.clsSingleton l r +notation:max "[" l " ↣ " r "]" => State.lmbSingleton l r @[simp] lemma State.union_parts_left : - (State.mk v c ∪ st₂ = State.mk (v ∪ st₂.vals) (c ∪ st₂.closures)) := by + (State.mk v c ∪ st₂ = State.mk (v ∪ st₂.vals) (c ∪ st₂.lambdas)) := by rfl @[simp] lemma State.union_parts_right : - (st₂ ∪ State.mk v c = State.mk (st₂.vals ∪ v) (st₂.closures ∪ c)) := by + (st₂ ∪ State.mk v c = State.mk (st₂.vals ∪ v) (st₂.lambdas ∪ c)) := by rfl @[simp] lemma State.union_parts : - st₁ ∪ st₂ = State.mk (st₁.vals ∪ st₂.vals) (st₁.closures ∪ st₂.closures) := by + st₁ ∪ st₂ = State.mk (st₁.vals ∪ st₂.vals) (st₁.lambdas ∪ st₂.lambdas) := by rfl @[simp] @@ -111,6 +111,6 @@ lemma State.union_vals {st₁ st₂ : State p} : @[simp] lemma State.union_closures {st₁ st₂ : State p} : - (st₁ ∪ st₂).closures = (st₁.closures ∪ st₂.closures) := by rfl + (st₁ ∪ st₂).lambdas = (st₁.lambdas ∪ st₂.lambdas) := by rfl end Lampe diff --git a/Lampe/Lampe/Tactic/SeparationLogic.lean b/Lampe/Lampe/Tactic/SeparationLogic.lean index 960ac29..b016f22 100644 --- a/Lampe/Lampe/Tactic/SeparationLogic.lean +++ b/Lampe/Lampe/Tactic/SeparationLogic.lean @@ -14,7 +14,7 @@ inductive SLTerm where | star : Expr → SLTerm → SLTerm → SLTerm | lift : Expr → SLTerm | singleton : Expr → Expr → SLTerm -| clsSingleton : Expr → Expr → SLTerm +| lmbSingleton : Expr → Expr → SLTerm | mvar : Expr → SLTerm | all : Expr → SLTerm | exi : Expr → SLTerm @@ -29,7 +29,7 @@ def SLTerm.toString : SLTerm → String | star _ a b => s!"({a.toString} ⋆ {b.toString})" | lift e => s!"⟦{e.dbgToString}⟧" | singleton e₁ _ => s!"[{e₁.dbgToString} ↦ _]" -| clsSingleton e₁ _ => s!"[{e₁.dbgToString} ↣ _]" +| lmbSingleton e₁ _ => s!"[{e₁.dbgToString} ↣ _]" | mvar e => s!"MV{e.dbgToString}" | unrecognized e => s!"" @@ -189,11 +189,11 @@ partial def parseSLExpr (e: Expr): TacticM SLTerm := do let fst ← args[1]? let snd ← args[2]? return SLTerm.singleton fst snd - else if e.isAppOf ``State.clsSingleton then + else if e.isAppOf ``State.lmbSingleton then let args := e.getAppArgs let fst ← args[1]? let snd ← args[2]? - return SLTerm.clsSingleton fst snd + return SLTerm.lmbSingleton fst snd else if e.isAppOf ``SLP.top then return SLTerm.top else if e.isAppOf ``SLP.lift then @@ -333,7 +333,7 @@ theorem singleton_congr_mv {p} {r} {v₁ v₂ : AnyValue p} : (v₁ = v₂) → simp apply SLP.entails_self -theorem clsSingleton_congr_mv {p} {r} {l₁ l₂ : Lambda} : (l₁ = l₂) → (([r ↣ l₁] : SLP (State p)) ⊢ [r ↣ l₂] ⋆ ⟦⟧) := by +theorem lmbSingleton_congr_mv {p} {r} {l₁ l₂ : Lambda} : (l₁ = l₂) → (([r ↣ l₁] : SLP (State p)) ⊢ [r ↣ l₂] ⋆ ⟦⟧) := by rintro rfl simp apply SLP.entails_self @@ -343,7 +343,7 @@ theorem singleton_star_congr {p} {r} {v₁ v₂ : AnyValue p} {R} : (v₁ = v₂ rintro rfl apply SLP.entails_self -theorem clsSingleton_star_congr {p} {r} {v₁ v₂ : Lambda} {R : SLP (State p)} : +theorem lmbSingleton_star_congr {p} {r} {v₁ v₂ : Lambda} {R : SLP (State p)} : (v₁ = v₂) → ([r ↣ v₁] ⋆ R ⊢ [r ↣ v₂] ⋆ R) := by rintro rfl apply SLP.entails_self @@ -365,9 +365,9 @@ partial def solveSingletonStarMV (goal : MVarId) (lhs : SLTerm) (rhs : Expr): Ta catch _ => pure [newGoal] pure $ newGoal ++ newGoals else throwError "not equal" - | SLTerm.clsSingleton v _ => + | SLTerm.lmbSingleton v _ => if v == rhs then - let newGoals ← goal.apply (←mkConstWithFreshMVarLevels ``clsSingleton_congr_mv) + let newGoals ← goal.apply (←mkConstWithFreshMVarLevels ``lmbSingleton_congr_mv) let newGoal ← newGoals[0]? let newGoal ← try newGoal.refl; pure [] catch _ => pure [newGoal] @@ -388,10 +388,10 @@ partial def solveSingletonStarMV (goal : MVarId) (lhs : SLTerm) (rhs : Expr): Ta let newGoal ← newGoals[0]? let new' ← solveSingletonStarMV newGoal r rhs return new' ++ newGoals - | SLTerm.clsSingleton v _ => do + | SLTerm.lmbSingleton v _ => do if v == rhs then -- [TODO] This should use EQ, not ent_self as well - let newGoals ← goal.apply (←mkConstWithFreshMVarLevels ``clsSingleton_star_congr) + let newGoals ← goal.apply (←mkConstWithFreshMVarLevels ``lmbSingleton_star_congr) let newGoal ← newGoals[0]? let newGoal ← try newGoal.refl; pure [] catch _ => pure [newGoal] @@ -432,14 +432,14 @@ partial def solvePureStarMV (goal : MVarId) (lhs : SLTerm): TacticM (List MVarId return ng ++ goals | .singleton _ _ => goal.apply (←mkConstWithFreshMVarLevels ``skip_evidence_pure) - | .clsSingleton _ _ => + | .lmbSingleton _ _ => goal.apply (←mkConstWithFreshMVarLevels ``skip_evidence_pure) | _ => throwError "not a lift {lhs}" partial def solveStarMV (goal : MVarId) (lhs : SLTerm) (rhsNonMv : SLTerm): TacticM (List MVarId) := do match rhsNonMv with | .singleton v _ => solveSingletonStarMV goal lhs v - | .clsSingleton v _ => solveSingletonStarMV goal lhs v + | .lmbSingleton v _ => solveSingletonStarMV goal lhs v | .lift _ => solvePureStarMV goal lhs | _ => throwError "not a singleton srry {rhsNonMv}" From fbca229a69291b8ade8365c90672042019e30231 Mon Sep 17 00:00:00 2001 From: utkn Date: Tue, 26 Nov 2024 12:53:38 +0300 Subject: [PATCH 25/39] Renamed newLambda to lam --- Lampe/Lampe/Hoare/SepTotal.lean | 2 +- Lampe/Lampe/Semantics.lean | 4 ++-- Lampe/Lampe/Tactic/SeparationLogic.lean | 6 +++--- 3 files changed, 6 insertions(+), 6 deletions(-) diff --git a/Lampe/Lampe/Hoare/SepTotal.lean b/Lampe/Lampe/Hoare/SepTotal.lean index 44ab122..388c248 100644 --- a/Lampe/Lampe/Hoare/SepTotal.lean +++ b/Lampe/Lampe/Hoare/SepTotal.lean @@ -330,7 +330,7 @@ theorem callLambda_intro {lambdaBody} : exact h simp only [SLP.true_star, SLP.entails_self] -theorem newLambda_intro : +theorem lam_intro : STHoare p Γ ⟦⟧ (.lambda argTps outTp lambdaBody) fun v => [v ↣ ⟨_, argTps, outTp, lambdaBody⟩] := by unfold STHoare THoare diff --git a/Lampe/Lampe/Semantics.lean b/Lampe/Lampe/Semantics.lean index 2304b14..280bd29 100644 --- a/Lampe/Lampe/Semantics.lean +++ b/Lampe/Lampe/Semantics.lean @@ -48,7 +48,7 @@ inductive Omni : (p : Prime) → Env → State p → Expr (Tp.denote p) tp → ( lambdas.lookup ref = some ⟨Tp.denote p, argTps, outTp, lambdaBody⟩ → Omni p Γ ⟨vh, lambdas⟩ (lambdaBody args) Q → Omni p Γ ⟨vh, lambdas⟩ (Expr.call h![] argTps outTp (.lambda ref) args) Q -| newLambda {Q} : +| lam {Q} : (∀ ref, ref ∉ lambdas → Q (some (⟨vh, lambdas.insert ref ⟨_, argTps, outTp, lambdaBody⟩⟩, ref))) → Omni p Γ ⟨vh, lambdas⟩ (Expr.lambda argTps outTp lambdaBody) Q | loopDone : @@ -182,7 +182,7 @@ theorem Omni.frame {p Γ tp} {st₁ st₂ : State p} {e : Expr (Tp.denote p) tp} simp only [LawfulHeap.disjoint] at hd simp only [Finmap.lookup_union_left (Finmap.mem_of_lookup_eq_some h)] tauto - | newLambda => + | lam => intros h simp only [LawfulHeap.disjoint, State.union_parts_left] at * obtain ⟨_, hd⟩ := h diff --git a/Lampe/Lampe/Tactic/SeparationLogic.lean b/Lampe/Lampe/Tactic/SeparationLogic.lean index b016f22..3aaa5c1 100644 --- a/Lampe/Lampe/Tactic/SeparationLogic.lean +++ b/Lampe/Lampe/Tactic/SeparationLogic.lean @@ -508,7 +508,7 @@ macro "stephelper1" : tactic => `(tactic|( | apply assert_intro | apply skip_intro | apply callLambda_intro - | apply newLambda_intro + | apply lam_intro -- memory builtins | apply var_intro | apply ref_intro @@ -562,7 +562,7 @@ macro "stephelper2" : tactic => `(tactic|( | apply consequence_frame_left assert_intro -- | apply consequence_frame_left skip_intro | apply consequence_frame_left callLambda_intro - | apply consequence_frame_left newLambda_intro + | apply consequence_frame_left lam_intro -- memory builtins | apply consequence_frame_left var_intro | apply consequence_frame_left ref_intro @@ -617,7 +617,7 @@ macro "stephelper3" : tactic => `(tactic|( | apply ramified_frame_top assert_intro | apply ramified_frame_top skip_intro | apply ramified_frame_top callLambda_intro - | apply ramified_frame_top newLambda_intro + | apply ramified_frame_top lam_intro -- memory builtins | apply ramified_frame_top var_intro | apply ramified_frame_top ref_intro From 427f728954bfa28efdeea7f089f4ea9f0e0ebf7b Mon Sep 17 00:00:00 2001 From: utkn Date: Tue, 26 Nov 2024 18:07:22 +0300 Subject: [PATCH 26/39] callLambda intro updated and proven --- Lampe/Lampe.lean | 21 ++++++++++++++++++--- Lampe/Lampe/Hoare/Builtins.lean | 4 ++-- Lampe/Lampe/Hoare/SepTotal.lean | 26 ++++++++++++++++++-------- 3 files changed, 38 insertions(+), 13 deletions(-) diff --git a/Lampe/Lampe.lean b/Lampe/Lampe.lean index b26d06c..bc00425 100644 --- a/Lampe/Lampe.lean +++ b/Lampe/Lampe.lean @@ -100,6 +100,21 @@ nr_def simple_lambda<>(x : Field) -> Field { example {p Γ x} : STHoare p Γ ⟦⟧ (simple_lambda.fn.body (Tp.denote p) h![] h![x]) fun v => v = x := by simp only [simple_lambda] - steps <;> try tauto - . apply STHoare.var_intro - . sorry -- unsolvable + steps + rotate_left 2 + exact (fun v => v = x) + rotate_left 1 + . steps + intros + subst_vars + rfl + . apply STHoare.ramified_frame_top + apply STHoare.callLambda_intro + rotate_left 2 + exact ⟦⟧ + exact (fun v => v = x) + rotate_left 2 + exact (fun h![a] => (Expr.var a)) + rotate_left 1 + apply STHoare.var_intro + sorry diff --git a/Lampe/Lampe/Hoare/Builtins.lean b/Lampe/Lampe/Hoare/Builtins.lean index f4f937d..7fc45c8 100644 --- a/Lampe/Lampe/Hoare/Builtins.lean +++ b/Lampe/Lampe/Hoare/Builtins.lean @@ -29,8 +29,8 @@ lemma pureBuiltin_intro_consequence {A : Type} {a : A} {sgn desc args} {Q : Tp.denote p outTp → Prop} (h1 : argTps = (sgn a).fst) (h2 : outTp = (sgn a).snd) - (hp : (h: (desc a (h1 ▸ args)).fst) → Q (h2 ▸ (desc a (h1 ▸ args)).snd h)) - : STHoare p Γ ⟦⟧ + (hp : (h: (desc a (h1 ▸ args)).fst) → Q (h2 ▸ (desc a (h1 ▸ args)).snd h)) : + STHoare p Γ ⟦⟧ (.call h![] argTps outTp (.builtin (Builtin.newGenericPureBuiltin sgn desc)) args) fun v => Q v := by subst_vars diff --git a/Lampe/Lampe/Hoare/SepTotal.lean b/Lampe/Lampe/Hoare/SepTotal.lean index 388c248..8e94d36 100644 --- a/Lampe/Lampe/Hoare/SepTotal.lean +++ b/Lampe/Lampe/Hoare/SepTotal.lean @@ -306,11 +306,11 @@ theorem skip_intro : . apply SLP.ent_star_top tauto -theorem callLambda_intro {lambdaBody} : - @STHoare outTp p Γ ⟦⟧ (lambdaBody args) (fun v => v = v') → - STHoare p Γ [ref ↣ ⟨_, argTps, outTp, lambdaBody⟩] +theorem callLambda_intro {lambdaBody} {P : SLP (State p)} {Q : Tp.denote p outTp → SLP (State p)}: + @STHoare outTp p Γ P (lambdaBody args) Q → + STHoare p Γ (P ⋆ [ref ↣ ⟨_, argTps, outTp, lambdaBody⟩]) (Expr.call h![] argTps outTp (.lambda ref) args) - (fun v => ⟦v = v'⟧ ⋆ [ref ↣ ⟨_, argTps, outTp, lambdaBody⟩]) := by + (fun v => (Q v) ⋆ [ref ↣ ⟨_, argTps, outTp, lambdaBody⟩]) := by intros rename_i h unfold STHoare THoare @@ -319,14 +319,24 @@ theorem callLambda_intro {lambdaBody} : unfold SLP.star at * . rename_i st h obtain ⟨st₁, ⟨st₂, ⟨_, h₂, h₃, _⟩⟩⟩ := h - simp only [State.union_parts, h₃] at h₂ - simp only [h₂] - simp_all + obtain ⟨st₁', ⟨st₂', _⟩⟩ := h₃ + simp_all only [State.union_parts, Finmap.mem_union, Finmap.mem_singleton, or_true, + Finmap.lookup_union_left] + generalize hL : (⟨_, _, _, _⟩ : Lambda) = lmb at * + have _ : ref ∉ st₁'.lambdas := by + rename_i h₃ + obtain ⟨hi₁, _, _, hi₂⟩ := h₃ + simp only [State.lmbSingleton] at hi₂ + simp only [LawfulHeap.disjoint] at hi₁ + obtain ⟨_, hj₁⟩ := hi₁ + rw [hi₂] at hj₁ + tauto + simp [Finmap.lookup_union_right (by tauto)] . apply consequence <;> tauto apply consequence_frame_left rotate_left 2 - exact ⟦⟧ + exact P exact h simp only [SLP.true_star, SLP.entails_self] From d0dab1bb06a12b91d4fd618955394e822e2adc81 Mon Sep 17 00:00:00 2001 From: utkn Date: Tue, 26 Nov 2024 23:14:39 +0300 Subject: [PATCH 27/39] steps tactic can now work with callLambda intro rule --- Lampe/Lampe.lean | 24 +++++++----------------- Lampe/Lampe/Hoare/SepTotal.lean | 16 ++++++++++++++-- Lampe/Lampe/Tactic/SeparationLogic.lean | 15 +++++++++++---- 3 files changed, 32 insertions(+), 23 deletions(-) diff --git a/Lampe/Lampe.lean b/Lampe/Lampe.lean index bc00425..39d5d68 100644 --- a/Lampe/Lampe.lean +++ b/Lampe/Lampe.lean @@ -101,20 +101,10 @@ example {p Γ x} : STHoare p Γ ⟦⟧ (simple_lambda.fn.body (Tp.denote p) h![] fun v => v = x := by simp only [simple_lambda] steps - rotate_left 2 - exact (fun v => v = x) - rotate_left 1 - . steps - intros - subst_vars - rfl - . apply STHoare.ramified_frame_top - apply STHoare.callLambda_intro - rotate_left 2 - exact ⟦⟧ - exact (fun v => v = x) - rotate_left 2 - exact (fun h![a] => (Expr.var a)) - rotate_left 1 - apply STHoare.var_intro - sorry + steps + simp only [SLP.true_star, SLP.entails_self] + intro + sl + tauto + sl + aesop diff --git a/Lampe/Lampe/Hoare/SepTotal.lean b/Lampe/Lampe/Hoare/SepTotal.lean index 8e94d36..9b110a7 100644 --- a/Lampe/Lampe/Hoare/SepTotal.lean +++ b/Lampe/Lampe/Hoare/SepTotal.lean @@ -91,6 +91,19 @@ theorem consequence_frame_left {H H₁ H₂ : SLP (State p)} rw [SLP.star_comm] apply SLP.ent_star_top +-- [TODO] find a better name for this +theorem utkans_thm {Q : _ → SLP (State p)} + (h_hoare_imp : STHoare p Γ P₁ e₁ Q → STHoare p Γ (P₁ ⋆ H) e₂ (fun v => Q v ⋆ H)) + (h_hoare : STHoare p Γ P₁ e₁ Q) + (h_ent_pre : H ⊢ P₁ ⋆ H) + (h_ent_post : ∀ v, ((Q v) ⋆ H) ⋆ ⊤ ⊢ (Q v) ⋆ ⊤) -- [TODO] get rid of this. + : + STHoare p Γ H e₂ Q := by + have h_hoare' := h_hoare_imp h_hoare + apply consequence h_ent_pre (fun v => SLP.entails_self) + apply consequence SLP.entails_self h_ent_post + tauto + theorem var_intro {v : Tp.denote p tp}: STHoare p Γ ⟦⟧ (.var v) (fun v' => ⟦v' = v⟧) := by unfold STHoare @@ -332,8 +345,7 @@ theorem callLambda_intro {lambdaBody} {P : SLP (State p)} {Q : Tp.denote p outTp rw [hi₂] at hj₁ tauto simp [Finmap.lookup_union_right (by tauto)] - . apply consequence - <;> tauto + . apply consequence <;> tauto apply consequence_frame_left rotate_left 2 exact P diff --git a/Lampe/Lampe/Tactic/SeparationLogic.lean b/Lampe/Lampe/Tactic/SeparationLogic.lean index 3aaa5c1..8b754f2 100644 --- a/Lampe/Lampe/Tactic/SeparationLogic.lean +++ b/Lampe/Lampe/Tactic/SeparationLogic.lean @@ -41,6 +41,10 @@ def SLTerm.isTop : SLTerm → Bool | SLTerm.top => true | _ => false +def SLTerm.isForAll : SLTerm → Bool +| SLTerm.all _ => true +| _ => false + instance : ToString SLTerm := ⟨SLTerm.toString⟩ instance : Inhabited SLTerm := ⟨SLTerm.top⟩ @@ -131,7 +135,10 @@ theorem SLP.pure_star_pure [LawfulHeap α] {P Q : Prop} : (P ⋆ Q) = (⟦P ∧ all_goals simp_all [LawfulHeap.disjoint_empty] macro "h_norm" : tactic => `(tactic|( - try simp only [SLP.star_assoc, pluck_pure_l, pluck_pure_l_assoc, pluck_pure_all_l, SLP.star_true, SLP.true_star, star_exists, exists_star]; + try simp only [SLP.star_assoc, + pluck_pure_l, pluck_pure_l_assoc, pluck_pure_all_l, + SLP.star_true, SLP.true_star, + star_exists, exists_star]; -- repeat (apply STHoare.pure_left; intro_cases); -- repeat (apply SLP.pure_left; intro_cases); subst_vars; @@ -473,6 +480,8 @@ partial def solveEntailment (goal : MVarId): TacticM (List MVarId) := do let g' ← g[0]? let ng ← solveEntailment g' pure $ ng ++ g + else if r.isForAll then + throwError "cannot solve forall" else throwError "todo {l} {r}" | SLTerm.singleton _ _ => -- [TODO] handle pure on the left @@ -507,7 +516,7 @@ macro "stephelper1" : tactic => `(tactic|( | apply fresh_intro | apply assert_intro | apply skip_intro - | apply callLambda_intro + | apply STHoare.utkans_thm STHoare.callLambda_intro | apply lam_intro -- memory builtins | apply var_intro @@ -561,7 +570,6 @@ macro "stephelper2" : tactic => `(tactic|( | apply consequence_frame_left Lampe.STHoare.litU_intro | apply consequence_frame_left assert_intro -- | apply consequence_frame_left skip_intro - | apply consequence_frame_left callLambda_intro | apply consequence_frame_left lam_intro -- memory builtins | apply consequence_frame_left var_intro @@ -616,7 +624,6 @@ macro "stephelper3" : tactic => `(tactic|( | apply ramified_frame_top Lampe.STHoare.litU_intro | apply ramified_frame_top assert_intro | apply ramified_frame_top skip_intro - | apply ramified_frame_top callLambda_intro | apply ramified_frame_top lam_intro -- memory builtins | apply ramified_frame_top var_intro From da4de6fc5a5c6384d36243702f34a8b438fcc096 Mon Sep 17 00:00:00 2001 From: utkn Date: Wed, 27 Nov 2024 10:36:50 +0300 Subject: [PATCH 28/39] updates in the steps tactic --- Lampe/Lampe.lean | 19 +++++++++++-------- Lampe/Lampe/Hoare/SepTotal.lean | 11 ++++++----- Lampe/Lampe/SeparationLogic/LawfulHeap.lean | 11 +++++++++++ Lampe/Lampe/SeparationLogic/SLP.lean | 16 ++++++++++------ 4 files changed, 38 insertions(+), 19 deletions(-) diff --git a/Lampe/Lampe.lean b/Lampe/Lampe.lean index 39d5d68..ec51ee6 100644 --- a/Lampe/Lampe.lean +++ b/Lampe/Lampe.lean @@ -92,19 +92,22 @@ example {p Γ x y}: STHoare p Γ ⟦⟧ (simple_if_else.fn.body (Tp.denote p) h! contradiction . aesop -nr_def simple_lambda<>(x : Field) -> Field { - let foo = |a|: Field -> Field { a }; - ^foo(x) : Field; +nr_def simple_lambda<>(x : Field, y : Field) -> Field { + let add = |a, b|: Field, Field -> Field { #add(a, b) : Field }; + ^add(x, y) : Field; } -example {p Γ x} : STHoare p Γ ⟦⟧ (simple_lambda.fn.body (Tp.denote p) h![] h![x]) - fun v => v = x := by +example {p Γ x y} : STHoare p Γ ⟦⟧ (simple_lambda.fn.body (Tp.denote p) h![] h![x, y]) + fun v => v = x + y := by simp only [simple_lambda] steps + simp_all steps - simp only [SLP.true_star, SLP.entails_self] - intro + simp_all + rotate_left 2 + simp_all [SLP.entails_self] + exact (fun v => v = x + y) sl - tauto + aesop sl aesop diff --git a/Lampe/Lampe/Hoare/SepTotal.lean b/Lampe/Lampe/Hoare/SepTotal.lean index 9b110a7..046247a 100644 --- a/Lampe/Lampe/Hoare/SepTotal.lean +++ b/Lampe/Lampe/Hoare/SepTotal.lean @@ -91,14 +91,15 @@ theorem consequence_frame_left {H H₁ H₂ : SLP (State p)} rw [SLP.star_comm] apply SLP.ent_star_top + -- [TODO] find a better name for this theorem utkans_thm {Q : _ → SLP (State p)} - (h_hoare_imp : STHoare p Γ P₁ e₁ Q → STHoare p Γ (P₁ ⋆ H) e₂ (fun v => Q v ⋆ H)) - (h_hoare : STHoare p Γ P₁ e₁ Q) - (h_ent_pre : H ⊢ P₁ ⋆ H) - (h_ent_post : ∀ v, ((Q v) ⋆ H) ⋆ ⊤ ⊢ (Q v) ⋆ ⊤) -- [TODO] get rid of this. - : + (h_hoare_imp : STHoare p Γ P e₁ Q → STHoare p Γ (P ⋆ H) e₂ (fun v => Q v ⋆ H)) + (h_hoare : STHoare p Γ P e₁ Q) + (h_ent_pre : H ⊢ P ⋆ H) : STHoare p Γ H e₂ Q := by + have h_ent_post : ∀ v, ((Q v) ⋆ H) ⋆ ⊤ ⊢ (Q v) ⋆ ⊤ := by + simp [SLP.ent_drop_left] have h_hoare' := h_hoare_imp h_hoare apply consequence h_ent_pre (fun v => SLP.entails_self) apply consequence SLP.entails_self h_ent_post diff --git a/Lampe/Lampe/SeparationLogic/LawfulHeap.lean b/Lampe/Lampe/SeparationLogic/LawfulHeap.lean index 0acb084..37feccb 100644 --- a/Lampe/Lampe/SeparationLogic/LawfulHeap.lean +++ b/Lampe/Lampe/SeparationLogic/LawfulHeap.lean @@ -28,6 +28,17 @@ theorem LawfulHeap_union_comm_of_disjoint [LawfulHeap α] {s₁ s₂ : α}: simp [Union.union] exact LawfulHeap.union_comm_of_disjoint +@[simp] +theorem LawfulHeap_disjoint_empty [LawfulHeap α] {a : α} : + LawfulHeap.disjoint a ∅ := by + rw [LawfulHeap.disjoint_symm_iff] + apply LawfulHeap.disjoint_empty + +@[simp] +theorem LawfulHeap_empty_disjoint [LawfulHeap α] {a : α} : + LawfulHeap.disjoint ∅ a := by + apply LawfulHeap.disjoint_empty + @[simp] theorem LawfulHeap_union_empty [LawfulHeap α] {a : α} : a ∪ ∅ = a := by diff --git a/Lampe/Lampe/SeparationLogic/SLP.lean b/Lampe/Lampe/SeparationLogic/SLP.lean index 552243b..f7de9e5 100644 --- a/Lampe/Lampe/SeparationLogic/SLP.lean +++ b/Lampe/Lampe/SeparationLogic/SLP.lean @@ -121,8 +121,7 @@ theorem true_star [LawfulHeap α] {H : SLP α} : (⟦⟧ ⋆ H) = H := by · simp_all · intro exists ∅, ?_ - simp_all [LawfulHeap.disjoint_empty] - apply LawfulHeap.disjoint_empty + simp_all @[simp] theorem star_true [LawfulHeap α] {H : SLP α} : (H ⋆ ⟦⟧) = H := by rw [star_comm]; simp @@ -169,8 +168,7 @@ theorem ent_star_top [LawfulHeap α] {H : SLP α} : H ⊢ H ⋆ ⊤ := by intro _ _ exists ?_, ∅ rw [LawfulHeap.disjoint_symm_iff] - simp_all [LawfulHeap.disjoint_empty] - apply LawfulHeap.disjoint_empty + simp_all theorem star_mono_r [LawfulHeap α] {P Q R : SLP α} : (P ⊢ Q) → (P ⋆ R ⊢ Q ⋆ R) := by unfold star entails @@ -210,8 +208,14 @@ theorem top_star_top [LawfulHeap α] : (top ⋆ (⊤ : SLP α)) = (⊤ : SLP α) funext x simp exists ∅, x - simp [LawfulHeap.disjoint_empty] - apply LawfulHeap.disjoint_empty + simp + +theorem ent_drop_left [LawfulHeap α] {Q H : SLP α} : Q ⋆ H ⊢ Q ⋆ ⊤ := by + unfold star entails + intros st h + obtain ⟨s₁, ⟨s₂, _⟩⟩ := h + use s₁, s₂ + repeat apply And.intro <;> tauto end star From 0edfe4be07478ef980803df2e0e6795199376f22 Mon Sep 17 00:00:00 2001 From: utkn Date: Wed, 27 Nov 2024 11:25:30 +0300 Subject: [PATCH 29/39] renamed utkans_thm to nested_triple --- Lampe/Lampe/Hoare/SepTotal.lean | 4 +--- Lampe/Lampe/Tactic/SeparationLogic.lean | 2 +- 2 files changed, 2 insertions(+), 4 deletions(-) diff --git a/Lampe/Lampe/Hoare/SepTotal.lean b/Lampe/Lampe/Hoare/SepTotal.lean index f44ddbb..2051708 100644 --- a/Lampe/Lampe/Hoare/SepTotal.lean +++ b/Lampe/Lampe/Hoare/SepTotal.lean @@ -91,9 +91,7 @@ theorem consequence_frame_left {H H₁ H₂ : SLP (State p)} rw [SLP.star_comm] apply SLP.ent_star_top - --- [TODO] find a better name for this -theorem utkans_thm {Q : _ → SLP (State p)} +theorem nested_triple {Q : _ → SLP (State p)} (h_hoare_imp : STHoare p Γ P e₁ Q → STHoare p Γ (P ⋆ H) e₂ (fun v => Q v ⋆ H)) (h_hoare : STHoare p Γ P e₁ Q) (h_ent_pre : H ⊢ P ⋆ H) : diff --git a/Lampe/Lampe/Tactic/SeparationLogic.lean b/Lampe/Lampe/Tactic/SeparationLogic.lean index 8b754f2..052ccf4 100644 --- a/Lampe/Lampe/Tactic/SeparationLogic.lean +++ b/Lampe/Lampe/Tactic/SeparationLogic.lean @@ -516,7 +516,7 @@ macro "stephelper1" : tactic => `(tactic|( | apply fresh_intro | apply assert_intro | apply skip_intro - | apply STHoare.utkans_thm STHoare.callLambda_intro + | apply STHoare.nested_triple STHoare.callLambda_intro | apply lam_intro -- memory builtins | apply var_intro From 30aeae3404b68e16578e2fbc55ee8feda55f9e21 Mon Sep 17 00:00:00 2001 From: utkn Date: Thu, 28 Nov 2024 20:22:26 +0300 Subject: [PATCH 30/39] lambda singleton notation changed --- Lampe/Lampe/Hoare/SepTotal.lean | 6 +++--- Lampe/Lampe/SeparationLogic/State.lean | 2 +- Lampe/Lampe/SeparationLogic/ValHeap.lean | 4 ---- Lampe/Lampe/Tactic/SeparationLogic.lean | 6 +++--- 4 files changed, 7 insertions(+), 11 deletions(-) diff --git a/Lampe/Lampe/Hoare/SepTotal.lean b/Lampe/Lampe/Hoare/SepTotal.lean index 2051708..0548411 100644 --- a/Lampe/Lampe/Hoare/SepTotal.lean +++ b/Lampe/Lampe/Hoare/SepTotal.lean @@ -320,9 +320,9 @@ theorem skip_intro : theorem callLambda_intro {lambdaBody} {P : SLP (State p)} {Q : Tp.denote p outTp → SLP (State p)}: @STHoare outTp p Γ P (lambdaBody args) Q → - STHoare p Γ (P ⋆ [ref ↣ ⟨_, argTps, outTp, lambdaBody⟩]) + STHoare p Γ (P ⋆ [λref ↦ ⟨_, argTps, outTp, lambdaBody⟩]) (Expr.call h![] argTps outTp (.lambda ref) args) - (fun v => (Q v) ⋆ [ref ↣ ⟨_, argTps, outTp, lambdaBody⟩]) := by + (fun v => (Q v) ⋆ [λref ↦ ⟨_, argTps, outTp, lambdaBody⟩]) := by intros rename_i h unfold STHoare THoare @@ -353,7 +353,7 @@ theorem callLambda_intro {lambdaBody} {P : SLP (State p)} {Q : Tp.denote p outTp theorem lam_intro : STHoare p Γ ⟦⟧ (.lambda argTps outTp lambdaBody) - fun v => [v ↣ ⟨_, argTps, outTp, lambdaBody⟩] := by + fun v => [λv ↦ ⟨_, argTps, outTp, lambdaBody⟩] := by unfold STHoare THoare intros H st h constructor diff --git a/Lampe/Lampe/SeparationLogic/State.lean b/Lampe/Lampe/SeparationLogic/State.lean index 81ff226..b2e19f3 100644 --- a/Lampe/Lampe/SeparationLogic/State.lean +++ b/Lampe/Lampe/SeparationLogic/State.lean @@ -88,7 +88,7 @@ notation:max "[" l " ↦ " r "]" => State.valSingleton l r def State.lmbSingleton (r : Ref) (v : Lambda) : SLP (State p) := fun st => st.lambdas = Finmap.singleton r v -notation:max "[" l " ↣ " r "]" => State.lmbSingleton l r +notation:max "[" "λ" l " ↦ " r "]" => State.lmbSingleton l r @[simp] lemma State.union_parts_left : diff --git a/Lampe/Lampe/SeparationLogic/ValHeap.lean b/Lampe/Lampe/SeparationLogic/ValHeap.lean index 7a59596..81c7f53 100644 --- a/Lampe/Lampe/SeparationLogic/ValHeap.lean +++ b/Lampe/Lampe/SeparationLogic/ValHeap.lean @@ -40,8 +40,4 @@ instance : LawfulHeap (ValHeap p) where disjoint_union_left := Finmap.disjoint_union_left disjoint_empty := Finmap.disjoint_empty --- def ValHeap.singleton (r : Ref) (v : AnyValue p) : SLP (ValHeap p) := fun st => st = Finmap.singleton r v - --- notation:max "[" l " ↦ " r "]" => ValHeap.singleton l r - end Lampe diff --git a/Lampe/Lampe/Tactic/SeparationLogic.lean b/Lampe/Lampe/Tactic/SeparationLogic.lean index 052ccf4..e3af14f 100644 --- a/Lampe/Lampe/Tactic/SeparationLogic.lean +++ b/Lampe/Lampe/Tactic/SeparationLogic.lean @@ -29,7 +29,7 @@ def SLTerm.toString : SLTerm → String | star _ a b => s!"({a.toString} ⋆ {b.toString})" | lift e => s!"⟦{e.dbgToString}⟧" | singleton e₁ _ => s!"[{e₁.dbgToString} ↦ _]" -| lmbSingleton e₁ _ => s!"[{e₁.dbgToString} ↣ _]" +| lmbSingleton e₁ _ => s!"[λ {e₁.dbgToString} ↦ _]" | mvar e => s!"MV{e.dbgToString}" | unrecognized e => s!"" @@ -340,7 +340,7 @@ theorem singleton_congr_mv {p} {r} {v₁ v₂ : AnyValue p} : (v₁ = v₂) → simp apply SLP.entails_self -theorem lmbSingleton_congr_mv {p} {r} {l₁ l₂ : Lambda} : (l₁ = l₂) → (([r ↣ l₁] : SLP (State p)) ⊢ [r ↣ l₂] ⋆ ⟦⟧) := by +theorem lmbSingleton_congr_mv {p} {r} {l₁ l₂ : Lambda} : (l₁ = l₂) → (([λr ↦ l₁] : SLP (State p)) ⊢ [λr ↦ l₂] ⋆ ⟦⟧) := by rintro rfl simp apply SLP.entails_self @@ -351,7 +351,7 @@ theorem singleton_star_congr {p} {r} {v₁ v₂ : AnyValue p} {R} : (v₁ = v₂ apply SLP.entails_self theorem lmbSingleton_star_congr {p} {r} {v₁ v₂ : Lambda} {R : SLP (State p)} : - (v₁ = v₂) → ([r ↣ v₁] ⋆ R ⊢ [r ↣ v₂] ⋆ R) := by + (v₁ = v₂) → ([λr ↦ v₁] ⋆ R ⊢ [λr ↦ v₂] ⋆ R) := by rintro rfl apply SLP.entails_self From fca5580f32b76976509222c9f2c3fab8ca89b159 Mon Sep 17 00:00:00 2001 From: utkn Date: Thu, 28 Nov 2024 20:28:44 +0300 Subject: [PATCH 31/39] variable p reintroduced to semantics. functionident's rep parameter is explicit now --- Lampe/Lampe/Ast.lean | 12 ++++---- Lampe/Lampe/Semantics.lean | 56 +++++++++++++++++++------------------- 2 files changed, 34 insertions(+), 34 deletions(-) diff --git a/Lampe/Lampe/Ast.lean b/Lampe/Lampe/Ast.lean index 1b7b8bb..c1b9119 100644 --- a/Lampe/Lampe/Ast.lean +++ b/Lampe/Lampe/Ast.lean @@ -25,11 +25,11 @@ structure TraitMethodImplRef where trait : TraitImplRef method : Ident -inductive FunctionIdent {rep : Tp → Type} : Type where -| builtin : Builtin → FunctionIdent -| decl : Ident → FunctionIdent -| lambda : rep .lambdaRef → FunctionIdent -| trait : TraitMethodImplRef → FunctionIdent +inductive FunctionIdent (rep : Tp → Type) : Type where +| builtin : Builtin → FunctionIdent rep +| decl : Ident → FunctionIdent rep +| lambda : rep .lambdaRef → FunctionIdent rep +| trait : TraitMethodImplRef → FunctionIdent rep inductive Member : Tp → List Tp → Type where | head : Member tp (tp :: tps) @@ -39,7 +39,7 @@ inductive Expr (rep : Tp → Type) : Tp → Type where | lit : (tp : Tp) → Nat → Expr rep tp | var : rep tp → Expr rep tp | letIn : Expr rep t₁ → (rep t₁ → Expr rep t₂) → Expr rep t₂ -| call : HList Kind.denote tyKinds → (argTypes : List Tp) → (res : Tp) → @FunctionIdent rep → HList rep argTypes → Expr rep res +| call : HList Kind.denote tyKinds → (argTypes : List Tp) → (res : Tp) → FunctionIdent rep → HList rep argTypes → Expr rep res | ite : rep .bool → Expr rep a → Expr rep a → Expr rep a | skip : Expr rep .unit | loop : rep (.u s) → rep (.u s) → (rep (.u s) → Expr rep r) → Expr rep .unit diff --git a/Lampe/Lampe/Semantics.lean b/Lampe/Lampe/Semantics.lean index ba418e6..12126c4 100644 --- a/Lampe/Lampe/Semantics.lean +++ b/Lampe/Lampe/Semantics.lean @@ -8,7 +8,7 @@ import Lampe.SeparationLogic.State namespace Lampe -variable (P : Prime) +variable (p : Prime) structure Env where functions : List (Ident × Function) @@ -34,57 +34,57 @@ inductive TraitResolution (Γ : Env): TraitImplRef → List (Ident × Function) (∀constraint ∈ impl.constraints implGenerics, TraitResolvable Γ constraint) → TraitResolution Γ ref (impl.impl implGenerics) -inductive Omni : (p : Prime) → Env → State p → Expr (Tp.denote p) tp → (Option (State p × Tp.denote p tp) → Prop) → Prop where -| litField {Q} : Q (some (st, n)) → Omni p Γ st (.lit .field n) Q -| litFalse {Q} : Q (some (st, false)) → Omni p Γ st (.lit .bool 0) Q -| litTrue {Q} : Q (some (st, true)) → Omni p Γ st (.lit .bool 1) Q -| litRef {Q} : Q (some (st, ⟨r⟩)) → Omni p Γ st (.lit (.ref tp) r) Q -| litU {Q} : Q (some (st, ↑n)) → Omni p Γ st (.lit (.u s) n) Q -| var {Q} : Q (some (st, v)) → Omni p Γ st (.var v) Q -| skip {Q} : Q (some (st, ())) → Omni p Γ st (.skip) Q +inductive Omni : Env → State p → Expr (Tp.denote p) tp → (Option (State p × Tp.denote p tp) → Prop) → Prop where +| litField {Q} : Q (some (st, n)) → Omni Γ st (.lit .field n) Q +| litFalse {Q} : Q (some (st, false)) → Omni Γ st (.lit .bool 0) Q +| litTrue {Q} : Q (some (st, true)) → Omni Γ st (.lit .bool 1) Q +| litRef {Q} : Q (some (st, ⟨r⟩)) → Omni Γ st (.lit (.ref tp) r) Q +| litU {Q} : Q (some (st, ↑n)) → Omni Γ st (.lit (.u s) n) Q +| var {Q} : Q (some (st, v)) → Omni Γ st (.var v) Q +| skip {Q} : Q (some (st, ())) → Omni Γ st (.skip) Q | iteTrue {mainBranch elseBranch} : - Omni p Γ st mainBranch Q → - Omni p Γ st (Expr.ite true mainBranch elseBranch) Q + Omni Γ st mainBranch Q → + Omni Γ st (Expr.ite true mainBranch elseBranch) Q | iteFalse {mainBranch elseBranch} : - Omni p Γ st elseBranch Q → - Omni p Γ st (Expr.ite false mainBranch elseBranch) Q + Omni Γ st elseBranch Q → + Omni Γ st (Expr.ite false mainBranch elseBranch) Q | letIn : - Omni p Γ st e Q₁ → - (∀v st, Q₁ (some (st, v)) → Omni p Γ st (b v) Q) → + Omni Γ st e Q₁ → + (∀v st, Q₁ (some (st, v)) → Omni Γ st (b v) Q) → (Q₁ none → Q none) → - Omni p Γ st (.letIn e b) Q + Omni Γ st (.letIn e b) Q | callLambda {lambdas : Lambdas} : lambdas.lookup ref = some ⟨Tp.denote p, argTps, outTp, lambdaBody⟩ → - Omni p Γ ⟨vh, lambdas⟩ (lambdaBody args) Q → - Omni p Γ ⟨vh, lambdas⟩ (Expr.call h![] argTps outTp (.lambda ref) args) Q + Omni Γ ⟨vh, lambdas⟩ (lambdaBody args) Q → + Omni Γ ⟨vh, lambdas⟩ (Expr.call h![] argTps outTp (.lambda ref) args) Q | lam {Q} : (∀ ref, ref ∉ lambdas → Q (some (⟨vh, lambdas.insert ref ⟨_, argTps, outTp, lambdaBody⟩⟩, ref))) → - Omni p Γ ⟨vh, lambdas⟩ (Expr.lambda argTps outTp lambdaBody) Q + Omni Γ ⟨vh, lambdas⟩ (Expr.lambda argTps outTp lambdaBody) Q | callBuiltin {Q} : (b.omni p st argTypes resType args (mapToValHeapCondition st.lambdas Q)) → - Omni p Γ st (Expr.call h![] argTypes resType (.builtin b) args) Q + Omni Γ st (Expr.call h![] argTypes resType (.builtin b) args) Q | callDecl: (fname, fn) ∈ Γ.functions → (hkc : fn.generics = tyKinds) → (htci : fn.inTps (hkc ▸ generics) = argTypes) → (htco : fn.outTp (hkc ▸ generics) = res) → - Omni p Γ st (htco ▸ fn.body _ (hkc ▸ generics) (htci ▸ args)) Q → - Omni p Γ st (@Expr.call _ tyKinds generics argTypes res (.decl fname) args) Q + Omni Γ st (htco ▸ fn.body _ (hkc ▸ generics) (htci ▸ args)) Q → + Omni Γ st (@Expr.call _ tyKinds generics argTypes res (.decl fname) args) Q | callTrait {impl}: TraitResolution Γ traitRef impl → (fname, fn) ∈ impl → (hkc : fn.generics = tyKinds) → (htci : fn.inTps (hkc ▸ generics) = argTypes) → (htco : fn.outTp (hkc ▸ generics) = res) → - Omni p Γ st (htco ▸ fn.body _ (hkc ▸ generics) (htci ▸ args)) Q → - Omni p Γ st (@Expr.call _ tyKinds generics argTypes res (.trait ⟨traitRef, fname⟩) args) Q + Omni Γ st (htco ▸ fn.body _ (hkc ▸ generics) (htci ▸ args)) Q → + Omni Γ st (@Expr.call _ tyKinds generics argTypes res (.trait ⟨traitRef, fname⟩) args) Q | loopDone : lo ≥ hi → - Omni p Γ st (.loop lo hi body) Q + Omni Γ st (.loop lo hi body) Q | loopNext {s} {lo hi : U s} {body} : lo < hi → - Omni p Γ st (.letIn (body lo) (fun _ => .loop (lo + 1) hi body)) Q → - Omni p Γ st (.loop lo hi body) Q + Omni Γ st (.letIn (body lo) (fun _ => .loop (lo + 1) hi body)) Q → + Omni Γ st (.loop lo hi body) Q theorem Omni.consequence {p Γ st tp} {e : Expr (Tp.denote p) tp} {Q Q'}: Omni p Γ st e Q → @@ -145,7 +145,7 @@ theorem Omni.frame {p Γ tp} {st₁ st₂ : State p} {e : Expr (Tp.denote p) tp} intros constructor simp only [State.union_closures, State.union_vals] - rename_i _ _ st₁ _ _ _ _ hd + rename_i _ st₁ _ _ _ _ hd have hf := b.frame hq (st₂ := st₂) unfold mapToValHeapCondition at * simp_all only [LawfulHeap.disjoint, true_implies] From 31217dd05147c824d49abfc480dcb0432a9dbbf5 Mon Sep 17 00:00:00 2001 From: utkn Date: Thu, 28 Nov 2024 21:33:20 +0300 Subject: [PATCH 32/39] Lambda syntax updated --- Lampe/Lampe.lean | 2 +- Lampe/Lampe/Syntax.lean | 20 +++++++++++++------- 2 files changed, 14 insertions(+), 8 deletions(-) diff --git a/Lampe/Lampe.lean b/Lampe/Lampe.lean index 6955b94..cc4ed2f 100644 --- a/Lampe/Lampe.lean +++ b/Lampe/Lampe.lean @@ -93,7 +93,7 @@ example {p Γ x y}: STHoare p Γ ⟦⟧ (simple_if_else.fn.body (Tp.denote p) h! . aesop nr_def simple_lambda<>(x : Field, y : Field) -> Field { - let add = |a, b|: Field, Field -> Field { #add(a, b) : Field }; + let add = |a : Field, b : Field| -> Field { #add(a, b) : Field }; ^add(x, y) : Field; } diff --git a/Lampe/Lampe/Syntax.lean b/Lampe/Lampe/Syntax.lean index d344258..b7d8802 100644 --- a/Lampe/Lampe/Syntax.lean +++ b/Lampe/Lampe/Syntax.lean @@ -12,6 +12,7 @@ declare_syntax_cat nr_ident declare_syntax_cat nr_type declare_syntax_cat nr_expr declare_syntax_cat nr_block_contents +declare_syntax_cat nr_param_decl syntax ident:nr_ident syntax ident"::"nr_ident : nr_ident @@ -82,6 +83,8 @@ partial def mkBuiltin [Monad m] [MonadQuotation m] [MonadExceptOf Exception m] [ | "write_ref" => ``(Builtin.writeRef) | _ => throwError "Unknown builtin {i}" +syntax ident ":" nr_type : nr_param_decl + syntax num ":" nr_type : nr_expr syntax ident : nr_expr syntax "#" nr_ident "(" nr_expr,* ")" ":" nr_type : nr_expr @@ -98,9 +101,11 @@ syntax "if" nr_expr nr_expr ("else" nr_expr)? : nr_expr syntax "for" ident "in" nr_expr ".." nr_expr nr_expr : nr_expr syntax "(" nr_expr ")" : nr_expr syntax "*(" nr_expr ")" : nr_expr -syntax "|" ident,* "|:" nr_type,* "->" nr_type nr_expr : nr_expr +syntax "|" nr_param_decl,* "|" "->" nr_type nr_expr : nr_expr syntax "^" nr_ident "(" nr_expr,* ")" ":" nr_type : nr_expr +syntax nr_fn_decl := nr_ident "<" ident,* ">" "(" nr_param_decl,* ")" "->" nr_type "{" sepBy(nr_expr, ";", ";", allowTrailingSep) "}" + def Expr.letMutIn (definition : Expr rep tp) (body : rep tp.ref → Expr rep tp'): Expr rep tp' := let refDef := Expr.letIn definition fun v => Expr.call h![] _ (tp.ref) (.builtin .ref) h![v] Expr.letIn refDef body @@ -198,10 +203,14 @@ partial def mkExpr [MonadSyntax m] (e : TSyntax `nr_expr) (vname : Option Lean.I mkExpr cond none fun cond => do let mainBody ← mkExpr mainBody none fun x => ``(Lampe.Expr.var $x) wrapSimple (←`(Lampe.Expr.ite $cond $mainBody (Lampe.Expr.skip))) vname k -| `(nr_expr| |$args,*|: $argTps,* -> $outTp $lambdaBody) => do +| `(nr_expr| | $params,* | -> $outTp $lambdaBody) => do let outTp ← mkNrType outTp - let argTps ← mkListLit (← argTps.getElems.toList.mapM mkNrType) - let args ← mkHListLit args.getElems.toList + let argTps ← mkListLit (← params.getElems.toList.mapM fun param => match param with + | `(nr_param_decl|$_:ident : $tp) => mkNrType tp + | _ => throwUnsupportedSyntax) + let args ← mkHListLit (← params.getElems.toList.mapM fun param => match param with + | `(nr_param_decl|$i:ident : $_) => `($i) + | _ => throwUnsupportedSyntax) let body ← mkExpr lambdaBody none fun x => ``(Lampe.Expr.var $x) wrapSimple (←`(Lampe.Expr.lambda $argTps $outTp (fun $args => $body))) vname k | `(nr_expr| ^ $i:ident ($args,*): $tp) => do @@ -211,9 +220,6 @@ partial def mkExpr [MonadSyntax m] (e : TSyntax `nr_expr) (vname : Option Lean.I end -declare_syntax_cat nr_param_decl -syntax ident ":" nr_type : nr_param_decl -syntax nr_fn_decl := nr_ident "<" ident,* ">" "(" nr_param_decl,* ")" "->" nr_type "{" sepBy(nr_expr, ";", ";", allowTrailingSep) "}" def mkFnDecl [Monad m] [MonadQuotation m] [MonadExceptOf Exception m] [MonadError m] : Syntax → m (String × TSyntax `term) | `(nr_fn_decl| $name < $generics,* > ( $params,* ) -> $outTp { $bExprs;* }) => do From 0c6af881aaa37b41f1a9b059ec30efd023380aca Mon Sep 17 00:00:00 2001 From: utkn Date: Thu, 28 Nov 2024 22:58:49 +0300 Subject: [PATCH 33/39] lawfulheap lemmas moved into namespace. class props now have a "thm_" prefix to distinguish them --- Lampe/Lampe/Hoare/Builtins.lean | 4 +- Lampe/Lampe/SeparationLogic/LawfulHeap.lean | 76 +++++++++++---------- Lampe/Lampe/SeparationLogic/SLP.lean | 24 +++---- Lampe/Lampe/SeparationLogic/State.lean | 12 ++-- Lampe/Lampe/SeparationLogic/ValHeap.lean | 12 ++-- Lampe/Lampe/Tactic/SeparationLogic.lean | 2 +- 6 files changed, 67 insertions(+), 63 deletions(-) diff --git a/Lampe/Lampe/Hoare/Builtins.lean b/Lampe/Lampe/Hoare/Builtins.lean index 7fc45c8..0a7e07c 100644 --- a/Lampe/Lampe/Hoare/Builtins.lean +++ b/Lampe/Lampe/Hoare/Builtins.lean @@ -356,9 +356,9 @@ theorem ref_intro: simp only [SLP.true_star] intro st hH r hr exists (⟨Finmap.singleton r ⟨tp, v⟩, st.lambdas⟩ ∪ st), ∅ - apply And.intro (by rw [LawfulHeap.disjoint_symm_iff]; apply LawfulHeap.disjoint_empty) + apply And.intro (by rw [LawfulHeap.disjoint_symm_iff]; apply LawfulHeap.empty_disjoint) constructor - . simp only [State.insertVal, Finmap.insert_eq_singleton_union, LawfulHeap_union_empty] + . simp only [State.insertVal, Finmap.insert_eq_singleton_union, LawfulHeap.union_empty] simp only [State.union_parts_left, Finmap.union_self] . apply And.intro ?_ (by simp) exists (⟨Finmap.singleton r ⟨tp, v⟩, ∅⟩), st diff --git a/Lampe/Lampe/SeparationLogic/LawfulHeap.lean b/Lampe/Lampe/SeparationLogic/LawfulHeap.lean index 37feccb..a1f2ff1 100644 --- a/Lampe/Lampe/SeparationLogic/LawfulHeap.lean +++ b/Lampe/Lampe/SeparationLogic/LawfulHeap.lean @@ -7,69 +7,73 @@ class LawfulHeap (α : Type _) where disjoint : α → α → Prop empty : α - union_empty {a : α} : union a empty = a - union_assoc {s₁ s₂ s₃ : α} : union (union s₁ s₂) s₃ = union s₁ (union s₂ s₃) - disjoint_symm_iff {a b : α} : disjoint a b ↔ disjoint b a - union_comm_of_disjoint {s₁ s₂ : α} : disjoint s₁ s₂ → union s₁ s₂ = union s₂ s₁ - disjoint_union_left (x y z : α) : disjoint (union x y) z ↔ disjoint x z ∧ disjoint y z - disjoint_empty (x : α) : disjoint empty x + private thm_union_empty {a : α} : union a empty = a + private thm_union_assoc {s₁ s₂ s₃ : α} : union (union s₁ s₂) s₃ = union s₁ (union s₂ s₃) + private thm_disjoint_symm_iff {a b : α} : disjoint a b ↔ disjoint b a + private thm_union_comm_of_disjoint {s₁ s₂ : α} : disjoint s₁ s₂ → union s₁ s₂ = union s₂ s₁ + private thm_disjoint_union_left (x y z : α) : disjoint (union x y) z ↔ disjoint x z ∧ disjoint y z + private thm_disjoint_empty (x : α) : disjoint empty x instance [LawfulHeap α] : Union α := ⟨LawfulHeap.union⟩ instance [LawfulHeap α] : EmptyCollection α := ⟨LawfulHeap.empty⟩ -theorem LawfulHeap_union_empty_commutative [LawfulHeap α] {a : α} : - a ∪ ∅ = ∅ ∪ a := by - have h := LawfulHeap.disjoint_empty a - simp [Union.union, EmptyCollection.emptyCollection] - rw [LawfulHeap.union_comm_of_disjoint h] +theorem LawfulHeap.disjoint_symm_iff [LawfulHeap α] {s₁ s₂ : α} : + LawfulHeap.disjoint s₁ s₂ ↔ LawfulHeap.disjoint s₂ s₁ := by + exact LawfulHeap.thm_disjoint_symm_iff -theorem LawfulHeap_union_comm_of_disjoint [LawfulHeap α] {s₁ s₂ : α}: +theorem LawfulHeap.union_comm_of_disjoint_ [LawfulHeap α] {s₁ s₂ : α} : LawfulHeap.disjoint s₁ s₂ → s₁ ∪ s₂ = s₂ ∪ s₁ := by simp [Union.union] - exact LawfulHeap.union_comm_of_disjoint + exact LawfulHeap.thm_union_comm_of_disjoint + +theorem LawfulHeap.union_empty_comm [LawfulHeap α] {a : α} : + a ∪ ∅ = ∅ ∪ a := by + have h := LawfulHeap.thm_disjoint_empty a + simp [Union.union, EmptyCollection.emptyCollection] + rw [LawfulHeap.thm_union_comm_of_disjoint h] @[simp] -theorem LawfulHeap_disjoint_empty [LawfulHeap α] {a : α} : - LawfulHeap.disjoint a ∅ := by - rw [LawfulHeap.disjoint_symm_iff] - apply LawfulHeap.disjoint_empty +theorem LawfulHeap.union_empty [LawfulHeap α] {a : α} : + a ∪ ∅ = a := by + apply LawfulHeap.thm_union_empty @[simp] -theorem LawfulHeap_empty_disjoint [LawfulHeap α] {a : α} : - LawfulHeap.disjoint ∅ a := by - apply LawfulHeap.disjoint_empty +theorem LawfulHeap.empty_union [LawfulHeap α] {a : α} : + ∅ ∪ a = a := by + rw [←LawfulHeap.union_empty_comm] + apply LawfulHeap.thm_union_empty @[simp] -theorem LawfulHeap_union_empty [LawfulHeap α] {a : α} : - a ∪ ∅ = a := by - apply LawfulHeap.union_empty +theorem LawfulHeap.disjoint_empty [LawfulHeap α] {a : α} : + LawfulHeap.disjoint a ∅ := by + rw [LawfulHeap.thm_disjoint_symm_iff] + apply LawfulHeap.thm_disjoint_empty @[simp] -theorem LawfulHeap_empty_union [LawfulHeap α] {a : α} : - ∅ ∪ a = a := by - rw [←LawfulHeap_union_empty_commutative] - apply LawfulHeap.union_empty +theorem LawfulHeap.empty_disjoint [LawfulHeap α] {a : α} : + LawfulHeap.disjoint ∅ a := by + apply LawfulHeap.thm_disjoint_empty -theorem LawfulHeap_union_assoc [LawfulHeap α] {s₁ s₂ s₃ : α} : +theorem LawfulHeap.union_assoc [LawfulHeap α] {s₁ s₂ s₃ : α} : s₁ ∪ s₂ ∪ s₃ = s₁ ∪ (s₂ ∪ s₃) := by simp [Union.union] - apply LawfulHeap.union_assoc + apply LawfulHeap.thm_union_assoc -theorem LawfulHeap_disjoint_union_left [LawfulHeap α] (x y z : α) : +theorem LawfulHeap.disjoint_union_left [LawfulHeap α] (x y z : α) : LawfulHeap.disjoint (x ∪ y) z ↔ LawfulHeap.disjoint x z ∧ LawfulHeap.disjoint y z := by simp [Union.union] - apply LawfulHeap.disjoint_union_left + apply LawfulHeap.thm_disjoint_union_left -theorem LawfulHeap_disjoint_union_right [LawfulHeap α] (x y z : α) : +theorem LawfulHeap.disjoint_union_right [LawfulHeap α] (x y z : α) : LawfulHeap.disjoint x (y ∪ z) ↔ LawfulHeap.disjoint x y ∧ LawfulHeap.disjoint x z := by conv => lhs - rw [LawfulHeap.disjoint_symm_iff] + rw [LawfulHeap.thm_disjoint_symm_iff] conv => rhs - rw [LawfulHeap.disjoint_symm_iff] + rw [LawfulHeap.thm_disjoint_symm_iff] rhs - rw [LawfulHeap.disjoint_symm_iff] - apply LawfulHeap.disjoint_union_left + rw [LawfulHeap.thm_disjoint_symm_iff] + apply LawfulHeap.thm_disjoint_union_left end Lampe diff --git a/Lampe/Lampe/SeparationLogic/SLP.lean b/Lampe/Lampe/SeparationLogic/SLP.lean index f7de9e5..a0000f3 100644 --- a/Lampe/Lampe/SeparationLogic/SLP.lean +++ b/Lampe/Lampe/SeparationLogic/SLP.lean @@ -76,8 +76,8 @@ theorem pure_right [LawfulHeap α] {H₁ H₂ : SLP α} : P → (H₁ ⊢ H₂) apply And.intro rfl apply_assumption assumption - . simp only [LawfulHeap_empty_union] - . apply LawfulHeap.disjoint_empty + . simp only [LawfulHeap.empty_union] + . apply LawfulHeap.empty_disjoint theorem entails_self [LawfulHeap α] {H : SLP α} : H ⊢ H := by tauto @@ -108,7 +108,7 @@ theorem star_comm [LawfulHeap α] {G H : SLP α} : (G ⋆ H) = (H ⋆ G) := by repeat apply Exists.intro rw [LawfulHeap.disjoint_symm_iff] apply And.intro (by assumption) - rw [LawfulHeap_union_comm_of_disjoint (by rw [LawfulHeap.disjoint_symm_iff]; assumption)] + rw [LawfulHeap.union_comm_of_disjoint_ (by rw [LawfulHeap.disjoint_symm_iff]; assumption)] tauto } @@ -134,8 +134,8 @@ theorem star_assoc [LawfulHeap α] {F G H : SLP α} : ((F ⋆ G) ⋆ H) = (F ⋆ apply Iff.intro · intro_cases subst_vars - rw [LawfulHeap_union_assoc] - simp only [LawfulHeap_disjoint_union_left] at * + rw [LawfulHeap.union_assoc] + simp only [LawfulHeap.disjoint_union_left] at * cases_type And repeat apply Exists.intro apply And.intro ?_ @@ -146,11 +146,11 @@ theorem star_assoc [LawfulHeap α] {F G H : SLP α} : ((F ⋆ G) ⋆ H) = (F ⋆ apply And.intro rfl simp_all assumption - simp_all [LawfulHeap_disjoint_union_right] + simp_all [LawfulHeap.disjoint_union_right] · intro_cases subst_vars - rw [←LawfulHeap_union_assoc] - simp only [LawfulHeap_disjoint_union_right] at * + rw [←LawfulHeap.union_assoc] + simp only [LawfulHeap.disjoint_union_right] at * cases_type And repeat apply Exists.intro apply And.intro ?_ @@ -161,7 +161,7 @@ theorem star_assoc [LawfulHeap α] {F G H : SLP α} : ((F ⋆ G) ⋆ H) = (F ⋆ apply And.intro rfl simp_all assumption - simp_all [LawfulHeap_disjoint_union_left] + simp_all [LawfulHeap.disjoint_union_left] @[simp] theorem ent_star_top [LawfulHeap α] {H : SLP α} : H ⊢ H ⋆ ⊤ := by @@ -188,7 +188,7 @@ theorem star_mono_l' [LawfulHeap α] {P Q : SLP α} : (⟦⟧ ⊢ Q) → (P ⊢ tauto simp rw [LawfulHeap.disjoint_symm_iff] - apply LawfulHeap.disjoint_empty + apply LawfulHeap.empty_disjoint theorem star_mono [LawfulHeap α] {H₁ H₂ Q₁ Q₂ : SLP α} : (H₁ ⊢ H₂) → (Q₁ ⊢ Q₂) → (H₁ ⋆ Q₁ ⊢ H₂ ⋆ Q₂) := by unfold star entails @@ -239,7 +239,7 @@ theorem wand_self_star [LawfulHeap α] {H : SLP α}: (H -⋆ H ⋆ top) = top := simp rotate_left rotate_left - rw [LawfulHeap_union_comm_of_disjoint (by assumption)] + rw [LawfulHeap.union_comm_of_disjoint_ (by assumption)] rw [LawfulHeap.disjoint_symm_iff] assumption @@ -256,7 +256,7 @@ theorem wand_cancel [LawfulHeap α] {P Q : SLP α} : (P ⋆ (P -⋆ Q)) ⊢ Q := intro_cases subst_vars rename_i h - rw [LawfulHeap_union_comm_of_disjoint (by assumption)] + rw [LawfulHeap.union_comm_of_disjoint_ (by assumption)] apply_assumption rw [LawfulHeap.disjoint_symm_iff] tauto diff --git a/Lampe/Lampe/SeparationLogic/State.lean b/Lampe/Lampe/SeparationLogic/State.lean index b2e19f3..f318e3d 100644 --- a/Lampe/Lampe/SeparationLogic/State.lean +++ b/Lampe/Lampe/SeparationLogic/State.lean @@ -49,17 +49,17 @@ instance : LawfulHeap (State p) where union := fun a b => ⟨a.vals ∪ b.vals, a.lambdas ∪ b.lambdas⟩ disjoint := fun a b => a.vals.Disjoint b.vals ∧ a.lambdas.Disjoint b.lambdas empty := ⟨∅, ∅⟩ - union_empty := by + thm_union_empty := by intros simp only [Finmap.union_empty] - union_assoc := by + thm_union_assoc := by intros simp only [Union.union, State.mk.injEq] apply And.intro . apply Finmap.union_assoc . apply Finmap.union_assoc - disjoint_symm_iff := by tauto - union_comm_of_disjoint := by + thm_disjoint_symm_iff := by tauto + thm_union_comm_of_disjoint := by intros simp only [Union.union, State.mk.injEq] apply And.intro @@ -67,8 +67,8 @@ instance : LawfulHeap (State p) where tauto . apply Finmap.union_comm_of_disjoint tauto - disjoint_empty := by tauto - disjoint_union_left := by + thm_disjoint_empty := by tauto + thm_disjoint_union_left := by intros x y z have h1 := (Finmap.disjoint_union_left x.vals y.vals z.vals) have h2 := (Finmap.disjoint_union_left x.lambdas y.lambdas z.lambdas) diff --git a/Lampe/Lampe/SeparationLogic/ValHeap.lean b/Lampe/Lampe/SeparationLogic/ValHeap.lean index 81c7f53..3ffe117 100644 --- a/Lampe/Lampe/SeparationLogic/ValHeap.lean +++ b/Lampe/Lampe/SeparationLogic/ValHeap.lean @@ -33,11 +33,11 @@ instance : LawfulHeap (ValHeap p) where union := fun a b => a ∪ b disjoint := fun a b => a.Disjoint b empty := ∅ - union_empty := Finmap.union_empty - union_assoc := Finmap.union_assoc - disjoint_symm_iff := by tauto - union_comm_of_disjoint := Finmap.union_comm_of_disjoint - disjoint_union_left := Finmap.disjoint_union_left - disjoint_empty := Finmap.disjoint_empty + thm_union_empty := Finmap.union_empty + thm_union_assoc := Finmap.union_assoc + thm_disjoint_symm_iff := by tauto + thm_union_comm_of_disjoint := Finmap.union_comm_of_disjoint + thm_disjoint_union_left := Finmap.disjoint_union_left + thm_disjoint_empty := Finmap.disjoint_empty end Lampe diff --git a/Lampe/Lampe/Tactic/SeparationLogic.lean b/Lampe/Lampe/Tactic/SeparationLogic.lean index e3af14f..0c74dc1 100644 --- a/Lampe/Lampe/Tactic/SeparationLogic.lean +++ b/Lampe/Lampe/Tactic/SeparationLogic.lean @@ -317,7 +317,7 @@ theorem exi_prop_l [LawfulHeap α] {P : Prop} {H : P → SLP α} {Q : SLP α} : apply h use ∅, st refine ⟨?_, ?_, ?_, ?_⟩ - apply LawfulHeap.disjoint_empty + apply LawfulHeap.empty_disjoint all_goals simp_all [LawfulHeap.disjoint_empty, SLP.lift] theorem use_right [LawfulHeap α] {R L G H : SLP α} : (R ⊢ G ⋆ H) → (L ⋆ R ⊢ G ⋆ L ⋆ H) := by From 4b3871973cd412ab9f246cf1e3328d9af1543cb0 Mon Sep 17 00:00:00 2001 From: utkn Date: Fri, 29 Nov 2024 18:53:39 +0300 Subject: [PATCH 34/39] Function and lambda types combined. `Rep` is now a parameter of functions. Accordingly, nrfn! macro had to be changed --- Lampe/Lampe.lean | 33 +++++++++++++++++---------------- Lampe/Lampe/Ast.lean | 24 +++++++++--------------- Lampe/Lampe/Hoare/SepTotal.lean | 15 ++++++++------- Lampe/Lampe/Semantics.lean | 18 ++++++++++-------- Lampe/Lampe/Syntax.lean | 23 ++++++++++------------- 5 files changed, 54 insertions(+), 59 deletions(-) diff --git a/Lampe/Lampe.lean b/Lampe/Lampe.lean index cc4ed2f..4657fb8 100644 --- a/Lampe/Lampe.lean +++ b/Lampe/Lampe.lean @@ -9,7 +9,7 @@ nr_def simple_muts<>(x : Field) -> Field { y } -example : STHoare p Γ ⟦⟧ (simple_muts.fn.body _ h![] h![x]) fun v => v = x := by +example : STHoare p Γ ⟦⟧ (simple_muts.fn.body h![] |>.body h![x]) fun v => v = x := by simp only [simple_muts] steps try (exact _) @@ -22,7 +22,7 @@ nr_def weirdEq(x : I, y : I) -> Unit { #assert(#eq(a, y) : bool) : Unit; } -example {P} {x y : Tp.denote P .field} : STHoare P Γ ⟦⟧ (weirdEq.fn.body _ h![.field] h![x, y]) fun _ => x = y := by +example {x y : Tp.denote p .field} : STHoare p Γ ⟦⟧ (weirdEq.fn.body h![.field] |>.body h![x, y]) fun _ => x = y := by simp only [weirdEq] steps simp_all @@ -40,7 +40,7 @@ lemma BitVec.add_toNat_of_lt_max {a b : BitVec w} (h: a.toNat + b.toNat < 2^w) : rw [Nat.mod_eq_of_lt] assumption -example {self that : Tp.denote P (.slice tp)} : STHoare P Γ ⟦⟧ (sliceAppend.fn.body _ h![tp] h![self, that]) fun v => v = self ++ that := by +example {self that : Tp.denote p (.slice tp)} : STHoare p Γ ⟦⟧ (sliceAppend.fn.body h![tp] |>.body h![self, that]) fun v => v = self ++ that := by simp only [sliceAppend] steps rename Tp.denote _ tp.slice.ref => selfRef @@ -49,7 +49,7 @@ example {self that : Tp.denote P (.slice tp)} : STHoare P Γ ⟦⟧ (sliceAppend steps have : (i + 1).toNat = i.toNat + 1 := by apply BitVec.add_toNat_of_lt_max - casesm* (Tp.denote P (.u 32)), (U _), Fin _ + casesm* (Tp.denote p (.u 32)), (U _), Fin _ simp at * linarith simp only [this, List.take_succ] @@ -67,7 +67,7 @@ nr_def simple_if<>(x : Field, y : Field) -> Field { z } -example {p Γ x y}: STHoare p Γ ⟦⟧ (simple_if.fn.body (Tp.denote p) h![] h![x, y]) +example {p Γ x y}: STHoare p Γ ⟦⟧ (simple_if.fn.body h![] |>.body h![x, y]) fun v => v = y := by simp only [simple_if] steps <;> tauto @@ -83,7 +83,7 @@ nr_def simple_if_else<>(x : Field, y : Field) -> Field { z } -example {p Γ x y}: STHoare p Γ ⟦⟧ (simple_if_else.fn.body (Tp.denote p) h![] h![x, y]) +example {p Γ x y}: STHoare p Γ ⟦⟧ (simple_if_else.fn.body h![] |>.body h![x, y]) fun v => v = x := by simp only [simple_if_else] steps @@ -97,8 +97,9 @@ nr_def simple_lambda<>(x : Field, y : Field) -> Field { ^add(x, y) : Field; } -example {p Γ x y} : STHoare p Γ ⟦⟧ (simple_lambda.fn.body (Tp.denote p) h![] h![x, y]) - fun v => v = x + y := by +example {p Γ} {x y : Tp.denote p Tp.field} : + STHoare p Γ ⟦⟧ (simple_lambda.fn.body h![] |>.body h![x, y]) + fun v => v = (x + y) := by simp only [simple_lambda] steps simp_all @@ -112,26 +113,26 @@ example {p Γ x y} : STHoare p Γ ⟦⟧ (simple_lambda.fn.body (Tp.denote p) h! sl aesop -def bulbulizeField : TraitImpl := { +def bulbulizeField (rep : Tp → Type) : TraitImpl := { traitGenericKinds := [], implGenericKinds := [], traitGenerics := fun _ => h![], constraints := fun _ => [], self := fun _ => .field, - impl := fun _ => [("bulbulize", nrfn![ + impl := fun _ => [("bulbulize", nrfn![rep; fn bulbulize<>(x : Field) -> Field { #add(x, x) : Field }].2 )] } -def bulbulizeU32 : TraitImpl := { +def bulbulizeU32 (rep : Tp → Type) : TraitImpl := { traitGenericKinds := [], implGenericKinds := [], traitGenerics := fun _ => h![], constraints := fun _ => [], self := fun _ => .u 32, - impl := fun _ => [("bulbulize", nrfn![ + impl := fun _ => [("bulbulize", nrfn![rep; fn bulbulize<>(_x : u32) -> u32 { 69 : u32 }].2 @@ -141,12 +142,12 @@ def bulbulizeU32 : TraitImpl := { def simpleTraitCall (tp : Tp) (arg : tp.denote P): Expr (Tp.denote P) tp := @Expr.call _ [] h![] [tp] tp (.trait ⟨⟨⟨"Bulbulize", [], h![]⟩, tp⟩, "bulbulize"⟩) h![arg] -def simpleTraitEnv : Env := { +def simpleTraitEnv (p : Prime) : Env := { functions := [], - traits := [("Bulbulize", bulbulizeField), ("Bulbulize", bulbulizeU32)] + traits := [("Bulbulize", bulbulizeField (Tp.denote p)), ("Bulbulize", bulbulizeU32 (Tp.denote p))] } -example : STHoare p simpleTraitEnv ⟦⟧ (simpleTraitCall .field arg) (fun v => v = 2 * arg) := by +example : STHoare p (simpleTraitEnv p) ⟦⟧ (simpleTraitCall .field arg) (fun v => v = 2 * arg) := by simp only [simpleTraitCall] apply STHoare.callTrait_intro · constructor @@ -163,7 +164,7 @@ example : STHoare p simpleTraitEnv ⟦⟧ (simpleTraitCall .field arg) (fun v => subst_vars ring -example : STHoare p simpleTraitEnv ⟦⟧ (simpleTraitCall (.u 32) arg) (fun v => v = 69) := by +example : STHoare p (simpleTraitEnv p) ⟦⟧ (simpleTraitCall (.u 32) arg) (fun v => v = 69) := by simp only [simpleTraitCall] apply STHoare.callTrait_intro · constructor diff --git a/Lampe/Lampe/Ast.lean b/Lampe/Lampe/Ast.lean index c1b9119..460c2f2 100644 --- a/Lampe/Lampe/Ast.lean +++ b/Lampe/Lampe/Ast.lean @@ -45,28 +45,22 @@ inductive Expr (rep : Tp → Type) : Tp → Type where | loop : rep (.u s) → rep (.u s) → (rep (.u s) → Expr rep r) → Expr rep .unit | lambda : (argTps : List Tp) → (outTp : Tp) → (HList rep argTps → Expr rep outTp) → Expr rep .lambdaRef +structure Lambda where + argTps : List Tp + outTp : Tp + rep : Tp → Type + body : HList rep argTps → Expr rep outTp + structure Function : Type _ where generics : List Kind - inTps : HList Kind.denote generics → List Tp - outTp : HList Kind.denote generics → Tp - body : (rep : Tp → Type) → (gs : HList Kind.denote generics) → - HList rep (inTps gs) → - Expr rep (outTp gs) + body : (HList Kind.denote generics) → Lambda /-- Polymorphic identity -/ -example : Function := { +example {rep : Tp → Type}: Function := { generics := [.type] - inTps := fun h![x] => [x] - outTp := fun h![x] => x - body := fun _ h![_] h![x] => .var x + body := fun h![tp] => ⟨[tp], tp, rep, fun h![x] => .var x⟩ } -structure Lambda where - rep : Tp → Type - argTps : List Tp - outTp : Tp - body : HList rep argTps → Expr rep outTp - structure FunctionDecl where name : Ident fn : Function diff --git a/Lampe/Lampe/Hoare/SepTotal.lean b/Lampe/Lampe/Hoare/SepTotal.lean index 0548411..ff685f5 100644 --- a/Lampe/Lampe/Hoare/SepTotal.lean +++ b/Lampe/Lampe/Hoare/SepTotal.lean @@ -320,9 +320,9 @@ theorem skip_intro : theorem callLambda_intro {lambdaBody} {P : SLP (State p)} {Q : Tp.denote p outTp → SLP (State p)}: @STHoare outTp p Γ P (lambdaBody args) Q → - STHoare p Γ (P ⋆ [λref ↦ ⟨_, argTps, outTp, lambdaBody⟩]) + STHoare p Γ (P ⋆ [λref ↦ ⟨argTps, outTp, _, lambdaBody⟩]) (Expr.call h![] argTps outTp (.lambda ref) args) - (fun v => (Q v) ⋆ [λref ↦ ⟨_, argTps, outTp, lambdaBody⟩]) := by + (fun v => (Q v) ⋆ [λref ↦ ⟨argTps, outTp, _, lambdaBody⟩]) := by intros rename_i h unfold STHoare THoare @@ -353,7 +353,7 @@ theorem callLambda_intro {lambdaBody} {P : SLP (State p)} {Q : Tp.denote p outTp theorem lam_intro : STHoare p Γ ⟦⟧ (.lambda argTps outTp lambdaBody) - fun v => [λv ↦ ⟨_, argTps, outTp, lambdaBody⟩] := by + fun v => [λv ↦ ⟨argTps, outTp, _, lambdaBody⟩] := by unfold STHoare THoare intros H st h constructor @@ -381,10 +381,11 @@ theorem lam_intro : theorem callTrait_intro {impl} {fname fn} (h_trait : TraitResolution Γ traitRef impl) (h_fn : (fname, fn) ∈ impl) - (h_kc : fn.generics = tyKinds) - (h_tci : fn.inTps (h_kc ▸ generics) = argTypes) - (h_tco : fn.outTp (h_kc ▸ generics) = res) - (h_hoare: STHoare p Γ H (h_tco ▸ fn.body _ (h_kc ▸ generics) (h_tci ▸ args)) Q): + (hkc : fn.generics = tyKinds) + (hrep : (fn.body (hkc ▸ generics) |>.rep) = Tp.denote p) + (htci : (fn.body (hkc ▸ generics) |>.argTps) = argTypes) + (htco : (fn.body (hkc ▸ generics) |>.outTp) = res) + (h_hoare: STHoare p Γ H (hrep ▸ htco ▸ (fn.body (hkc ▸ generics) |>.body (hrep ▸ htci ▸ args))) Q): STHoare p Γ H (@Expr.call _ tyKinds generics argTypes res (.trait ⟨traitRef, fname⟩) args) Q := by diff --git a/Lampe/Lampe/Semantics.lean b/Lampe/Lampe/Semantics.lean index 12126c4..eecc916 100644 --- a/Lampe/Lampe/Semantics.lean +++ b/Lampe/Lampe/Semantics.lean @@ -54,11 +54,11 @@ inductive Omni : Env → State p → Expr (Tp.denote p) tp → (Option (State p (Q₁ none → Q none) → Omni Γ st (.letIn e b) Q | callLambda {lambdas : Lambdas} : - lambdas.lookup ref = some ⟨Tp.denote p, argTps, outTp, lambdaBody⟩ → + lambdas.lookup ref = some ⟨argTps, outTp, Tp.denote p, lambdaBody⟩ → Omni Γ ⟨vh, lambdas⟩ (lambdaBody args) Q → Omni Γ ⟨vh, lambdas⟩ (Expr.call h![] argTps outTp (.lambda ref) args) Q | lam {Q} : - (∀ ref, ref ∉ lambdas → Q (some (⟨vh, lambdas.insert ref ⟨_, argTps, outTp, lambdaBody⟩⟩, ref))) → + (∀ ref, ref ∉ lambdas → Q (some (⟨vh, lambdas.insert ref ⟨argTps, outTp, Tp.denote p, lambdaBody⟩⟩, ref))) → Omni Γ ⟨vh, lambdas⟩ (Expr.lambda argTps outTp lambdaBody) Q | callBuiltin {Q} : (b.omni p st argTypes resType args (mapToValHeapCondition st.lambdas Q)) → @@ -66,17 +66,19 @@ inductive Omni : Env → State p → Expr (Tp.denote p) tp → (Option (State p | callDecl: (fname, fn) ∈ Γ.functions → (hkc : fn.generics = tyKinds) → - (htci : fn.inTps (hkc ▸ generics) = argTypes) → - (htco : fn.outTp (hkc ▸ generics) = res) → - Omni Γ st (htco ▸ fn.body _ (hkc ▸ generics) (htci ▸ args)) Q → + (hrep : (fn.body (hkc ▸ generics) |>.rep) = Tp.denote p) → + (htci : (fn.body (hkc ▸ generics) |>.argTps) = argTypes) → + (htco : (fn.body (hkc ▸ generics) |>.outTp) = res) → + Omni Γ st (hrep ▸ htco ▸ (fn.body (hkc ▸ generics) |>.body (hrep ▸ htci ▸ args))) Q → Omni Γ st (@Expr.call _ tyKinds generics argTypes res (.decl fname) args) Q | callTrait {impl}: TraitResolution Γ traitRef impl → (fname, fn) ∈ impl → (hkc : fn.generics = tyKinds) → - (htci : fn.inTps (hkc ▸ generics) = argTypes) → - (htco : fn.outTp (hkc ▸ generics) = res) → - Omni Γ st (htco ▸ fn.body _ (hkc ▸ generics) (htci ▸ args)) Q → + (hrep : (fn.body (hkc ▸ generics) |>.rep) = Tp.denote p) → + (htci : (fn.body (hkc ▸ generics) |>.argTps) = argTypes) → + (htco : (fn.body (hkc ▸ generics) |>.outTp) = res) → + Omni Γ st (hrep ▸ htco ▸ (fn.body (hkc ▸ generics) |>.body (hrep ▸ htci ▸ args))) Q → Omni Γ st (@Expr.call _ tyKinds generics argTypes res (.trait ⟨traitRef, fname⟩) args) Q | loopDone : lo ≥ hi → diff --git a/Lampe/Lampe/Syntax.lean b/Lampe/Lampe/Syntax.lean index b7d8802..e717941 100644 --- a/Lampe/Lampe/Syntax.lean +++ b/Lampe/Lampe/Syntax.lean @@ -221,7 +221,7 @@ partial def mkExpr [MonadSyntax m] (e : TSyntax `nr_expr) (vname : Option Lean.I end -def mkFnDecl [Monad m] [MonadQuotation m] [MonadExceptOf Exception m] [MonadError m] : Syntax → m (String × TSyntax `term) +def mkFnDecl [Monad m] [MonadQuotation m] [MonadExceptOf Exception m] [MonadError m] (syn : Syntax) (rep : Lean.Ident) : m (String × TSyntax `term) := match syn with | `(nr_fn_decl| $name < $generics,* > ( $params,* ) -> $outTp { $bExprs;* }) => do let name ← mkNrIdent name let generics := generics.getElems.toList.map fun i => (mkIdent $ Name.mkSimple i.getId.toString) @@ -229,27 +229,24 @@ def mkFnDecl [Monad m] [MonadQuotation m] [MonadExceptOf Exception m] [MonadErro let params : List (TSyntax `term × TSyntax `term) ← params.getElems.toList.mapM fun p => match p with | `(nr_param_decl|$i:ident : $tp) => do pure (i, ←mkNrType tp) | _ => throwUnsupportedSyntax - let inputsDecl ← `(fun generics => match generics with | $(←mkHListLit generics) => $(←mkListLit $ params.map Prod.snd )) - let outType ← mkNrType outTp - let outDecl ← `(fun generics => match generics with | $(←mkHListLit generics) => $outType) let body ← MonadSyntax.run ((mkBlock bExprs.getElems.toList) (fun x => `(Expr.var $x))) - let bodyDecl ← `( - fun rep generics => match generics with - | $(←mkHListLit generics) => fun arguments => match arguments with - | $(←mkHListLit $ params.map Prod.fst) => $body - ) - let syn : TSyntax `term ← `(FunctionDecl.mk $(Syntax.mkStrLit name) $ Function.mk $genericsDecl $inputsDecl $outDecl $bodyDecl) + let lambdaDecl ← `(fun generics => match generics with + | $(←mkHListLit generics) => ⟨$(←mkListLit $ params.map Prod.snd), $(←mkNrType outTp), $rep, fun args => match args with + | $(←mkHListLit $ params.map Prod.fst) => $body⟩) + let syn : TSyntax `term ← `(FunctionDecl.mk $(Syntax.mkStrLit name) $ Function.mk $genericsDecl $lambdaDecl) pure (name, syn) | _ => throwUnsupportedSyntax elab "expr![" expr:nr_expr "]" : term => do let term ← MonadSyntax.run $ mkExpr expr none fun x => ``(Expr.var $x) Elab.Term.elabTerm term.raw none -elab "nrfn![" "fn" fn:nr_fn_decl "]" : term => do Elab.Term.elabTerm (←mkFnDecl fn).2 none +elab "nrfn![" rep:ident ";" "fn" fn:nr_fn_decl "]" : term => do + Elab.Term.elabTerm (←mkFnDecl fn rep).2 none elab "nr_def" decl:nr_fn_decl : command => do - let (name, decl) ← mkFnDecl decl - let decl ← `(def $(mkIdent $ Name.mkSimple name) : Lampe.FunctionDecl := $decl) + let rep := mkIdent $ Name.mkSimple "rep" + let (name, decl) ← mkFnDecl decl rep + let decl ← `(def $(mkIdent $ Name.mkSimple name) {$rep : Lampe.Tp → Type} : Lampe.FunctionDecl := $decl) Elab.Command.elabCommand decl end Lampe From 00f335f2ad6eac1c1b469231c580d54e0381bee5 Mon Sep 17 00:00:00 2001 From: utkn Date: Fri, 29 Nov 2024 19:51:42 +0300 Subject: [PATCH 35/39] nested_triple theorem moved to the tactic code --- Lampe/Lampe/Hoare/SepTotal.lean | 12 ------------ Lampe/Lampe/Tactic/SeparationLogic.lean | 14 +++++++++++++- 2 files changed, 13 insertions(+), 13 deletions(-) diff --git a/Lampe/Lampe/Hoare/SepTotal.lean b/Lampe/Lampe/Hoare/SepTotal.lean index ff685f5..131da71 100644 --- a/Lampe/Lampe/Hoare/SepTotal.lean +++ b/Lampe/Lampe/Hoare/SepTotal.lean @@ -91,18 +91,6 @@ theorem consequence_frame_left {H H₁ H₂ : SLP (State p)} rw [SLP.star_comm] apply SLP.ent_star_top -theorem nested_triple {Q : _ → SLP (State p)} - (h_hoare_imp : STHoare p Γ P e₁ Q → STHoare p Γ (P ⋆ H) e₂ (fun v => Q v ⋆ H)) - (h_hoare : STHoare p Γ P e₁ Q) - (h_ent_pre : H ⊢ P ⋆ H) : - STHoare p Γ H e₂ Q := by - have h_ent_post : ∀ v, ((Q v) ⋆ H) ⋆ ⊤ ⊢ (Q v) ⋆ ⊤ := by - simp [SLP.ent_drop_left] - have h_hoare' := h_hoare_imp h_hoare - apply consequence h_ent_pre (fun v => SLP.entails_self) - apply consequence SLP.entails_self h_ent_post - tauto - theorem var_intro {v : Tp.denote p tp}: STHoare p Γ ⟦⟧ (.var v) (fun v' => ⟦v' = v⟧) := by unfold STHoare diff --git a/Lampe/Lampe/Tactic/SeparationLogic.lean b/Lampe/Lampe/Tactic/SeparationLogic.lean index 0c74dc1..c19eb00 100644 --- a/Lampe/Lampe/Tactic/SeparationLogic.lean +++ b/Lampe/Lampe/Tactic/SeparationLogic.lean @@ -355,6 +355,18 @@ theorem lmbSingleton_star_congr {p} {r} {v₁ v₂ : Lambda} {R : SLP (State p)} rintro rfl apply SLP.entails_self +lemma nested_triple {Q : _ → SLP (State p)} + (h_hoare_imp : STHoare p Γ P e₁ Q → STHoare p Γ (P ⋆ H) e₂ (fun v => Q v ⋆ H)) + (h_hoare : STHoare p Γ P e₁ Q) + (h_ent_pre : H ⊢ P ⋆ H) : + STHoare p Γ H e₂ Q := by + have h_ent_post : ∀ v, ((Q v) ⋆ H) ⋆ ⊤ ⊢ (Q v) ⋆ ⊤ := by + simp [SLP.ent_drop_left] + have h_hoare' := h_hoare_imp h_hoare + apply consequence h_ent_pre (fun v => SLP.entails_self) + apply consequence SLP.entails_self h_ent_post + tauto + def canSolveSingleton (lhs : SLTerm) (rhsV : Expr): Bool := match lhs with @@ -516,7 +528,7 @@ macro "stephelper1" : tactic => `(tactic|( | apply fresh_intro | apply assert_intro | apply skip_intro - | apply STHoare.nested_triple STHoare.callLambda_intro + | apply nested_triple STHoare.callLambda_intro | apply lam_intro -- memory builtins | apply var_intro From 5e0c7127b624010095d5a7db69645a25644ee21d Mon Sep 17 00:00:00 2001 From: utkn Date: Fri, 29 Nov 2024 22:36:19 +0300 Subject: [PATCH 36/39] nrfn! now returns (rep -> Function) --- Lampe/Lampe.lean | 14 +++++++------- Lampe/Lampe/Ast.lean | 2 +- Lampe/Lampe/Hoare/SepTotal.lean | 3 ++- Lampe/Lampe/Semantics.lean | 9 +++++---- Lampe/Lampe/Syntax.lean | 6 ++++-- 5 files changed, 19 insertions(+), 15 deletions(-) diff --git a/Lampe/Lampe.lean b/Lampe/Lampe.lean index 4657fb8..663c086 100644 --- a/Lampe/Lampe.lean +++ b/Lampe/Lampe.lean @@ -113,29 +113,29 @@ example {p Γ} {x y : Tp.denote p Tp.field} : sl aesop -def bulbulizeField (rep : Tp → Type) : TraitImpl := { +def bulbulizeField : TraitImpl := { traitGenericKinds := [], implGenericKinds := [], traitGenerics := fun _ => h![], constraints := fun _ => [], self := fun _ => .field, - impl := fun _ => [("bulbulize", nrfn![rep; + impl := fun _ => [("bulbulize", nrfn![ fn bulbulize<>(x : Field) -> Field { #add(x, x) : Field - }].2 + }] )] } -def bulbulizeU32 (rep : Tp → Type) : TraitImpl := { +def bulbulizeU32 : TraitImpl := { traitGenericKinds := [], implGenericKinds := [], traitGenerics := fun _ => h![], constraints := fun _ => [], self := fun _ => .u 32, - impl := fun _ => [("bulbulize", nrfn![rep; + impl := fun _ => [("bulbulize", nrfn![ fn bulbulize<>(_x : u32) -> u32 { 69 : u32 - }].2 + }] )] } @@ -144,7 +144,7 @@ def simpleTraitCall (tp : Tp) (arg : tp.denote P): Expr (Tp.denote P) tp := def simpleTraitEnv (p : Prime) : Env := { functions := [], - traits := [("Bulbulize", bulbulizeField (Tp.denote p)), ("Bulbulize", bulbulizeU32 (Tp.denote p))] + traits := [("Bulbulize", bulbulizeField), ("Bulbulize", bulbulizeU32)] } example : STHoare p (simpleTraitEnv p) ⟦⟧ (simpleTraitCall .field arg) (fun v => v = 2 * arg) := by diff --git a/Lampe/Lampe/Ast.lean b/Lampe/Lampe/Ast.lean index 460c2f2..52e652e 100644 --- a/Lampe/Lampe/Ast.lean +++ b/Lampe/Lampe/Ast.lean @@ -79,7 +79,7 @@ implGenericKinds : List Kind traitGenerics : HList Kind.denote implGenericKinds → HList Kind.denote traitGenericKinds constraints : HList Kind.denote implGenericKinds → List TraitImplRef self : HList Kind.denote implGenericKinds → Tp -impl : HList Kind.denote implGenericKinds → List (Ident × Function) +impl : HList Kind.denote implGenericKinds → List (Ident × ((Tp → Type) → Function)) -- @[reducible] -- def Struct.tp (s: Struct): HList Kind.denote s.tyArgKinds → Tp := diff --git a/Lampe/Lampe/Hoare/SepTotal.lean b/Lampe/Lampe/Hoare/SepTotal.lean index 131da71..89f28fe 100644 --- a/Lampe/Lampe/Hoare/SepTotal.lean +++ b/Lampe/Lampe/Hoare/SepTotal.lean @@ -368,7 +368,8 @@ theorem lam_intro : theorem callTrait_intro {impl} {fname fn} (h_trait : TraitResolution Γ traitRef impl) - (h_fn : (fname, fn) ∈ impl) + (h_fn : (fname, fn') ∈ impl) + (_ : fn = fn' (Tp.denote p)) (hkc : fn.generics = tyKinds) (hrep : (fn.body (hkc ▸ generics) |>.rep) = Tp.denote p) (htci : (fn.body (hkc ▸ generics) |>.argTps) = argTypes) diff --git a/Lampe/Lampe/Semantics.lean b/Lampe/Lampe/Semantics.lean index eecc916..baec610 100644 --- a/Lampe/Lampe/Semantics.lean +++ b/Lampe/Lampe/Semantics.lean @@ -24,7 +24,7 @@ inductive TraitResolvable (Γ : Env): TraitImplRef → Prop where (∀constraint ∈ impl.constraints implGenerics, TraitResolvable Γ constraint) → TraitResolvable Γ ref -inductive TraitResolution (Γ : Env): TraitImplRef → List (Ident × Function) → Prop where +inductive TraitResolution (Γ : Env): TraitImplRef → List (Ident × ((Tp → Type) → Function)) → Prop where | ok {ref impl}: (ref.trait.name, impl) ∈ Γ.traits → (ktc : ref.trait.traitGenericKinds = impl.traitGenericKinds) → @@ -73,13 +73,14 @@ inductive Omni : Env → State p → Expr (Tp.denote p) tp → (Option (State p Omni Γ st (@Expr.call _ tyKinds generics argTypes res (.decl fname) args) Q | callTrait {impl}: TraitResolution Γ traitRef impl → - (fname, fn) ∈ impl → + (fname, fn') ∈ impl → + fn = fn' _ → (hkc : fn.generics = tyKinds) → (hrep : (fn.body (hkc ▸ generics) |>.rep) = Tp.denote p) → (htci : (fn.body (hkc ▸ generics) |>.argTps) = argTypes) → (htco : (fn.body (hkc ▸ generics) |>.outTp) = res) → Omni Γ st (hrep ▸ htco ▸ (fn.body (hkc ▸ generics) |>.body (hrep ▸ htci ▸ args))) Q → - Omni Γ st (@Expr.call _ tyKinds generics argTypes res (.trait ⟨traitRef, fname⟩) args) Q + Omni Γ st (@Expr.call (Tp.denote p) tyKinds generics argTypes res (.trait ⟨traitRef, fname⟩) args) Q | loopDone : lo ≥ hi → Omni Γ st (.loop lo hi body) Q @@ -190,7 +191,7 @@ theorem Omni.frame {p Γ tp} {st₁ st₂ : State p} {e : Expr (Tp.denote p) tp} constructor all_goals (try assumption) tauto - | callTrait _ _ _ _ _ _ ih => + | callTrait _ _ _ _ _ _ _ => intro constructor all_goals (try assumption) diff --git a/Lampe/Lampe/Syntax.lean b/Lampe/Lampe/Syntax.lean index e717941..e8acedd 100644 --- a/Lampe/Lampe/Syntax.lean +++ b/Lampe/Lampe/Syntax.lean @@ -240,8 +240,10 @@ def mkFnDecl [Monad m] [MonadQuotation m] [MonadExceptOf Exception m] [MonadErro elab "expr![" expr:nr_expr "]" : term => do let term ← MonadSyntax.run $ mkExpr expr none fun x => ``(Expr.var $x) Elab.Term.elabTerm term.raw none -elab "nrfn![" rep:ident ";" "fn" fn:nr_fn_decl "]" : term => do - Elab.Term.elabTerm (←mkFnDecl fn rep).2 none +elab "nrfn![" "fn" fn:nr_fn_decl "]" : term => do + let rep := mkIdent $ Name.mkSimple "rep" + let stx ← `(fun $rep => $((←mkFnDecl fn rep).2).fn) + Elab.Term.elabTerm stx none elab "nr_def" decl:nr_fn_decl : command => do let rep := mkIdent $ Name.mkSimple "rep" From d3a844106db058cdf669306eeef5681105d65d04 Mon Sep 17 00:00:00 2001 From: utkn Date: Sat, 30 Nov 2024 11:10:55 +0300 Subject: [PATCH 37/39] small changes --- Lampe/Lampe.lean | 6 +++--- Lampe/Lampe/Hoare/SepTotal.lean | 4 ++-- 2 files changed, 5 insertions(+), 5 deletions(-) diff --git a/Lampe/Lampe.lean b/Lampe/Lampe.lean index 663c086..6e2be9c 100644 --- a/Lampe/Lampe.lean +++ b/Lampe/Lampe.lean @@ -142,12 +142,12 @@ def bulbulizeU32 : TraitImpl := { def simpleTraitCall (tp : Tp) (arg : tp.denote P): Expr (Tp.denote P) tp := @Expr.call _ [] h![] [tp] tp (.trait ⟨⟨⟨"Bulbulize", [], h![]⟩, tp⟩, "bulbulize"⟩) h![arg] -def simpleTraitEnv (p : Prime) : Env := { +def simpleTraitEnv : Env := { functions := [], traits := [("Bulbulize", bulbulizeField), ("Bulbulize", bulbulizeU32)] } -example : STHoare p (simpleTraitEnv p) ⟦⟧ (simpleTraitCall .field arg) (fun v => v = 2 * arg) := by +example : STHoare p simpleTraitEnv ⟦⟧ (simpleTraitCall .field arg) (fun v => v = 2 * arg) := by simp only [simpleTraitCall] apply STHoare.callTrait_intro · constructor @@ -164,7 +164,7 @@ example : STHoare p (simpleTraitEnv p) ⟦⟧ (simpleTraitCall .field arg) (fun subst_vars ring -example : STHoare p (simpleTraitEnv p) ⟦⟧ (simpleTraitCall (.u 32) arg) (fun v => v = 69) := by +example : STHoare p simpleTraitEnv ⟦⟧ (simpleTraitCall (.u 32) arg) (fun v => v = 69) := by simp only [simpleTraitCall] apply STHoare.callTrait_intro · constructor diff --git a/Lampe/Lampe/Hoare/SepTotal.lean b/Lampe/Lampe/Hoare/SepTotal.lean index 89f28fe..59e6e3f 100644 --- a/Lampe/Lampe/Hoare/SepTotal.lean +++ b/Lampe/Lampe/Hoare/SepTotal.lean @@ -368,8 +368,8 @@ theorem lam_intro : theorem callTrait_intro {impl} {fname fn} (h_trait : TraitResolution Γ traitRef impl) - (h_fn : (fname, fn') ∈ impl) - (_ : fn = fn' (Tp.denote p)) + (h_fn' : (fname, fn') ∈ impl) + (h_fn : fn = fn' (Tp.denote p)) (hkc : fn.generics = tyKinds) (hrep : (fn.body (hkc ▸ generics) |>.rep) = Tp.denote p) (htci : (fn.body (hkc ▸ generics) |>.argTps) = argTypes) From 4feede0fb017426fd09337e4245e1f7b4e0e95f2 Mon Sep 17 00:00:00 2001 From: utkn Date: Mon, 2 Dec 2024 15:24:51 +0300 Subject: [PATCH 38/39] Lambda is now parametrized over "rep" instead of having "rep" as its field. Accordingly, ugly hacks to acommodate this were fixed --- Lampe/Lampe.lean | 12 +++++----- Lampe/Lampe/Ast.lean | 11 +++++----- Lampe/Lampe/Hoare/SepTotal.lean | 20 ++++++++--------- Lampe/Lampe/Semantics.lean | 29 +++++++++++-------------- Lampe/Lampe/SeparationLogic/State.lean | 8 +++---- Lampe/Lampe/Syntax.lean | 14 +++++------- Lampe/Lampe/Tactic/SeparationLogic.lean | 4 ++-- 7 files changed, 45 insertions(+), 53 deletions(-) diff --git a/Lampe/Lampe.lean b/Lampe/Lampe.lean index 6e2be9c..919e701 100644 --- a/Lampe/Lampe.lean +++ b/Lampe/Lampe.lean @@ -9,7 +9,7 @@ nr_def simple_muts<>(x : Field) -> Field { y } -example : STHoare p Γ ⟦⟧ (simple_muts.fn.body h![] |>.body h![x]) fun v => v = x := by +example : STHoare p Γ ⟦⟧ (simple_muts.fn.body _ h![] |>.body h![x]) fun v => v = x := by simp only [simple_muts] steps try (exact _) @@ -22,7 +22,7 @@ nr_def weirdEq(x : I, y : I) -> Unit { #assert(#eq(a, y) : bool) : Unit; } -example {x y : Tp.denote p .field} : STHoare p Γ ⟦⟧ (weirdEq.fn.body h![.field] |>.body h![x, y]) fun _ => x = y := by +example {x y : Tp.denote p .field} : STHoare p Γ ⟦⟧ (weirdEq.fn.body _ h![.field] |>.body h![x, y]) fun _ => x = y := by simp only [weirdEq] steps simp_all @@ -40,7 +40,7 @@ lemma BitVec.add_toNat_of_lt_max {a b : BitVec w} (h: a.toNat + b.toNat < 2^w) : rw [Nat.mod_eq_of_lt] assumption -example {self that : Tp.denote p (.slice tp)} : STHoare p Γ ⟦⟧ (sliceAppend.fn.body h![tp] |>.body h![self, that]) fun v => v = self ++ that := by +example {self that : Tp.denote p (.slice tp)} : STHoare p Γ ⟦⟧ (sliceAppend.fn.body _ h![tp] |>.body h![self, that]) fun v => v = self ++ that := by simp only [sliceAppend] steps rename Tp.denote _ tp.slice.ref => selfRef @@ -67,7 +67,7 @@ nr_def simple_if<>(x : Field, y : Field) -> Field { z } -example {p Γ x y}: STHoare p Γ ⟦⟧ (simple_if.fn.body h![] |>.body h![x, y]) +example {p Γ x y}: STHoare p Γ ⟦⟧ (simple_if.fn.body _ h![] |>.body h![x, y]) fun v => v = y := by simp only [simple_if] steps <;> tauto @@ -83,7 +83,7 @@ nr_def simple_if_else<>(x : Field, y : Field) -> Field { z } -example {p Γ x y}: STHoare p Γ ⟦⟧ (simple_if_else.fn.body h![] |>.body h![x, y]) +example {p Γ x y}: STHoare p Γ ⟦⟧ (simple_if_else.fn.body _ h![] |>.body h![x, y]) fun v => v = x := by simp only [simple_if_else] steps @@ -98,7 +98,7 @@ nr_def simple_lambda<>(x : Field, y : Field) -> Field { } example {p Γ} {x y : Tp.denote p Tp.field} : - STHoare p Γ ⟦⟧ (simple_lambda.fn.body h![] |>.body h![x, y]) + STHoare p Γ ⟦⟧ (simple_lambda.fn.body _ h![] |>.body h![x, y]) fun v => v = (x + y) := by simp only [simple_lambda] steps diff --git a/Lampe/Lampe/Ast.lean b/Lampe/Lampe/Ast.lean index 52e652e..6c42747 100644 --- a/Lampe/Lampe/Ast.lean +++ b/Lampe/Lampe/Ast.lean @@ -45,20 +45,19 @@ inductive Expr (rep : Tp → Type) : Tp → Type where | loop : rep (.u s) → rep (.u s) → (rep (.u s) → Expr rep r) → Expr rep .unit | lambda : (argTps : List Tp) → (outTp : Tp) → (HList rep argTps → Expr rep outTp) → Expr rep .lambdaRef -structure Lambda where +structure Lambda (rep : Tp → Type) where argTps : List Tp outTp : Tp - rep : Tp → Type body : HList rep argTps → Expr rep outTp structure Function : Type _ where generics : List Kind - body : (HList Kind.denote generics) → Lambda + body : ∀ (rep : Tp → Type), (HList Kind.denote generics) → Lambda rep /-- Polymorphic identity -/ -example {rep : Tp → Type}: Function := { +example : Function := { generics := [.type] - body := fun h![tp] => ⟨[tp], tp, rep, fun h![x] => .var x⟩ + body := fun _ h![tp] => ⟨[tp], tp, fun h![x] => .var x⟩ } structure FunctionDecl where @@ -79,7 +78,7 @@ implGenericKinds : List Kind traitGenerics : HList Kind.denote implGenericKinds → HList Kind.denote traitGenericKinds constraints : HList Kind.denote implGenericKinds → List TraitImplRef self : HList Kind.denote implGenericKinds → Tp -impl : HList Kind.denote implGenericKinds → List (Ident × ((Tp → Type) → Function)) +impl : HList Kind.denote implGenericKinds → List (Ident × Function) -- @[reducible] -- def Struct.tp (s: Struct): HList Kind.denote s.tyArgKinds → Tp := diff --git a/Lampe/Lampe/Hoare/SepTotal.lean b/Lampe/Lampe/Hoare/SepTotal.lean index 59e6e3f..bb9dfd9 100644 --- a/Lampe/Lampe/Hoare/SepTotal.lean +++ b/Lampe/Lampe/Hoare/SepTotal.lean @@ -308,9 +308,9 @@ theorem skip_intro : theorem callLambda_intro {lambdaBody} {P : SLP (State p)} {Q : Tp.denote p outTp → SLP (State p)}: @STHoare outTp p Γ P (lambdaBody args) Q → - STHoare p Γ (P ⋆ [λref ↦ ⟨argTps, outTp, _, lambdaBody⟩]) + STHoare p Γ (P ⋆ [λref ↦ ⟨argTps, outTp, lambdaBody⟩]) (Expr.call h![] argTps outTp (.lambda ref) args) - (fun v => (Q v) ⋆ [λref ↦ ⟨argTps, outTp, _, lambdaBody⟩]) := by + (fun v => (Q v) ⋆ [λref ↦ ⟨argTps, outTp, lambdaBody⟩]) := by intros rename_i h unfold STHoare THoare @@ -322,7 +322,7 @@ theorem callLambda_intro {lambdaBody} {P : SLP (State p)} {Q : Tp.denote p outTp obtain ⟨st₁', ⟨st₂', _⟩⟩ := h₃ simp_all only [State.union_parts, Finmap.mem_union, Finmap.mem_singleton, or_true, Finmap.lookup_union_left] - generalize hL : (⟨_, _, _, _⟩ : Lambda) = lmb at * + generalize hL : (⟨_, _, _⟩ : Lambda _) = lmb at * have _ : ref ∉ st₁'.lambdas := by rename_i h₃ obtain ⟨hi₁, _, _, hi₂⟩ := h₃ @@ -341,14 +341,14 @@ theorem callLambda_intro {lambdaBody} {P : SLP (State p)} {Q : Tp.denote p outTp theorem lam_intro : STHoare p Γ ⟦⟧ (.lambda argTps outTp lambdaBody) - fun v => [λv ↦ ⟨argTps, outTp, _, lambdaBody⟩] := by + fun v => [λv ↦ ⟨argTps, outTp, lambdaBody⟩] := by unfold STHoare THoare intros H st h constructor intros simp_all only [SLP.true_star, SLP.star_assoc] rename Ref => r - generalize (⟨_, _, _, _⟩ : Lambda) = lambda + generalize (⟨_, _, _⟩ : Lambda _) = lambda exists ⟨∅, Finmap.singleton r lambda⟩, st refine ⟨?_, ?_, ?_, ?_⟩ . simp only [LawfulHeap.disjoint] @@ -368,13 +368,11 @@ theorem lam_intro : theorem callTrait_intro {impl} {fname fn} (h_trait : TraitResolution Γ traitRef impl) - (h_fn' : (fname, fn') ∈ impl) - (h_fn : fn = fn' (Tp.denote p)) + (h_fn : (fname, fn) ∈ impl) (hkc : fn.generics = tyKinds) - (hrep : (fn.body (hkc ▸ generics) |>.rep) = Tp.denote p) - (htci : (fn.body (hkc ▸ generics) |>.argTps) = argTypes) - (htco : (fn.body (hkc ▸ generics) |>.outTp) = res) - (h_hoare: STHoare p Γ H (hrep ▸ htco ▸ (fn.body (hkc ▸ generics) |>.body (hrep ▸ htci ▸ args))) Q): + (htci : (fn.body _ (hkc ▸ generics) |>.argTps) = argTypes) + (htco : (fn.body _ (hkc ▸ generics) |>.outTp) = res) + (h_hoare: STHoare p Γ H (htco ▸ (fn.body _ (hkc ▸ generics) |>.body (htci ▸ args))) Q): STHoare p Γ H (@Expr.call _ tyKinds generics argTypes res (.trait ⟨traitRef, fname⟩) args) Q := by diff --git a/Lampe/Lampe/Semantics.lean b/Lampe/Lampe/Semantics.lean index baec610..3c7dd1a 100644 --- a/Lampe/Lampe/Semantics.lean +++ b/Lampe/Lampe/Semantics.lean @@ -24,7 +24,7 @@ inductive TraitResolvable (Γ : Env): TraitImplRef → Prop where (∀constraint ∈ impl.constraints implGenerics, TraitResolvable Γ constraint) → TraitResolvable Γ ref -inductive TraitResolution (Γ : Env): TraitImplRef → List (Ident × ((Tp → Type) → Function)) → Prop where +inductive TraitResolution (Γ : Env): TraitImplRef → List (Ident × Function) → Prop where | ok {ref impl}: (ref.trait.name, impl) ∈ Γ.traits → (ktc : ref.trait.traitGenericKinds = impl.traitGenericKinds) → @@ -53,12 +53,12 @@ inductive Omni : Env → State p → Expr (Tp.denote p) tp → (Option (State p (∀v st, Q₁ (some (st, v)) → Omni Γ st (b v) Q) → (Q₁ none → Q none) → Omni Γ st (.letIn e b) Q -| callLambda {lambdas : Lambdas} : - lambdas.lookup ref = some ⟨argTps, outTp, Tp.denote p, lambdaBody⟩ → +| callLambda {lambdas : Lambdas p} : + lambdas.lookup ref = some ⟨argTps, outTp, lambdaBody⟩ → Omni Γ ⟨vh, lambdas⟩ (lambdaBody args) Q → Omni Γ ⟨vh, lambdas⟩ (Expr.call h![] argTps outTp (.lambda ref) args) Q | lam {Q} : - (∀ ref, ref ∉ lambdas → Q (some (⟨vh, lambdas.insert ref ⟨argTps, outTp, Tp.denote p, lambdaBody⟩⟩, ref))) → + (∀ ref, ref ∉ lambdas → Q (some (⟨vh, lambdas.insert ref ⟨argTps, outTp, lambdaBody⟩⟩, ref))) → Omni Γ ⟨vh, lambdas⟩ (Expr.lambda argTps outTp lambdaBody) Q | callBuiltin {Q} : (b.omni p st argTypes resType args (mapToValHeapCondition st.lambdas Q)) → @@ -66,20 +66,17 @@ inductive Omni : Env → State p → Expr (Tp.denote p) tp → (Option (State p | callDecl: (fname, fn) ∈ Γ.functions → (hkc : fn.generics = tyKinds) → - (hrep : (fn.body (hkc ▸ generics) |>.rep) = Tp.denote p) → - (htci : (fn.body (hkc ▸ generics) |>.argTps) = argTypes) → - (htco : (fn.body (hkc ▸ generics) |>.outTp) = res) → - Omni Γ st (hrep ▸ htco ▸ (fn.body (hkc ▸ generics) |>.body (hrep ▸ htci ▸ args))) Q → + (htci : (fn.body _ (hkc ▸ generics) |>.argTps) = argTypes) → + (htco : (fn.body _ (hkc ▸ generics) |>.outTp) = res) → + Omni Γ st (htco ▸ (fn.body _ (hkc ▸ generics) |>.body (htci ▸ args))) Q → Omni Γ st (@Expr.call _ tyKinds generics argTypes res (.decl fname) args) Q | callTrait {impl}: TraitResolution Γ traitRef impl → - (fname, fn') ∈ impl → - fn = fn' _ → + (fname, fn) ∈ impl → (hkc : fn.generics = tyKinds) → - (hrep : (fn.body (hkc ▸ generics) |>.rep) = Tp.denote p) → - (htci : (fn.body (hkc ▸ generics) |>.argTps) = argTypes) → - (htco : (fn.body (hkc ▸ generics) |>.outTp) = res) → - Omni Γ st (hrep ▸ htco ▸ (fn.body (hkc ▸ generics) |>.body (hrep ▸ htci ▸ args))) Q → + (htci : (fn.body _ (hkc ▸ generics) |>.argTps) = argTypes) → + (htco : (fn.body _ (hkc ▸ generics) |>.outTp) = res) → + Omni Γ st (htco ▸ (fn.body _ (hkc ▸ generics) |>.body (htci ▸ args))) Q → Omni Γ st (@Expr.call (Tp.denote p) tyKinds generics argTypes res (.trait ⟨traitRef, fname⟩) args) Q | loopDone : lo ≥ hi → @@ -224,10 +221,10 @@ theorem Omni.frame {p Γ tp} {st₁ st₂ : State p} {e : Expr (Tp.denote p) tp} constructor intros simp only - rename Lambdas => lmbs + rename Lambdas _ => lmbs rename ValHeap _ => vh rename Ref => r - generalize hL : (⟨_, _, _, _⟩ : Lambda) = lambda + generalize hL : (⟨_, _, _⟩ : Lambda _) = lambda have hi : r ∉ lmbs ∧ r ∉ st₂.lambdas := by aesop have hd₁ : Finmap.Disjoint lmbs (Finmap.singleton r lambda) := by apply Finmap.Disjoint.symm diff --git a/Lampe/Lampe/SeparationLogic/State.lean b/Lampe/Lampe/SeparationLogic/State.lean index f318e3d..11295a4 100644 --- a/Lampe/Lampe/SeparationLogic/State.lean +++ b/Lampe/Lampe/SeparationLogic/State.lean @@ -4,11 +4,11 @@ import Lampe.Ast namespace Lampe -abbrev Lambdas := Finmap fun _ : Ref => Lambda +abbrev Lambdas (p : Prime) := Finmap fun _ : Ref => Lambda (Tp.denote p) structure State (p : Prime) where vals : ValHeap p - lambdas : Lambdas + lambdas : Lambdas p instance : Membership Ref (State p) where mem := fun a e => e ∈ a.vals @@ -22,7 +22,7 @@ instance : Coe (State p) (ValHeap p) where /-- Maps a post-condition on `State`s to a post-condition on `ValHeap`s by keeping the lambdas fixed -/ @[reducible] def mapToValHeapCondition - (lambdas : Lambdas) + (lambdas : Lambdas p) (Q : Option (State p × T) → Prop) : Option (ValHeap p × T) → Prop := fun vv => Q (vv.map (fun (vals, t) => ⟨⟨vals, lambdas⟩, t⟩)) @@ -85,7 +85,7 @@ def State.valSingleton (r : Ref) (v : AnyValue p) : SLP (State p) := notation:max "[" l " ↦ " r "]" => State.valSingleton l r @[reducible] -def State.lmbSingleton (r : Ref) (v : Lambda) : SLP (State p) := +def State.lmbSingleton (r : Ref) (v : Lambda (Tp.denote p)) : SLP (State p) := fun st => st.lambdas = Finmap.singleton r v notation:max "[" "λ" l " ↦ " r "]" => State.lmbSingleton l r diff --git a/Lampe/Lampe/Syntax.lean b/Lampe/Lampe/Syntax.lean index e8acedd..1911765 100644 --- a/Lampe/Lampe/Syntax.lean +++ b/Lampe/Lampe/Syntax.lean @@ -221,7 +221,7 @@ partial def mkExpr [MonadSyntax m] (e : TSyntax `nr_expr) (vname : Option Lean.I end -def mkFnDecl [Monad m] [MonadQuotation m] [MonadExceptOf Exception m] [MonadError m] (syn : Syntax) (rep : Lean.Ident) : m (String × TSyntax `term) := match syn with +def mkFnDecl [Monad m] [MonadQuotation m] [MonadExceptOf Exception m] [MonadError m] (syn : Syntax) : m (String × TSyntax `term) := match syn with | `(nr_fn_decl| $name < $generics,* > ( $params,* ) -> $outTp { $bExprs;* }) => do let name ← mkNrIdent name let generics := generics.getElems.toList.map fun i => (mkIdent $ Name.mkSimple i.getId.toString) @@ -230,8 +230,8 @@ def mkFnDecl [Monad m] [MonadQuotation m] [MonadExceptOf Exception m] [MonadErro | `(nr_param_decl|$i:ident : $tp) => do pure (i, ←mkNrType tp) | _ => throwUnsupportedSyntax let body ← MonadSyntax.run ((mkBlock bExprs.getElems.toList) (fun x => `(Expr.var $x))) - let lambdaDecl ← `(fun generics => match generics with - | $(←mkHListLit generics) => ⟨$(←mkListLit $ params.map Prod.snd), $(←mkNrType outTp), $rep, fun args => match args with + let lambdaDecl ← `(fun rep generics => match generics with + | $(←mkHListLit generics) => ⟨$(←mkListLit $ params.map Prod.snd), $(←mkNrType outTp), fun args => match args with | $(←mkHListLit $ params.map Prod.fst) => $body⟩) let syn : TSyntax `term ← `(FunctionDecl.mk $(Syntax.mkStrLit name) $ Function.mk $genericsDecl $lambdaDecl) pure (name, syn) @@ -241,14 +241,12 @@ elab "expr![" expr:nr_expr "]" : term => do let term ← MonadSyntax.run $ mkExpr expr none fun x => ``(Expr.var $x) Elab.Term.elabTerm term.raw none elab "nrfn![" "fn" fn:nr_fn_decl "]" : term => do - let rep := mkIdent $ Name.mkSimple "rep" - let stx ← `(fun $rep => $((←mkFnDecl fn rep).2).fn) + let stx ← `($((←mkFnDecl fn).2).fn) Elab.Term.elabTerm stx none elab "nr_def" decl:nr_fn_decl : command => do - let rep := mkIdent $ Name.mkSimple "rep" - let (name, decl) ← mkFnDecl decl rep - let decl ← `(def $(mkIdent $ Name.mkSimple name) {$rep : Lampe.Tp → Type} : Lampe.FunctionDecl := $decl) + let (name, decl) ← mkFnDecl decl + let decl ← `(def $(mkIdent $ Name.mkSimple name) : Lampe.FunctionDecl := $decl) Elab.Command.elabCommand decl end Lampe diff --git a/Lampe/Lampe/Tactic/SeparationLogic.lean b/Lampe/Lampe/Tactic/SeparationLogic.lean index c19eb00..f2dbcfd 100644 --- a/Lampe/Lampe/Tactic/SeparationLogic.lean +++ b/Lampe/Lampe/Tactic/SeparationLogic.lean @@ -340,7 +340,7 @@ theorem singleton_congr_mv {p} {r} {v₁ v₂ : AnyValue p} : (v₁ = v₂) → simp apply SLP.entails_self -theorem lmbSingleton_congr_mv {p} {r} {l₁ l₂ : Lambda} : (l₁ = l₂) → (([λr ↦ l₁] : SLP (State p)) ⊢ [λr ↦ l₂] ⋆ ⟦⟧) := by +theorem lmbSingleton_congr_mv {p} {r} {l₁ l₂ : Lambda _} : (l₁ = l₂) → (([λr ↦ l₁] : SLP (State p)) ⊢ [λr ↦ l₂] ⋆ ⟦⟧) := by rintro rfl simp apply SLP.entails_self @@ -350,7 +350,7 @@ theorem singleton_star_congr {p} {r} {v₁ v₂ : AnyValue p} {R} : (v₁ = v₂ rintro rfl apply SLP.entails_self -theorem lmbSingleton_star_congr {p} {r} {v₁ v₂ : Lambda} {R : SLP (State p)} : +theorem lmbSingleton_star_congr {p} {r} {v₁ v₂ : Lambda _} {R : SLP (State p)} : (v₁ = v₂) → ([λr ↦ v₁] ⋆ R ⊢ [λr ↦ v₂] ⋆ R) := by rintro rfl apply SLP.entails_self From 6c7d4f57e513f170c38eb67507ea7b2859c35593 Mon Sep 17 00:00:00 2001 From: utkn Date: Mon, 2 Dec 2024 15:33:21 +0300 Subject: [PATCH 39/39] small fix in the proof in the main --- Lampe/Lampe.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/Lampe/Lampe.lean b/Lampe/Lampe.lean index 919e701..dce6382 100644 --- a/Lampe/Lampe.lean +++ b/Lampe/Lampe.lean @@ -12,7 +12,6 @@ nr_def simple_muts<>(x : Field) -> Field { example : STHoare p Γ ⟦⟧ (simple_muts.fn.body _ h![] |>.body h![x]) fun v => v = x := by simp only [simple_muts] steps - try (exact _) simp_all nr_def weirdEq(x : I, y : I) -> Unit {