Skip to content

Commit 1da1ab7

Browse files
committed
Some clean up and preperation for the last sorry
1 parent 678b3b2 commit 1da1ab7

File tree

3 files changed

+32
-121
lines changed

3 files changed

+32
-121
lines changed

InferenceInLean/Basic.lean

Lines changed: 9 additions & 55 deletions
Original file line numberDiff line numberDiff line change
@@ -92,19 +92,6 @@ theorem groundResolution_soundness {A : Atom sig Empty} {C D : Clause sig Empty}
9292

9393
/- ### 3.10 General Resolution -/
9494

95-
-- TODO: do we need to add that freeVars ∩ freeVars = ∅?
96-
@[simp]
97-
def GeneralResolutionRule [inst : DecidableEq X] (A B : Atom sig X) (C D : Clause sig X)
98-
(σ : Substitution sig X) :
99-
Inference sig X :=
100-
⟨{.neg A :: C, .pos B :: D}, (C ++ D).substitute σ, MostGeneralUnifier [(A, B)] σ⟩
101-
102-
@[simp]
103-
def GeneralFactorizationRule [inst : DecidableEq X] (A B : Atom sig X) (C : Clause sig X)
104-
(σ : Substitution sig X) :
105-
Inference sig X :=
106-
⟨{.pos A :: .pos B :: C}, Clause.substitute σ (.pos A :: C), MostGeneralUnifier [(A, B)] σ⟩
107-
10895
lemma validclosed_of_valid [DecidableEq X] {C: Clause sig X} {I : Interpretation sig univ} :
10996
ValidIn C.toFormula I → ValidIn C.toClosedFormula I := by
11097
intro heval
@@ -114,6 +101,7 @@ lemma validclosed_of_valid [DecidableEq X] {C: Clause sig X} {I : Interpretation
114101
have := (@three_three_seven sig X univ _ n C I xs hxsnodup rfl).mpr
115102
exact fun β => this heval β
116103

104+
/- Direct formalization of Proposition 3.10.14 -/
117105
theorem generalResolutionRuleSound [DecidableEq X] (A B : Atom sig X) (C D : Clause sig X)
118106
(σ : Substitution sig X) (hmgu : MostGeneralUnifier [(A, B)] σ) :
119107
@Entails _ _ univ _
@@ -144,7 +132,8 @@ theorem generalResolutionRuleSound [DecidableEq X] (A B : Atom sig X) (C D : Cla
144132

145133
let leftys : List X := (leftinner.substitute σ).freeVarsList
146134
let leftm : ℕ := leftys.length
147-
have hleftysnodup : leftys.Nodup := by exact nodup_clauseFreeVarsList sig X (Clause.substitute σ leftinner)
135+
have hleftysnodup : leftys.Nodup := by
136+
exact nodup_clauseFreeVarsList sig X (Clause.substitute σ leftinner)
148137

149138
have hleft338 := @three_three_eight univ sig X _
150139
leftinner I σ leftn leftm leftxs leftys hleftxsnodup rfl hleftysnodup rfl hleftvalid
@@ -164,14 +153,14 @@ theorem generalResolutionRuleSound [DecidableEq X] (A B : Atom sig X) (C D : Cla
164153

165154
-- use 3.3.7
166155
have hleftinnersub : @ValidIn _ X _ _ (leftinner.substitute σ) I := by
167-
exact (three_three_seven leftys.length (Clause.toFormula sig X (Clause.substitute σ leftinner)) I
168-
leftys hleftysnodup rfl).mp
169-
hleft338
156+
exact (three_three_seven leftys.length
157+
(Clause.toFormula sig X (Clause.substitute σ leftinner)) I leftys hleftysnodup rfl).mp
158+
hleft338
170159

171160
have hrightinnersub : @ValidIn _ X _ _ (rightinner.substitute σ) I := by
172-
exact (three_three_seven rightys.length (Clause.toFormula sig X (Clause.substitute σ rightinner)) I
173-
rightys hrightysnodup rfl).mp
174-
hright338
161+
exact (three_three_seven rightys.length
162+
(Clause.toFormula sig X (Clause.substitute σ rightinner)) I rightys hrightysnodup rfl).mp
163+
hright338
175164

176165
have hDσ_of_negBσ : ∀ β : Assignment X univ, ¬Atom.eval I β (B.substitute σ) →
177166
Formula.eval I β (D.substitute σ) := by
@@ -228,38 +217,3 @@ theorem generalResolutionRuleSound [DecidableEq X] (A B : Atom sig X) (C D : Cla
228217
aesop
229218

230219
exact validclosed_of_valid hCDσ β
231-
232-
theorem generalResolution_soundness [inst : DecidableEq X] {A B : Atom sig X} {C D : Clause sig X}
233-
{σ : Substitution sig X} :
234-
@Soundness _ _ univ _ ([GeneralResolutionRule A B C D σ, GeneralFactorizationRule A B C σ]):= by
235-
rw [Soundness]
236-
intro rule h_rule_general hcond I
237-
238-
intro β h_premise_true
239-
simp_all only [GeneralResolutionRule, Clause, List.append_eq, GeneralFactorizationRule]
240-
rw [List.mem_cons, List.mem_singleton] at h_rule_general
241-
cases h_rule_general
242-
-- proof of resolution rule
243-
next h_res_rule =>
244-
subst h_res_rule
245-
simp only [GeneralResolutionRule] at h_premise_true
246-
simp_all only [Clause, GeneralResolutionRule, List.append_eq, Substitution.eq_1,
247-
MostGeneralUnifier, Unifier, Equality.eq_1, EqualityProblem.eq_1, List.mem_singleton,
248-
forall_eq, MoreGeneral, Set.mem_insert_iff, Set.mem_singleton_iff,
249-
EntailsInterpret, forall_eq_or_imp]
250-
rcases hcond with ⟨hunif, hmgu⟩
251-
repeat rw [← EntailsInterpret] at h_premise_true
252-
clear hmgu
253-
rw [eval_append_iff_eval_or_subst, Formula.eval]
254-
obtain ⟨first, second⟩ := h_premise_true
255-
have hclosed : Formula.closed (Clause.toFormula sig X (Literal.neg A :: C)) := sorry
256-
have hclosed₂ : Formula.closed (Clause.toFormula sig X (Literal.pos B :: D)) := sorry
257-
have s := validIn_of_entails_closed I _ hclosed (by use β)
258-
have t := validIn_of_entails_closed I _ hclosed₂ (by use β)
259-
apply valid_sub_of_valid _ σ at s
260-
apply valid_sub_of_valid _ σ at t
261-
specialize s β
262-
specialize t β
263-
aesop
264-
next h_fact_rule =>
265-
sorry

InferenceInLean/Semantics.lean

Lines changed: 0 additions & 37 deletions
Original file line numberDiff line numberDiff line change
@@ -159,8 +159,6 @@ lemma Assignment.eval_closed_term {sig : Signature} {X : Variables} {I : Interpr
159159
rw [ih]
160160
rw [← hargsareequal]
161161

162-
#check Term.subterms_closed
163-
164162
@[simp]
165163
lemma Assignment.eval_term_with_one_free {univ : Universes} {sig : Signature} {X : Variables}
166164
{I : Interpretation sig univ} [DecidableEq X] {t : Term sig X}
@@ -184,27 +182,6 @@ lemma Assignment.eval_term_with_one_free {univ : Universes} {sig : Signature} {X
184182
· exact a
185183
rw [hargsareequal]
186184

187-
/- @[simp]
188-
def Term.eval_without_free {sig : Signature} {X : Variables} [DecidableEq X] (t : Term sig X) :
189-
t.freeVars = {} → ∀ (f : sig.funs) (args : List (Term sig X)),
190-
t = Term.func f args → args = [] := by
191-
intro hclosed f args heq
192-
by_contra harg
193-
induction' t using Term.induction with x args ih f generalizing args
194-
· sorry
195-
·
196-
simp_all only [imp_false, not_not, Term.func.injEq]
197-
obtain ⟨left, right⟩ := heq
198-
subst left right
199-
apply harg
200-
have hex : ∃ (t : Term sig X), t ∈ args := List.exists_mem_of_ne_nil args harg
201-
rcases hex with ⟨arg, harg⟩
202-
specialize ih (Term.func f args)
203-
have h : ∀ term ∈ args,
204-
Term.freeVars sig X term = ∅ → ∀ (args : List (Term sig X)), term = Term.func f args := sorry
205-
specialize ih h -/
206-
207-
208185
lemma List.reduce_to_empty {α β: Type} {xs : List α} {as : List β}
209186
(hlen : xs.length = as.length) (hzero : as.length = 0 ∨ xs.length = 0) : xs = [] ∧ as = [] := by
210187
simp_all only [List.length_eq_zero, or_self, and_true, List.length_nil]
@@ -286,19 +263,6 @@ lemma Assignment.bigModify_add_nondup {X : Variables} {univ : Universes} [Decida
286263
((β.bigModify xs as).modify x a) x = a := by
287264
simp_all only [modify, ↓reduceIte]
288265

289-
#check List.drop
290-
/- lemma Assignment.bigModify_nodup_erase {X : Variables} {univ : Universes} [DecidableEq X]
291-
(β : Assignment X univ) (xs : List X) (_huniq : xs.Nodup) (as : List univ)
292-
(_hlen : xs.length = as.length) (x : X) (a : univ) (_hx : x ∈ xs)
293-
(i : ℕ) (hiinbounds : i < xs.length) (hi : xs[i] = x) (ha : as[i] = a) :
294-
∀ (y : X), β.bigModify (xs.eraseIdx i) (as.eraseIdx i) y = if x == y then β x -/
295-
#check List.Nodup.erase
296-
/- lemma Assignment.bigModify_add_nondup {X : Variables} {univ : Universes} [DecidableEq X]
297-
(β : Assignment X univ) (xs : List X) (_huniq : xs.Nodup) (as : List univ)
298-
(_hlen : xs.length = as.length) (x : X) (a : univ) (_hx : x ∈ xs) :
299-
(i : ℕ) (hi : xs[i] = x) (a' : univ) (ha : as[i] = a')
300-
(bigModify (xs.dropIdx i) (as.dropIdx i)).modify x a = β := by sorry -/
301-
302266
-- y ∉ [x1, ..., xn] → β[x1 ↦ a1, ..., xn ↦ an] y = β y
303267
@[simp]
304268
lemma Assignment.bigModify_nonmem {X : Variables} {univ : Universes} [DecidableEq X]
@@ -566,7 +530,6 @@ lemma Atom.eval_of_many_free {univ : Universes} {sig : Signature} {X : Variables
566530
exact heval
567531
rw [hargsareequal]
568532

569-
#check Term.freeVars
570533
lemma Formula.eval_of_many_free {univ : Universes} {sig : Signature} {X : Variables}
571534
[inst : DecidableEq X] (I : Interpretation sig univ) (F : Formula sig X) (xs : List X)
572535
(as : List univ) (hlen : xs.length = as.length)

InferenceInLean/Syntax.lean

Lines changed: 23 additions & 29 deletions
Original file line numberDiff line numberDiff line change
@@ -61,6 +61,21 @@ def Term.freeVars : Term sig X -> Set X
6161
| .func _ [] => ∅
6262
| .func f (a :: args) => Term.freeVars a ∪ Term.freeVars (Term.func f args)
6363

64+
lemma Term.arg_subset_of_freeVars {sig : Signature} {X : Variables}
65+
[_inst : DecidableEq X] (args : List (Term sig X)) (f : sig.funs) :
66+
∀ (arg : Term sig X) (_harg : arg ∈ args),
67+
Term.freeVars sig X arg ⊆ Term.freeVars sig X (Term.func f args) := by
68+
intro arg harg
69+
induction' args with arg' args ih
70+
· simp_all only [List.not_mem_nil]
71+
· simp only [List.mem_cons] at harg
72+
rcases harg with harg | harg
73+
· subst harg
74+
simp_all only [Term.freeVars.eq_3, Set.subset_union_left]
75+
· specialize ih harg
76+
rw [Term.freeVars]
77+
exact Set.subset_union_of_subset_right ih (Term.freeVars sig X arg')
78+
6479
inductive Atom where
6580
| pred (p : sig.preds) (args : List (Term sig X))
6681

@@ -145,7 +160,8 @@ instance : Coe (Set <| Clause sig X) (Set <| Formula sig X) :=
145160
def Term.freeVarsList [DecidableEq X] : Term sig X -> List X
146161
| Term.var x => [x]
147162
| Term.func _ [] => []
148-
| Term.func f (a :: args) => List.dedup ((Term.freeVarsList a).append (Term.freeVarsList (Term.func f args)))
163+
| Term.func f (a :: args) => List.dedup
164+
((Term.freeVarsList a).append (Term.freeVarsList (Term.func f args)))
149165

150166
@[simp]
151167
lemma Term.freeVars_sub_freeVarsList [DecidableEq X] (t : Term sig X) :
@@ -181,17 +197,16 @@ def Atom.freeVars_sub_freeVarsList [DecidableEq X] (a : Atom sig X) :
181197
| nil => aesop
182198
| cons head tail ih =>
183199
simp only [freeVars, Set.union_subset_iff]
184-
constructor
200+
unfold freeVarsList
201+
constructor <;> intro x hxinfree
185202
· have hfree_subset := Term.freeVars_sub_freeVarsList sig X head
186-
unfold freeVarsList
187-
intro x hxinfree
188-
simp_all only [List.coe_toFinset, List.append_eq, List.mem_dedup, List.mem_append, Set.mem_setOf_eq]
203+
simp_all only [List.coe_toFinset, List.append_eq, List.mem_dedup, List.mem_append,
204+
Set.mem_setOf_eq]
189205
apply Or.inl
190206
apply hfree_subset
191207
simp_all only
192-
· unfold freeVarsList
193-
intro x hxinfree
194-
simp_all only [List.coe_toFinset, List.append_eq, List.mem_dedup, List.mem_append, Set.mem_setOf_eq]
208+
· simp_all only [List.coe_toFinset, List.append_eq, List.mem_dedup, List.mem_append,
209+
Set.mem_setOf_eq]
195210
apply Or.inr
196211
apply ih
197212
exact hxinfree
@@ -399,24 +414,3 @@ theorem idempotent_iff_inter_empty {sig : Signature} {X : Variables} [BEq X] [BE
399414
rw [Substitution.compose]
400415
refine Substitution.id_of_free_not_in_domain σ (σ x) ?_
401416
aesop
402-
403-
/-
404-
Maybe continue with?
405-
3.3 LEMMA If a is an idempotent substitution and z is an arbitrary substitution, then a is
406-
more general than z iff za = z.
407-
-/
408-
409-
lemma Term.arg_subset_of_freeVars {sig : Signature} {X : Variables}
410-
[_inst : DecidableEq X] (args : List (Term sig X)) (f : sig.funs) :
411-
∀ (arg : Term sig X) (_harg : arg ∈ args),
412-
Term.freeVars sig X arg ⊆ Term.freeVars sig X (Term.func f args) := by
413-
intro arg harg
414-
induction' args with arg' args ih
415-
· simp_all only [List.not_mem_nil]
416-
· simp only [List.mem_cons] at harg
417-
rcases harg with harg | harg
418-
· subst harg
419-
simp_all only [Term.freeVars.eq_3, Set.subset_union_left]
420-
· specialize ih harg
421-
rw [Term.freeVars]
422-
exact Set.subset_union_of_subset_right ih (Term.freeVars sig X arg')

0 commit comments

Comments
 (0)