datatype form = VAR "num"
| FALSE
| DISJ "form" "form"
| NOT "form"
| DIAM "form"
record 'a frame =
world :: "'a ⇒ bool"
rel :: "'a ⇒ 'a ⇒ bool"
record 'a model = "'a frame" +
valt :: "num ⇒ 'a ⇒ bool"
fun satis :: "'a model ⇒ 'a ⇒ form ⇒ bool"
where
"satis M w (VAR p) ⟷ (valt M p w ∧ world M w)"
| "satis M w FALSE ⟷ False"
| "satis M w (NOT f) ⟷ world M w ∧ ¬ satis M w f"
| "satis M w (DISJ f1 f2) ⟷ satis M w f1 ∨ satis M w f2"
| "satis M w (DIAM f) ⟷ world M w ∧ (∃ v. rel M w v ∧ world M v ∧ satis M v f)"
lemma satis_in_worlds: "satis M w f ⟹ world M w"
proof (induct rule:form.induct)
case (DISJ x1a x2)
then show ?case by auto (* not sure if this works, but auto is roughly simp + blast *)
next
case (DIAM x)
then show ?case
by fastforce
qed simp_all
definition BOX_def : "BOX f = NOT (DIAM (NOT f))"
definition AND_def : "AND f1 f2 = NOT (DISJ (NOT f1) (NOT f2))"
definition TRUE_def : "TRUE = NOT FALSE"
fun subst :: "(num ⇒ form) ⇒ form ⇒ form"
where
"subst s FALSE = FALSE"
| "subst s (VAR p) = s p"
| "subst s (DISJ f1 f2) = DISJ (subst s f1) (subst s f2)"
| "subst s (NOT f) = NOT (subst s f)"
| "subst s (DIAM f) = DIAM (subst s f)"
lemma subst_BOX[simp] : "subst f (BOX form) = BOX (subst f form)"
by (simp add: BOX_def)
lemma satis_AND : "satis M w (AND f1 f2) ⟷ satis M w f1 ∧ satis M w f2"
proof -
show ?thesis apply (simp add: AND_def)
by (meson satis_in_worlds)
qed
definition satis_set_def :
"satis_set M w Σ ≡ (⋀a. a ∈ Σ ⟹ satis M w a)"
(*
val demodalize_def = Define`
demodalize FALSE = FALSE /\
demodalize (VAR p) = VAR p /\
demodalize (DISJ form1 form2) = DISJ (demodalize form1) (demodalize form2) /\
demodalize (NOT form) = NOT (demodalize form) /\
demodalize (DIAM form) = demodalize form`;
*)
fun demodalize :: "form ⇒ form"
where
"demodalize FALSE = FALSE"
| "demodalize (VAR p) = VAR p"
| "demodalize (DISJ form1 form2) = DISJ (demodalize form1) (demodalize form2)"
| "demodalize (NOT form) = NOT (demodalize form)"
| "demodalize (DIAM form) = demodalize form"
thm HOL.eq_reflection
fun propform :: "form ⇒ bool"
where
"propform FALSE = True"
| "propform (VAR p) = True"
| "propform (DISJ f1 f2) = (propform f1 ∧ propform f2)"
| "propform (NOT f) = propform f"
| "propform (DIAM f) = False"
definition IMP_def : "IMP f1 f2 = DISJ (NOT f1) f2"
fun peval :: "(num ⇒ bool) ⇒ form ⇒ bool" where
"peval σ (VAR p) = σ p"
| "peval σ (DISJ f1 f2) = (peval σ f1 ∨ peval σ f2)"
| "peval σ FALSE = False"
| "peval σ (NOT f) = (¬ peval σ f)"
| "peval σ (DIAM f) = False"
definition ptaut_def : "ptaut f ⟷ propform f ∧ (∀σ. peval σ f)"
inductive Kproof :: "form list ⇒ bool" where
"Kproof []"
| "Kproof p ∧ List.member p (IMP f1 f2) ∧ List.member p f1 ⟹ Kproof (p @ [f2])"
| "Kproof p ∧ List.member p f ⟹ Kproof (p @ [subst σ f])"
| "Kproof p ∧ List.member p f ⟹ Kproof (p @ [BOX f])"
| "Kproof p ⟹ Kproof (p @ [IMP (BOX (IMP form1 form2)) (IMP (BOX form1) (BOX form2))])"
| "Kproof p ⟹ Kproof (p @ [IMP (DIAM form) (NOT (BOX (NOT form)))])"
| "Kproof p ∧ ptaut f ⟹ Kproof (p @ [f])"
lemma demodalize_IMP : "demodalize (IMP f1 f2) = IMP (demodalize f1) (demodalize f2)"
proof -
show ?thesis
by (simp add: IMP_def)
qed
lemma ptaut_MP : assumes a1 : "ptaut (IMP f1 f2)" and a2: "ptaut f1" shows b: "ptaut f2"
proof -
show ?thesis using a1 a2 by (simp add: IMP_def ptaut_def)
qed
find_theorems "List.member (a @ [b])"
definition valid_frame_state_def:
"valid_frame_state Fr w phi ≡ (⋀M. frame.truncate M = Fr ⟹ satis M w phi)"
lemma propform_demodalize : "ptaut f ⟹ ptaut (demodalize f)"
sorry
lemma demodalize_subst: "demodalize (subst f form) = demodalize (subst f (demodalize form))"
sorry
lemma peval_demodalize_subst_eq :
" propform form ⟹
(peval σ (demodalize (subst f form)) \<equiv> peval (peval σ o demodalize o f) form)"
sorry
thm disjE
lemma exercise_1_6_2 : "Kproof p ⟹ List.member p f ⟹ ptaut (demodalize f)"
proof (induction arbitrary: f rule: Kproof.induct)
case 1
then show ?case
using member_rec(2) by force
next
case a:(2 p f1 f2)
hence a1 : "List.member p f ∨ f = f2"
by (metis in_set_member rotate1.simps(2) set_ConsD set_rotate1)
then
show ?case
by (metis a.IH demodalize_IMP ptaut_MP)
next
case a : (3 p f σ g)
hence a1 : "List.member p g ∨ g = subst σ f"
by (metis member_def rotate1.simps(2) set_ConsD set_rotate1)
with a show ?case
proof
assume "List.member p g" show ?case sorry
next
assume "g = subst σ f" show ?case sorry
qed
(*apply (rule disjE)*)
next
case (4 p f)
then show ?case sorry
next
case (5 p form1 form2)
then show ?case sorry
next
case (6 p form)
then show ?case sorry
next
case (7 p f)
then show ?case sorry
qed
The last theorem is where I am messed up.
I proved "A \/ B", and want to do case splitting.
I tried writing:
case a : (3 p f σ g)
hence a1 : "List.member p g ∨ g = subst σ f"
by (metis member_def rotate1.simps(2) set_ConsD set_rotate1)
with a show ?case
proof
assume "List.member p g" show ?case sorry
next
assume "g = subst σ f" show ?case sorry
qed
following prog-prove page 45.
the usual pattern (I use, see @Mathias Fleury 's below) is:
moreover {
assume A
have "foo"
}
moreover {
assume B
have "foo"
}
ultimately show "foo"
Yong Kiam said:
the usual pattern is:
moreover { assume A have "foo" } moreover { assume B have "foo" } ultimately show "foo"
Thanks! I will try this one as well.
In one min. I still want to know why the syntax I wrote is wrong.
I tried apply as well:
image.png
Does not work either.
Yong Kiam said:
the usual pattern is:
moreover { assume A have "foo" } moreover { assume B have "foo" } ultimately show "foo"
May I please ask what is wrong here?
In general I prefer the consider pattern or relying on cases to do the splitting
Yiming Xu said:
Yong Kiam said:
the usual pattern is:
moreover { assume A have "foo" } moreover { assume B have "foo" } ultimately show "foo"
May I please ask what is wrong here?
Did you read the error message?
the error message says "I am expecting a proof here", not Isar commands
Mathias Fleury said:
In general I prefer the consider pattern or relying on cases to do the splitting
ah it looks like consider
is nicer indeed, thanks
yes, but the error is 3 lines above: you didn't finish the proof of the have
like in nearly every system on earth: the first error message is the most important one…
I am mostly copying from the examples, which line caused the proof mode?
which of the lines should be followed by a proof?
by.
There are only 2 options: assume or have
I see.
I was just about to say by is apply...done so I will be confused if it cause it...
That is the apply one.
That is the one in prog-prove. I see. Thank you!
Yiming Xu said:
this one uses unrestricted proof
which is frowned upon by some people... maybe proof (rule disj1E)
or something instead
Thanks! I will avoid this style and type up the rules explicitly.
Last updated: Feb 01 2025 at 20:19 UTC