Skip to content

Commit a8a6f71

Browse files
authored
fix: add monotonicity lemmas for universal quantifiers (#8403)
This PR adds missing monotonicity lemmas for universal quantifiers, that are used in defining (co)inductive predicates.
1 parent 9ad4414 commit a8a6f71

File tree

2 files changed

+32
-0
lines changed

2 files changed

+32
-0
lines changed

src/Init/Internal/Order/Basic.lean

Lines changed: 12 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -877,6 +877,12 @@ instance ImplicationOrder.instCompleteLattice : CompleteLattice ImplicationOrder
877877
@monotone _ _ _ ImplicationOrder.instOrder (fun x => (Exists (f x))) :=
878878
fun x y hxy ⟨w, hw⟩ => ⟨w, monotone_apply w f h x y hxy hw⟩
879879

880+
@[partial_fixpoint_monotone] theorem implication_order_monotone_forall
881+
{α} [PartialOrder α] {β} (f : α → β → ImplicationOrder)
882+
(h : monotone f) :
883+
@monotone _ _ _ ImplicationOrder.instOrder (fun x => ∀ y, f x y) :=
884+
fun x y hxy h₂ y₁ => monotone_apply y₁ f h x y hxy (h₂ y₁)
885+
880886
@[partial_fixpoint_monotone] theorem implication_order_monotone_and
881887
{α} [PartialOrder α] (f₁ : α → ImplicationOrder) (f₂ : α → ImplicationOrder)
882888
(h₁ : @monotone _ _ _ ImplicationOrder.instOrder f₁)
@@ -931,6 +937,12 @@ def ReverseImplicationOrder.instCompleteLattice : CompleteLattice ReverseImplica
931937
@monotone _ _ _ ReverseImplicationOrder.instOrder (fun x => Exists (f x)) :=
932938
fun x y hxy ⟨w, hw⟩ => ⟨w, monotone_apply w f h x y hxy hw⟩
933939

940+
@[partial_fixpoint_monotone] theorem coind_monotone_forall
941+
{α} [PartialOrder α] {β} (f : α → β → ReverseImplicationOrder)
942+
(h : monotone f) :
943+
@monotone _ _ _ ReverseImplicationOrder.instOrder (fun x => ∀ y, f x y) :=
944+
fun x y hxy h₂ y₁ => monotone_apply y₁ f h x y hxy (h₂ y₁)
945+
934946
@[partial_fixpoint_monotone] theorem coind_monotone_and
935947
{α} [PartialOrder α] (f₁ : α → Prop) (f₂ : α → Prop)
936948
(h₁ : @monotone _ _ _ ReverseImplicationOrder.instOrder f₁)

tests/lean/run/coinductive_predicates.lean

Lines changed: 20 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -156,3 +156,23 @@ theorem infseq_coinduction_principle_2:
156156
exact rel2
157157
apply Exists.intro mid
158158
exact ⟨ s, h₆ ⟩
159+
160+
-- Automata theory example that involves forall quantifier
161+
def DFA (Q : Type) (A : Type) : Type := Q → (Bool × (A → Q))
162+
163+
def language_equivalent (automaton : DFA Q A) (q₁ q₂ : Q) : Prop :=
164+
let ⟨o₁, t₁⟩ := automaton q₁
165+
let ⟨o₂, t₂⟩ := automaton q₂
166+
o₁ = o₂ ∧ (∀ a : A, language_equivalent automaton (t₁ a) (t₂ a))
167+
greatest_fixpoint
168+
169+
/--
170+
info: language_equivalent.fixpoint_induct {Q A : Type} (automaton : DFA Q A) (x : Q → Q → Prop)
171+
(y :
172+
∀ (x_1 x_2 : Q),
173+
x x_1 x_2 →
174+
(automaton x_1).fst = (automaton x_2).fst ∧ ∀ (a : A), x ((automaton x_1).snd a) ((automaton x_2).snd a))
175+
(x✝ x✝¹ : Q) : x x✝ x✝¹ → language_equivalent automaton x✝ x✝¹
176+
-/
177+
#guard_msgs in
178+
#check language_equivalent.fixpoint_induct

0 commit comments

Comments
 (0)