Skip to content

Commit

Permalink
strengthen the claim of likeLemmaFour
Browse files Browse the repository at this point in the history
  • Loading branch information
m4lvin committed Oct 5, 2023
1 parent f7f318c commit 00d27e6
Showing 1 changed file with 12 additions and 7 deletions.
19 changes: 12 additions & 7 deletions Pdl/Unravel.lean
Original file line number Diff line number Diff line change
Expand Up @@ -51,9 +51,10 @@ def nsub : Formula → List Formula
theorem likeLemmaFour :
∀ (W M) (a : Program) (w v : W) (X' X : List Formula) (P : Formula),
X = X' ++ {~⌈a⌉ P} →
(M, w)⊨Con X → relate M a w v → (M, v)⊨(~P) → ∃ Y ∈ {X'}⊎unravel (~⌈a⌉ P), (M, w)⊨Con Y :=
-- TODO: ∧ ∃ a_1 ... a_n: ~⌈a_1⌉...⌈a_n⌉P ∈ Y ∧ relate M (a_1;…;a_n) w v
-- first needs a "list of boxes in front of formula def"
(M, w)⊨Con X → relate M a w v → (M, v)⊨(~P) →
∃ Y ∈ {X'}⊎unravel (~⌈a⌉ P), (M, w)⊨Con Y
∧ ∃ as : List Program, (~ Formula.boxes as P) ∈ Y
∧ relate M (Program.steps as) w v :=
by
intro W M a
-- 'induction' tactic does not support mutually inductive types, ...
Expand All @@ -68,7 +69,9 @@ theorem likeLemmaFour :
constructor
· subst X_def
rfl
assumption
constructor
· assumption
· sorry
case sequence -- a -- b c -- IHb
b c =>
intro w v X' X P X_def w_sat_X w_bc_v v_sat_nP
Expand Down Expand Up @@ -101,13 +104,15 @@ theorem likeLemmaFour :
use y
tauto
· tauto
· sorry
-- need (M, u)⊨~⌈c⌉ P
· sorry -- need (M, u)⊨~⌈c⌉ P -- make and use IHc here ?
rcases IHb with ⟨Y, Y_in, w_conY⟩
use Y
constructor
sorry
exact w_conY
constructor
· tauto
· use [c]
sorry
case union =>
intro w v X' X P X_def w_sat_X w_a_v v_sat_nP; subst X_def
sorry
Expand Down

0 comments on commit 00d27e6

Please sign in to comment.