src/HOL/Tools/Meson/meson_clausify.ML
changeset 42347 53e444ecb525
parent 42336 d63d43e85879
child 42351 ad89f5462cdc
equal deleted inserted replaced
42346:be52d9bed9f6 42347:53e444ecb525
   327     if new_skolemizer then
   327     if new_skolemizer then
   328       let
   328       let
   329         fun skolemize choice_ths =
   329         fun skolemize choice_ths =
   330           skolemize_with_choice_theorems ctxt choice_ths
   330           skolemize_with_choice_theorems ctxt choice_ths
   331           #> simplify (ss_only @{thms all_simps[symmetric]})
   331           #> simplify (ss_only @{thms all_simps[symmetric]})
       
   332         val no_choice = null choice_ths
   332         val pull_out =
   333         val pull_out =
   333           simplify (ss_only @{thms all_simps[symmetric] ex_simps[symmetric]})
   334           if no_choice then
   334         val no_choice = null choice_ths
   335             simplify (ss_only @{thms all_simps[symmetric] ex_simps[symmetric]})
       
   336           else
       
   337             skolemize choice_ths
       
   338         val discharger_th = th |> pull_out
   335         val discharger_th =
   339         val discharger_th =
   336           th |> (if no_choice then pull_out else skolemize choice_ths)
   340           discharger_th |> has_too_many_clauses ctxt (concl_of discharger_th)
       
   341                            ? (to_definitional_cnf_with_quantifiers ctxt
       
   342                               #> pull_out)
   337         val zapped_th =
   343         val zapped_th =
   338           discharger_th |> prop_of |> rename_bound_vars_to_be_zapped ax_no
   344           discharger_th |> prop_of |> rename_bound_vars_to_be_zapped ax_no
   339           |> (if no_choice then
   345           |> (if no_choice then
   340                 Skip_Proof.make_thm thy #> skolemize [cheat_choice] #> cprop_of
   346                 Skip_Proof.make_thm thy #> skolemize [cheat_choice] #> cprop_of
   341               else
   347               else
   362           end
   368           end
   363        else
   369        else
   364          (NONE, (th, ctxt))
   370          (NONE, (th, ctxt))
   365       end
   371       end
   366     else
   372     else
   367       (NONE, (th, ctxt))
   373       (NONE, (th |> has_too_many_clauses ctxt (concl_of th)
       
   374                     ? to_definitional_cnf_with_quantifiers ctxt, ctxt))
   368   end
   375   end
   369 
       
   370 val all_out_ss =
       
   371   Simplifier.clear_ss HOL_basic_ss addsimps @{thms all_simps[symmetric]}
       
   372 
       
   373 val meta_allI = @{lemma "ALL x. P x ==> (!!x. P x)" by auto}
       
   374 
       
   375 fun repeat f x =
       
   376   case try f x of
       
   377     SOME y => repeat f y
       
   378   | NONE => x
       
   379 
       
   380 fun close_thm thy th =
       
   381   fold (Thm.forall_intr o cterm_of thy o Var) (Term.add_vars (prop_of th) []) th
       
   382 
       
   383 fun cnf_axiom_xxx ctxt0 new_skolemizer ax_no th =
       
   384   let
       
   385     val ctxt0 = Variable.global_thm_context th
       
   386     val thy = ProofContext.theory_of ctxt0
       
   387     val choice_ths = choice_theorems thy
       
   388     val (opt, (nnf_th, ctxt)) =
       
   389       nnf_axiom choice_ths new_skolemizer ax_no th ctxt0
       
   390     val skolem_ths = map (old_skolem_theorem_from_def thy) o old_skolem_defs
       
   391     fun make_cnf th = Meson.make_cnf (skolem_ths th) th
       
   392     val (cnf_ths, ctxt) =
       
   393        make_cnf nnf_th ctxt
       
   394        |> (fn (cnf, _) =>
       
   395               let
       
   396                 val _ = tracing ("nondef CNF: " ^ string_of_int (length cnf) ^ " clause(s)")
       
   397                 val sko_th =
       
   398                   nnf_th |> Simplifier.simplify all_out_ss
       
   399                          |> repeat (fn th => th RS meta_allI)
       
   400                          |> Meson.make_xxx_skolem ctxt (skolem_ths nnf_th)
       
   401                          |> close_thm thy
       
   402                          |> Conv.fconv_rule Object_Logic.atomize
       
   403                          |> to_definitional_cnf_with_quantifiers ctxt
       
   404               in make_cnf sko_th ctxt end
       
   405            | p => p)
       
   406     fun intr_imp ct th =
       
   407       Thm.instantiate ([], map (pairself (cterm_of thy))
       
   408                                [(Var (("i", 0), @{typ nat}),
       
   409                                  HOLogic.mk_nat ax_no)])
       
   410                       (zero_var_indexes @{thm skolem_COMBK_D})
       
   411       RS Thm.implies_intr ct th
       
   412   in
       
   413     (NONE (*###*),
       
   414      cnf_ths |> map (introduce_combinators_in_theorem
       
   415                      #> (case opt of SOME (_, ct) => intr_imp ct | NONE => I))
       
   416              |> Variable.export ctxt ctxt0
       
   417              |> finish_cnf
       
   418              |> map Thm.close_derivation)
       
   419   end
       
   420 
       
   421 
   376 
   422 (* Convert a theorem to CNF, with additional premises due to skolemization. *)
   377 (* Convert a theorem to CNF, with additional premises due to skolemization. *)
   423 fun cnf_axiom ctxt0 new_skolemizer ax_no th =
   378 fun cnf_axiom ctxt0 new_skolemizer ax_no th =
   424   let
   379   let
   425     val thy = ProofContext.theory_of ctxt0
   380     val thy = ProofContext.theory_of ctxt0
   426     val choice_ths = choice_theorems thy
   381     val choice_ths = choice_theorems thy
   427     val (opt, (nnf_th, ctxt)) =
   382     val (opt, (nnf_th, ctxt)) =
   428       nnf_axiom choice_ths new_skolemizer ax_no th ctxt0
   383       nnf_axiom choice_ths new_skolemizer ax_no th ctxt0
   429     fun clausify th =
   384     fun clausify th =
   430       make_cnf (if new_skolemizer orelse null choice_ths then
   385       make_cnf (if new_skolemizer orelse null choice_ths then []
   431                   []
   386                 else map (old_skolem_theorem_from_def thy) (old_skolem_defs th))
   432                 else
   387                th ctxt
   433                   map (old_skolem_theorem_from_def thy)
   388     val (cnf_ths, ctxt) = clausify nnf_th
   434                       (old_skolem_defs th)) th ctxt
       
   435     val (cnf_ths, ctxt) =
       
   436       clausify nnf_th
       
   437       |> (fn ([], _) =>
       
   438              if new_skolemizer then
       
   439                let
       
   440 val _ = tracing ("**** NEW CLAUSIFIER, DEF-CNF: " ^ Display.string_of_thm ctxt (to_definitional_cnf_with_quantifiers ctxt nnf_th)) (*###*)
       
   441 val _ = tracing (" *** th: " ^ Display.string_of_thm ctxt th) (*###*)
       
   442                  val (_, (th1, ctxt)) = nnf_axiom choice_ths true ax_no th ctxt0
       
   443 val _ = tracing (" *** th1: " ^ Display.string_of_thm ctxt th1) (*###*)
       
   444                  val th2 = to_definitional_cnf_with_quantifiers ctxt th1
       
   445 val _ = tracing (" *** th2: " ^ Display.string_of_thm ctxt th2) (*###*)
       
   446                  val (ths3, ctxt) = clausify th2
       
   447 val _ = tracing (" *** hd ths3: " ^ Display.string_of_thm ctxt (hd ths3)) (*###*)
       
   448                in (ths3, ctxt) end
       
   449              else
       
   450 let
       
   451 val _ = tracing ("**** OLD CLAUSIFIER, DEF-CNF: " ^ Display.string_of_thm ctxt (to_definitional_cnf_with_quantifiers ctxt nnf_th))
       
   452 (*###*) in
       
   453                clausify (to_definitional_cnf_with_quantifiers ctxt nnf_th)
       
   454 end
       
   455            | p => p)
       
   456     fun intr_imp ct th =
   389     fun intr_imp ct th =
   457       Thm.instantiate ([], map (pairself (cterm_of thy))
   390       Thm.instantiate ([], map (pairself (cterm_of thy))
   458                                [(Var (("i", 0), @{typ nat}),
   391                                [(Var (("i", 0), @{typ nat}),
   459                                  HOLogic.mk_nat ax_no)])
   392                                  HOLogic.mk_nat ax_no)])
   460                       (zero_var_indexes @{thm skolem_COMBK_D})
   393                       (zero_var_indexes @{thm skolem_COMBK_D})