File ‹Tools/cnf.ML›
signature CNF =
sig
val is_atom: term -> bool
val is_literal: term -> bool
val is_clause: term -> bool
val clause_is_trivial: term -> bool
val clause2raw_thm: Proof.context -> thm -> thm
val make_nnf_thm: Proof.context -> term -> thm
val weakening_tac: Proof.context -> int -> tactic
val make_cnf_thm: Proof.context -> term -> thm
val make_cnfx_thm: Proof.context -> term -> thm
val cnf_rewrite_tac: Proof.context -> int -> tactic
val cnfx_rewrite_tac: Proof.context -> int -> tactic
end;
structure CNF : CNF =
struct
fun is_atom \<^Const_>‹False› = false
| is_atom \<^Const_>‹True› = false
| is_atom \<^Const_>‹conj for _ _› = false
| is_atom \<^Const_>‹disj for _ _› = false
| is_atom \<^Const_>‹implies for _ _› = false
| is_atom \<^Const_>‹HOL.eq \<^Type>‹bool› for _ _› = false
| is_atom \<^Const_>‹Not for _› = false
| is_atom _ = true;
fun is_literal \<^Const_>‹Not for x› = is_atom x
| is_literal x = is_atom x;
fun is_clause \<^Const_>‹disj for x y› = is_clause x andalso is_clause y
| is_clause x = is_literal x;
fun clause_is_trivial c =
let
fun dual \<^Const_>‹Not for x› = x
| dual x = \<^Const>‹Not for x›
fun has_duals [] = false
| has_duals (x::xs) = member (op =) xs (dual x) orelse has_duals xs
in
has_duals (HOLogic.disjuncts c)
end;
fun clause2raw_thm ctxt clause =
let
fun not_disj_to_prem i thm =
if i > Thm.nprems_of thm then
thm
else
not_disj_to_prem (i+1)
(Seq.hd (REPEAT_DETERM (resolve_tac ctxt @{thms cnf.clause2raw_not_disj} i) thm))
in
(@{thm cnf.clause2raw_notE} OF [clause])
|> not_disj_to_prem 1
|> Seq.hd o TRYALL (resolve_tac ctxt @{thms cnf.clause2raw_not_not})
|> Thm.assume_prems ~1
end;
fun inst_thm ctxt ts thm =
Thm.instantiate' [] (map (SOME o Thm.cterm_of ctxt) ts) thm;
fun make_nnf_thm ctxt \<^Const_>‹conj for x y› =
let
val thm1 = make_nnf_thm ctxt x
val thm2 = make_nnf_thm ctxt y
in
@{thm cnf.conj_cong} OF [thm1, thm2]
end
| make_nnf_thm ctxt \<^Const_>‹disj for x y› =
let
val thm1 = make_nnf_thm ctxt x
val thm2 = make_nnf_thm ctxt y
in
@{thm cnf.disj_cong} OF [thm1, thm2]
end
| make_nnf_thm ctxt \<^Const_>‹implies for x y› =
let
val thm1 = make_nnf_thm ctxt \<^Const>‹Not for x›
val thm2 = make_nnf_thm ctxt y
in
@{thm cnf.make_nnf_imp} OF [thm1, thm2]
end
| make_nnf_thm ctxt \<^Const_>‹HOL.eq \<^Type>‹bool› for x y› =
let
val thm1 = make_nnf_thm ctxt x
val thm2 = make_nnf_thm ctxt \<^Const>‹Not for x›
val thm3 = make_nnf_thm ctxt y
val thm4 = make_nnf_thm ctxt \<^Const>‹Not for y›
in
@{thm cnf.make_nnf_iff} OF [thm1, thm2, thm3, thm4]
end
| make_nnf_thm _ \<^Const_>‹Not for \<^Const_>‹False›› =
@{thm cnf.make_nnf_not_false}
| make_nnf_thm _ \<^Const_>‹Not for \<^Const_>‹True›› =
@{thm cnf.make_nnf_not_true}
| make_nnf_thm ctxt \<^Const_>‹Not for \<^Const_>‹conj for x y›› =
let
val thm1 = make_nnf_thm ctxt \<^Const>‹Not for x›
val thm2 = make_nnf_thm ctxt \<^Const>‹Not for y›
in
@{thm cnf.make_nnf_not_conj} OF [thm1, thm2]
end
| make_nnf_thm ctxt \<^Const_>‹Not for \<^Const_>‹disj for x y›› =
let
val thm1 = make_nnf_thm ctxt \<^Const>‹Not for x›
val thm2 = make_nnf_thm ctxt \<^Const>‹Not for y›
in
@{thm cnf.make_nnf_not_disj} OF [thm1, thm2]
end
| make_nnf_thm ctxt \<^Const_>‹Not for \<^Const_>‹implies for x y›› =
let
val thm1 = make_nnf_thm ctxt x
val thm2 = make_nnf_thm ctxt \<^Const>‹Not for y›
in
@{thm cnf.make_nnf_not_imp} OF [thm1, thm2]
end
| make_nnf_thm ctxt \<^Const_>‹Not for \<^Const_>‹HOL.eq \<^Type>‹bool› for x y›› =
let
val thm1 = make_nnf_thm ctxt x
val thm2 = make_nnf_thm ctxt \<^Const>‹Not for x›
val thm3 = make_nnf_thm ctxt y
val thm4 = make_nnf_thm ctxt \<^Const>‹Not for y›
in
@{thm cnf.make_nnf_not_iff} OF [thm1, thm2, thm3, thm4]
end
| make_nnf_thm ctxt \<^Const_>‹Not for \<^Const_>‹Not for x›› =
let
val thm1 = make_nnf_thm ctxt x
in
@{thm cnf.make_nnf_not_not} OF [thm1]
end
| make_nnf_thm ctxt t = inst_thm ctxt [t] @{thm cnf.iff_refl};
fun make_under_quantifiers ctxt make t =
let
fun conv ctxt ct =
(case Thm.term_of ct of
Const _ $ Abs _ => Conv.comb_conv (conv ctxt) ct
| Abs _ => Conv.abs_conv (conv o snd) ctxt ct
| Const _ => Conv.all_conv ct
| t => make t RS @{thm eq_reflection})
in HOLogic.mk_obj_eq (conv ctxt (Thm.cterm_of ctxt t)) end
fun make_nnf_thm_under_quantifiers ctxt =
make_under_quantifiers ctxt (make_nnf_thm ctxt)
fun simp_True_False_thm ctxt \<^Const_>‹conj for x y› =
let
val thm1 = simp_True_False_thm ctxt x
val x'= (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm1
in
if x' = \<^Const>‹False› then
@{thm cnf.simp_TF_conj_False_l} OF [thm1]
else
let
val thm2 = simp_True_False_thm ctxt y
val y' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm2
in
if x' = \<^Const>‹True› then
@{thm cnf.simp_TF_conj_True_l} OF [thm1, thm2]
else if y' = \<^Const>‹False› then
@{thm cnf.simp_TF_conj_False_r} OF [thm2]
else if y' = \<^Const>‹True› then
@{thm cnf.simp_TF_conj_True_r} OF [thm1, thm2]
else
@{thm cnf.conj_cong} OF [thm1, thm2]
end
end
| simp_True_False_thm ctxt \<^Const_>‹disj for x y› =
let
val thm1 = simp_True_False_thm ctxt x
val x' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm1
in
if x' = \<^Const>‹True› then
@{thm cnf.simp_TF_disj_True_l} OF [thm1]
else
let
val thm2 = simp_True_False_thm ctxt y
val y' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm2
in
if x' = \<^Const>‹False› then
@{thm cnf.simp_TF_disj_False_l} OF [thm1, thm2]
else if y' = \<^Const>‹True› then
@{thm cnf.simp_TF_disj_True_r} OF [thm2]
else if y' = \<^Const>‹False› then
@{thm cnf.simp_TF_disj_False_r} OF [thm1, thm2]
else
@{thm cnf.disj_cong} OF [thm1, thm2]
end
end
| simp_True_False_thm ctxt t = inst_thm ctxt [t] @{thm cnf.iff_refl};
fun make_cnf_thm ctxt t =
let
fun make_cnf_thm_from_nnf \<^Const_>‹conj for x y› =
let
val thm1 = make_cnf_thm_from_nnf x
val thm2 = make_cnf_thm_from_nnf y
in
@{thm cnf.conj_cong} OF [thm1, thm2]
end
| make_cnf_thm_from_nnf \<^Const_>‹disj for x y› =
let
fun make_cnf_disj_thm \<^Const_>‹conj for x1 x2› y' =
let
val thm1 = make_cnf_disj_thm x1 y'
val thm2 = make_cnf_disj_thm x2 y'
in
@{thm cnf.make_cnf_disj_conj_l} OF [thm1, thm2]
end
| make_cnf_disj_thm x' \<^Const_>‹conj for y1 y2› =
let
val thm1 = make_cnf_disj_thm x' y1
val thm2 = make_cnf_disj_thm x' y2
in
@{thm cnf.make_cnf_disj_conj_r} OF [thm1, thm2]
end
| make_cnf_disj_thm x' y' =
inst_thm ctxt [\<^Const>‹disj for x' y'›] @{thm cnf.iff_refl}
val thm1 = make_cnf_thm_from_nnf x
val thm2 = make_cnf_thm_from_nnf y
val x' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm1
val y' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm2
val disj_thm = @{thm cnf.disj_cong} OF [thm1, thm2]
in
@{thm cnf.iff_trans} OF [disj_thm, make_cnf_disj_thm x' y']
end
| make_cnf_thm_from_nnf t = inst_thm ctxt [t] @{thm cnf.iff_refl}
val nnf_thm = make_nnf_thm_under_quantifiers ctxt t
val nnf = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) nnf_thm
val simp_thm = simp_True_False_thm ctxt nnf
val simp = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) simp_thm
val cnf_thm = make_under_quantifiers ctxt make_cnf_thm_from_nnf simp
in
@{thm cnf.iff_trans} OF [@{thm cnf.iff_trans} OF [nnf_thm, simp_thm], cnf_thm]
end;
fun make_cnfx_thm ctxt t =
let
val var_id = Unsynchronized.ref 0
fun new_free () =
Free ("cnfx_" ^ string_of_int (Unsynchronized.inc var_id), \<^Type>‹bool›)
fun make_cnfx_thm_from_nnf \<^Const_>‹conj for x y› =
let
val thm1 = make_cnfx_thm_from_nnf x
val thm2 = make_cnfx_thm_from_nnf y
in
@{thm cnf.conj_cong} OF [thm1, thm2]
end
| make_cnfx_thm_from_nnf \<^Const_>‹disj for x y› =
if is_clause x andalso is_clause y then
inst_thm ctxt [\<^Const>‹disj for x y›] @{thm cnf.iff_refl}
else if is_literal y orelse is_literal x then
let
fun make_cnfx_disj_thm \<^Const_>‹conj for x1 x2› y' =
let
val thm1 = make_cnfx_disj_thm x1 y'
val thm2 = make_cnfx_disj_thm x2 y'
in
@{thm cnf.make_cnf_disj_conj_l} OF [thm1, thm2]
end
| make_cnfx_disj_thm x' \<^Const_>‹conj for y1 y2› =
let
val thm1 = make_cnfx_disj_thm x' y1
val thm2 = make_cnfx_disj_thm x' y2
in
@{thm cnf.make_cnf_disj_conj_r} OF [thm1, thm2]
end
| make_cnfx_disj_thm \<^Const_>‹Ex \<^Type>‹bool› for x'› y' =
let
val thm1 = inst_thm ctxt [x', y'] @{thm cnf.make_cnfx_disj_ex_l}
val var = new_free ()
val thm2 = make_cnfx_disj_thm (betapply (x', var)) y'
val thm3 = Thm.forall_intr (Thm.cterm_of ctxt var) thm2
val thm4 = Thm.strip_shyps (thm3 COMP allI)
val thm5 = Thm.strip_shyps (thm4 RS @{thm cnf.make_cnfx_ex_cong})
in
@{thm cnf.iff_trans} OF [thm1, thm5]
end
| make_cnfx_disj_thm x' \<^Const_>‹Ex \<^Type>‹bool› for y'› =
let
val thm1 = inst_thm ctxt [x', y'] @{thm cnf.make_cnfx_disj_ex_r}
val var = new_free ()
val thm2 = make_cnfx_disj_thm x' (betapply (y', var))
val thm3 = Thm.forall_intr (Thm.cterm_of ctxt var) thm2
val thm4 = Thm.strip_shyps (thm3 COMP allI)
val thm5 = Thm.strip_shyps (thm4 RS @{thm cnf.make_cnfx_ex_cong})
in
@{thm cnf.iff_trans} OF [thm1, thm5]
end
| make_cnfx_disj_thm x' y' =
inst_thm ctxt [\<^Const>‹disj for x' y'›] @{thm cnf.iff_refl}
val thm1 = make_cnfx_thm_from_nnf x
val thm2 = make_cnfx_thm_from_nnf y
val x' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm1
val y' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm2
val disj_thm = @{thm cnf.disj_cong} OF [thm1, thm2]
in
@{thm cnf.iff_trans} OF [disj_thm, make_cnfx_disj_thm x' y']
end
else
let
val thm1 = inst_thm ctxt [x, y] @{thm cnf.make_cnfx_newlit}
val var = new_free ()
val body = \<^Const>‹conj for \<^Const>‹disj for x var› \<^Const>‹disj for y \<^Const>‹Not for var›››
val thm2 = make_cnfx_thm_from_nnf body
val thm3 = Thm.forall_intr (Thm.cterm_of ctxt var) thm2
val thm4 = Thm.strip_shyps (thm3 COMP allI)
val thm5 = Thm.strip_shyps (thm4 RS @{thm cnf.make_cnfx_ex_cong})
in
@{thm cnf.iff_trans} OF [thm1, thm5]
end
| make_cnfx_thm_from_nnf t = inst_thm ctxt [t] @{thm cnf.iff_refl}
val nnf_thm = make_nnf_thm_under_quantifiers ctxt t
val nnf = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) nnf_thm
val simp_thm = simp_True_False_thm ctxt nnf
val simp = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) simp_thm
val _ = (var_id := Frees.fold (fn ((name, _), _) => fn max =>
let
val idx =
(case try (unprefix "cnfx_") name of
SOME s => Int.fromString s
| NONE => NONE)
in
Int.max (max, the_default 0 idx)
end) (Frees.build (Frees.add_frees simp)) 0)
val cnfx_thm = make_under_quantifiers ctxt make_cnfx_thm_from_nnf simp
in
@{thm cnf.iff_trans} OF [@{thm cnf.iff_trans} OF [nnf_thm, simp_thm], cnfx_thm]
end;
fun weakening_tac ctxt i =
dresolve_tac ctxt @{thms cnf.weakening_thm} i THEN assume_tac ctxt (i+1);
fun cnf_rewrite_tac ctxt i =
Subgoal.FOCUS (fn {prems, context = ctxt', ...} =>
let
val cnf_thms = map (make_cnf_thm ctxt' o HOLogic.dest_Trueprop o Thm.prop_of) prems
val cut_thms = map (fn (th, pr) => @{thm cnf.cnftac_eq_imp} OF [th, pr]) (cnf_thms ~~ prems)
in
cut_facts_tac cut_thms 1
end) ctxt i
THEN SELECT_GOAL (fn thm =>
let
val n = Logic.count_prems ((Term.strip_all_body o fst o Logic.dest_implies o Thm.prop_of) thm)
in
PRIMITIVE (funpow (n div 2) (Seq.hd o weakening_tac ctxt 1)) thm
end) i;
fun cnfx_rewrite_tac ctxt i =
Subgoal.FOCUS (fn {prems, context = ctxt', ...} =>
let
val cnfx_thms = map (make_cnfx_thm ctxt' o HOLogic.dest_Trueprop o Thm.prop_of) prems
val cut_thms = map (fn (th, pr) => @{thm cnf.cnftac_eq_imp} OF [th, pr]) (cnfx_thms ~~ prems)
in
cut_facts_tac cut_thms 1
end) ctxt i
THEN SELECT_GOAL (fn thm =>
let
val n = Logic.count_prems ((Term.strip_all_body o fst o Logic.dest_implies o Thm.prop_of) thm)
in
PRIMITIVE (funpow (n div 2) (Seq.hd o weakening_tac ctxt 1)) thm
end) i;
end;