| author | wenzelm | 
| Sat, 09 Mar 2024 19:57:08 +0100 | |
| changeset 79836 | c69ae2b8987e | 
| parent 69989 | bcba61d92558 | 
| permissions | -rw-r--r-- | 
| 59720 | 1 | (* Title: HOL/TPTP/TPTP_Proof_Reconstruction_Test_Units.thy | 
| 55596 | 2 | Author: Nik Sultana, Cambridge University Computer Laboratory | 
| 3 | ||
| 4 | Unit tests for proof reconstruction module. | |
| 5 | *) | |
| 6 | ||
| 60649 
e007aa6a8aa2
more explicit use of context and elimination of Thm.theory_of_thm, although unclear (and untested?) situations remain;
 wenzelm parents: 
59720diff
changeset | 7 | theory TPTP_Proof_Reconstruction_Test_Units | 
| 55596 | 8 | imports TPTP_Test TPTP_Proof_Reconstruction | 
| 9 | begin | |
| 10 | ||
| 56281 | 11 | declare [[ML_exception_trace, ML_print_depth = 200]] | 
| 55596 | 12 | |
| 60926 | 13 | declare [[tptp_trace_reconstruction = true]] | 
| 55596 | 14 | |
| 15 | lemma "! (X1 :: bool) (X2 :: bool) (X3 :: bool) (X4 :: bool) (X5 :: bool). P \<Longrightarrow> ! (X1 :: bool) (X3 :: bool) (X5 :: bool). P" | |
| 63167 | 16 | apply (tactic \<open>canonicalise_qtfr_order @{context} 1\<close>)
 | 
| 55596 | 17 | oops | 
| 18 | ||
| 19 | lemma "! (X1 :: bool) (X2 :: bool) (X3 :: bool) (X4 :: bool) (X5 :: bool). P \<Longrightarrow> ! (X1 :: bool) (X3 :: bool) (X5 :: bool). P" | |
| 63167 | 20 | apply (tactic \<open>canonicalise_qtfr_order @{context} 1\<close>)
 | 
| 55596 | 21 | apply (rule allI)+ | 
| 63167 | 22 | apply (tactic \<open>nominal_inst_parametermatch_tac @{context} @{thm allE} 1\<close>)+
 | 
| 55596 | 23 | oops | 
| 24 | ||
| 25 | (*Could test bind_tac further with NUM667^1 inode43*) | |
| 26 | ||
| 27 | (* | |
| 28 | (* SEU581^2.p_nux *) | |
| 29 |      (* (Annotated_step ("inode1", "bind"), *)
 | |
| 61076 | 30 | lemma "\<forall>(SV5::TPTP_Interpret.ind \<Rightarrow> bool) | 
| 31 | SV6::TPTP_Interpret.ind. | |
| 55596 | 32 | (bnd_in (bnd_dsetconstr bnd_sK1_A bnd_sK2_SY15) | 
| 33 | (bnd_powerset bnd_sK1_A) = | |
| 34 | bnd_in (bnd_dsetconstr SV6 SV5) | |
| 35 | (bnd_powerset SV6)) = | |
| 36 | False \<Longrightarrow> | |
| 37 | (bnd_in (bnd_dsetconstr bnd_sK1_A bnd_sK2_SY15) | |
| 38 | (bnd_powerset bnd_sK1_A) = | |
| 39 | bnd_in (bnd_dsetconstr bnd_sK1_A bnd_sK2_SY15) | |
| 40 | (bnd_powerset bnd_sK1_A)) = | |
| 41 | False" | |
| 42 | ML_prf {*
 | |
| 43 | open TPTP_Syntax; | |
| 44 | open TPTP_Proof; | |
| 45 | ||
| 46 | ||
| 47 | val binds = | |
| 48 | [Bind ("SV6", Atom (THF_Atom_term (Term_Func (Uninterpreted "sK1_A", [])))), Bind ("SV5", Quant (Lambda, [("SX0", SOME (Fmla_type (Atom (THF_Atom_term (Term_Func (TypeSymbol Type_Ind, []))))))], Fmla (Interpreted_ExtraLogic Apply, [Atom (THF_Atom_term (Term_Func (Uninterpreted "sK2_SY15", []))), Atom (THF_Atom_term (Term_Var "SX0"))])))]
 | |
| 49 | (* |> TPTP_Reconstruct.permute *) | |
| 50 | ||
| 51 | (* | |
| 52 | val binds = | |
| 53 | [Bind ("SV5", Quant (Lambda, [("SX0", SOME (Fmla_type (Atom (THF_Atom_term (Term_Func (TypeSymbol Type_Ind, []))))))], Fmla (Interpreted_ExtraLogic Apply, [Atom (THF_Atom_term (Term_Func (Uninterpreted "sK2_SY15", []))), Atom (THF_Atom_term (Term_Var "SX0"))]))),
 | |
| 54 | Bind ("SV6", Atom (THF_Atom_term (Term_Func (Uninterpreted "sK1_A", []))))
 | |
| 55 | ] | |
| 56 | *) | |
| 57 | ||
| 58 | val tec = | |
| 59 | (* | |
| 60 |   map (bind_tac @{context} (hd prob_names)) binds
 | |
| 61 | |> FIRST | |
| 62 | *) | |
| 63 |   bind_tac @{context} (hd prob_names) binds
 | |
| 64 | *} | |
| 65 | apply (tactic {*tec*})
 | |
| 66 | done | |
| 67 | ||
| 68 |      (* (Annotated_step ("inode2", "bind"), *)
 | |
| 61076 | 69 | lemma "\<forall>(SV7::TPTP_Interpret.ind) SV8::TPTP_Interpret.ind. | 
| 55596 | 70 | (bnd_subset SV8 SV7 = | 
| 71 | bnd_subset (bnd_dsetconstr bnd_sK1_A bnd_sK2_SY15) | |
| 72 | bnd_sK1_A) = | |
| 73 | False \<or> | |
| 74 | bnd_in SV8 (bnd_powerset SV7) = False \<Longrightarrow> | |
| 75 | (bnd_subset (bnd_dsetconstr bnd_sK1_A bnd_sK2_SY15) | |
| 76 | bnd_sK1_A = | |
| 77 | bnd_subset (bnd_dsetconstr bnd_sK1_A bnd_sK2_SY15) | |
| 78 | bnd_sK1_A) = | |
| 79 | False \<or> | |
| 80 | bnd_in (bnd_dsetconstr bnd_sK1_A bnd_sK2_SY15) | |
| 81 | (bnd_powerset bnd_sK1_A) = | |
| 82 | False" | |
| 83 | ML_prf {*
 | |
| 84 | open TPTP_Syntax; | |
| 85 | open TPTP_Proof; | |
| 86 | ||
| 87 | ||
| 88 | val binds = | |
| 89 | [Bind ("SV8", Fmla (Interpreted_ExtraLogic Apply, [Fmla (Interpreted_ExtraLogic Apply, [Atom (THF_Atom_term (Term_Func (Uninterpreted "dsetconstr", []))), Atom (THF_Atom_term (Term_Func (Uninterpreted "sK1_A", [])))]), Quant (Lambda, [("SX0", SOME (Fmla_type (Atom (THF_Atom_term (Term_Func (TypeSymbol Type_Ind, []))))))], Fmla (Interpreted_ExtraLogic Apply, [Atom (THF_Atom_term (Term_Func (Uninterpreted "sK2_SY15", []))), Atom (THF_Atom_term (Term_Var "SX0"))]))])), Bind ("SV7", Atom (THF_Atom_term (Term_Func (Uninterpreted "sK1_A", []))))]
 | |
| 90 | (* |> TPTP_Reconstruct.permute *) | |
| 91 | ||
| 92 | val tec = | |
| 93 | (* | |
| 94 |   map (bind_tac @{context} (hd prob_names)) binds
 | |
| 95 | |> FIRST | |
| 96 | *) | |
| 97 |   bind_tac @{context} (hd prob_names) binds
 | |
| 98 | *} | |
| 99 | apply (tactic {*tec*})
 | |
| 100 | done | |
| 101 | *) | |
| 102 | ||
| 103 | (* | |
| 104 | from SEU897^5 | |
| 105 | lemma " | |
| 106 | \<forall>SV9 SV10 SV11 SV12 SV13 SV14. | |
| 107 | (((((bnd_sK5_SY14 SV14 SV13 SV12 = SV11) = False \<or> | |
| 108 | (bnd_sK4_SX0 = SV10 (bnd_sK5_SY14 SV9 SV10 SV11)) = | |
| 109 | False) \<or> | |
| 110 | bnd_cR SV14 = False) \<or> | |
| 111 | (SV12 = SV13 SV14) = False) \<or> | |
| 112 | bnd_cR SV9 = False) \<or> | |
| 113 | (SV11 = SV10 SV9) = False \<Longrightarrow> | |
| 114 | \<forall>SV14 SV13 SV12 SV10 SV9. | |
| 115 | (((((bnd_sK5_SY14 SV14 SV13 SV12 = | |
| 116 | bnd_sK5_SY14 SV14 SV13 SV12) = | |
| 117 | False \<or> | |
| 118 | (bnd_sK4_SX0 = | |
| 119 | SV10 | |
| 120 | (bnd_sK5_SY14 SV9 SV10 | |
| 121 | (bnd_sK5_SY14 SV14 SV13 SV12))) = | |
| 122 | False) \<or> | |
| 123 | bnd_cR SV14 = False) \<or> | |
| 124 | (SV12 = SV13 SV14) = False) \<or> | |
| 125 | bnd_cR SV9 = False) \<or> | |
| 126 | (bnd_sK5_SY14 SV14 SV13 SV12 = SV10 SV9) = False" | |
| 127 | ML_prf {*
 | |
| 128 | open TPTP_Syntax; | |
| 129 | open TPTP_Proof; | |
| 130 | ||
| 131 | val binds = | |
| 132 | [Bind ("SV11", Fmla (Interpreted_ExtraLogic Apply, [Fmla (Interpreted_ExtraLogic Apply, [Fmla (Interpreted_ExtraLogic Apply, [Atom (THF_Atom_term (Term_Func (Uninterpreted "sK5_SY14", []))), Atom (THF_Atom_term (Term_Var "SV14"))]), Atom (THF_Atom_term (Term_Var "SV13"))]), Atom (THF_Atom_term (Term_Var "SV12"))]))]
 | |
| 133 | ||
| 134 | val tec = bind_tac @{context} (hd prob_names) binds
 | |
| 135 | *} | |
| 136 | apply (tactic {*tec*})
 | |
| 137 | done | |
| 138 | *) | |
| 139 | ||
| 140 | ||
| 141 | subsection "Interpreting the inferences" | |
| 142 | ||
| 143 | (*from SET598^5 | |
| 144 | lemma "(bnd_sK1_X = (\<lambda>SY17. bnd_sK2_Y SY17 \<and> bnd_sK3_Z SY17) \<longrightarrow> | |
| 145 | ((\<forall>SY25. bnd_sK1_X SY25 \<longrightarrow> bnd_sK2_Y SY25) \<and> | |
| 146 | (\<forall>SY26. bnd_sK1_X SY26 \<longrightarrow> bnd_sK3_Z SY26)) \<and> | |
| 147 | (\<forall>SY27. | |
| 148 | (\<forall>SY21. SY27 SY21 \<longrightarrow> bnd_sK2_Y SY21) \<and> | |
| 149 | (\<forall>SY15. SY27 SY15 \<longrightarrow> bnd_sK3_Z SY15) \<longrightarrow> | |
| 150 | (\<forall>SY30. SY27 SY30 \<longrightarrow> bnd_sK1_X SY30))) = | |
| 151 | False \<Longrightarrow> | |
| 152 | (\<not> (bnd_sK1_X = (\<lambda>SY17. bnd_sK2_Y SY17 \<and> bnd_sK3_Z SY17) \<longrightarrow> | |
| 153 | ((\<forall>SY25. bnd_sK1_X SY25 \<longrightarrow> bnd_sK2_Y SY25) \<and> | |
| 154 | (\<forall>SY26. bnd_sK1_X SY26 \<longrightarrow> bnd_sK3_Z SY26)) \<and> | |
| 155 | (\<forall>SY27. | |
| 156 | (\<forall>SY21. SY27 SY21 \<longrightarrow> bnd_sK2_Y SY21) \<and> | |
| 157 | (\<forall>SY15. SY27 SY15 \<longrightarrow> bnd_sK3_Z SY15) \<longrightarrow> | |
| 158 | (\<forall>SY30. SY27 SY30 \<longrightarrow> bnd_sK1_X SY30)))) = | |
| 159 | True" | |
| 160 | apply (tactic {*polarity_switch_tac @{context}*})
 | |
| 161 | done | |
| 162 | lemma " | |
| 163 | (((\<forall>SY25. bnd_sK1_X SY25 \<longrightarrow> bnd_sK2_Y SY25) \<and> | |
| 164 | (\<forall>SY26. bnd_sK1_X SY26 \<longrightarrow> bnd_sK3_Z SY26)) \<and> | |
| 165 | (\<forall>SY27. | |
| 166 | (\<forall>SY21. SY27 SY21 \<longrightarrow> bnd_sK2_Y SY21) \<and> | |
| 167 | (\<forall>SY15. SY27 SY15 \<longrightarrow> bnd_sK3_Z SY15) \<longrightarrow> | |
| 168 | (\<forall>SY30. SY27 SY30 \<longrightarrow> bnd_sK1_X SY30)) \<longrightarrow> | |
| 169 | bnd_sK1_X = (\<lambda>SY17. bnd_sK2_Y SY17 \<and> bnd_sK3_Z SY17)) = | |
| 170 | False \<Longrightarrow> | |
| 171 | (\<not> (((\<forall>SY25. bnd_sK1_X SY25 \<longrightarrow> bnd_sK2_Y SY25) \<and> | |
| 172 | (\<forall>SY26. bnd_sK1_X SY26 \<longrightarrow> bnd_sK3_Z SY26)) \<and> | |
| 173 | (\<forall>SY27. | |
| 174 | (\<forall>SY21. SY27 SY21 \<longrightarrow> bnd_sK2_Y SY21) \<and> | |
| 175 | (\<forall>SY15. SY27 SY15 \<longrightarrow> bnd_sK3_Z SY15) \<longrightarrow> | |
| 176 | (\<forall>SY30. SY27 SY30 \<longrightarrow> bnd_sK1_X SY30)) \<longrightarrow> | |
| 177 | bnd_sK1_X = (\<lambda>SY17. bnd_sK2_Y SY17 \<and> bnd_sK3_Z SY17))) = | |
| 178 | True" | |
| 179 | apply (tactic {*polarity_switch_tac @{context}*})
 | |
| 180 | done | |
| 181 | *) | |
| 182 | ||
| 183 | (* beware lack of type annotations | |
| 184 | (* lemma "!!x. (A x = B x) = False ==> (B x = A x) = False" *) | |
| 185 | (* lemma "!!x. (A x = B x) = True ==> (B x = A x) = True" *) | |
| 186 | (* lemma "((A x) = (B x)) = True ==> ((B x) = (A x)) = True" *) | |
| 187 | lemma "(A = B) = True ==> (B = A) = True" | |
| 188 | *) | |
| 189 | lemma "!!x. ((A x :: bool) = B x) = False ==> (B x = A x) = False" | |
| 63167 | 190 | apply (tactic \<open>expander_animal @{context} 1\<close>)
 | 
| 55596 | 191 | oops | 
| 192 | ||
| 193 | lemma "(A & B) ==> ~(~A | ~B)" | |
| 63167 | 194 | by (tactic \<open>expander_animal @{context} 1\<close>)
 | 
| 55596 | 195 | |
| 196 | lemma "(A | B) ==> ~(~A & ~B)" | |
| 63167 | 197 | by (tactic \<open>expander_animal @{context} 1\<close>)
 | 
| 55596 | 198 | |
| 199 | lemma "(A | B) | C ==> A | (B | C)" | |
| 63167 | 200 | by (tactic \<open>expander_animal @{context} 1\<close>)
 | 
| 55596 | 201 | |
| 202 | lemma "(~~A) = B ==> A = B" | |
| 63167 | 203 | by (tactic \<open>expander_animal @{context} 1\<close>)
 | 
| 55596 | 204 | |
| 205 | lemma "~ ~ (A = True) ==> A = True" | |
| 63167 | 206 | by (tactic \<open>expander_animal @{context} 1\<close>)
 | 
| 55596 | 207 | |
| 208 | (*This might not be a goal which might realistically arise: | |
| 209 | lemma "((~~A) = B) & (B = (~~A)) ==> ~(~(A = B) | ~(B = A))" *) | |
| 210 | lemma "((~~A) = True) ==> ~(~(A = True) | ~(True = A))" | |
| 63167 | 211 | apply (tactic \<open>expander_animal @{context} 1\<close>)+
 | 
| 55596 | 212 | apply (rule conjI) | 
| 213 | apply assumption | |
| 214 | apply (rule sym, assumption) | |
| 215 | done | |
| 216 | ||
| 217 | lemma "A = B ==> ((~~A) = B) & (B = (~~A)) ==> ~(~(A = B) | ~(B = A))" | |
| 63167 | 218 | by (tactic \<open>expander_animal @{context} 1\<close>)+
 | 
| 55596 | 219 | |
| 220 | (*some lemmas assume constants in the signature of PUZ114^5*) | |
| 221 | consts | |
| 222 | PUZ114_5_bnd_sK1 :: "TPTP_Interpret.ind" | |
| 223 | PUZ114_5_bnd_sK2 :: "TPTP_Interpret.ind" | |
| 224 | PUZ114_5_bnd_sK3 :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool" | |
| 225 | PUZ114_5_bnd_sK4 :: "(TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool) \<Rightarrow> TPTP_Interpret.ind" | |
| 226 | PUZ114_5_bnd_sK5 :: "(TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool) \<Rightarrow> TPTP_Interpret.ind" | |
| 227 | PUZ114_5_bnd_s :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind" | |
| 228 | PUZ114_5_bnd_c1 :: TPTP_Interpret.ind | |
| 229 | ||
| 230 | (*testing logical expansion*) | |
| 231 | lemma "!! SY30. (SY30 PUZ114_5_bnd_c1 PUZ114_5_bnd_c1 \<and> | |
| 232 | (\<forall>Xj Xk. | |
| 233 | SY30 Xj Xk \<longrightarrow> | |
| 234 | SY30 (PUZ114_5_bnd_s (PUZ114_5_bnd_s Xj)) Xk \<and> | |
| 235 | SY30 (PUZ114_5_bnd_s Xj) (PUZ114_5_bnd_s Xk)) \<longrightarrow> | |
| 236 | SY30 PUZ114_5_bnd_sK1 PUZ114_5_bnd_sK2) | |
| 237 | ==> ( | |
| 238 | (~ SY30 PUZ114_5_bnd_c1 PUZ114_5_bnd_c1) | |
| 239 | | (~ (\<forall>Xj Xk. | |
| 240 | SY30 Xj Xk \<longrightarrow> | |
| 241 | SY30 (PUZ114_5_bnd_s (PUZ114_5_bnd_s Xj)) Xk \<and> | |
| 242 | SY30 (PUZ114_5_bnd_s Xj) (PUZ114_5_bnd_s Xk))) | |
| 243 | | SY30 PUZ114_5_bnd_sK1 PUZ114_5_bnd_sK2 | |
| 244 | )" | |
| 63167 | 245 | by (tactic \<open>expander_animal @{context} 1\<close>)+
 | 
| 55596 | 246 | |
| 247 | (* | |
| 248 | extcnf_forall_pos: | |
| 249 | ||
| 250 | (! X. L1) | ... | Ln | |
| 251 | ---------------------------- X' fresh | |
| 252 | ! X'. (L1[X'/X] | ... | Ln) | |
| 253 | ||
| 254 | After elimination rule has been applied we'll have a subgoal which looks like this: | |
| 255 | (! X. L1) | |
| 256 | ---------------------------- X' fresh | |
| 257 | ! X'. (L1[X'/X] | ... | Ln) | |
| 258 | and we need to transform it so that, in Isabelle, we go from | |
| 259 | (! X. L1) ==> ! X'. (L1[X'/X] | ... | Ln) | |
| 260 | to | |
| 261 | \<And> X'. L1[X'/X] ==> (L1[X'/X] | ... | Ln) | |
| 262 | (where X' is fresh, or renamings are done suitably).*) | |
| 263 | ||
| 264 | lemma "A | B \<Longrightarrow> A | B | C" | |
| 63167 | 265 | apply (tactic \<open>flip_conclusion_tac @{context} 1\<close>)+
 | 
| 266 | apply (tactic \<open>break_hypotheses 1\<close>)+ | |
| 55596 | 267 | oops | 
| 268 | ||
| 269 | consts | |
| 270 | CSR122_1_bnd_lBill_THFTYPE_i :: TPTP_Interpret.ind | |
| 271 | CSR122_1_bnd_lMary_THFTYPE_i :: TPTP_Interpret.ind | |
| 272 | CSR122_1_bnd_lSue_THFTYPE_i :: TPTP_Interpret.ind | |
| 273 | CSR122_1_bnd_n2009_THFTYPE_i :: TPTP_Interpret.ind | |
| 274 | CSR122_1_bnd_lYearFn_THFTYPE_IiiI :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind" | |
| 275 | CSR122_1_bnd_holdsDuring_THFTYPE_IiooI :: | |
| 276 | "TPTP_Interpret.ind \<Rightarrow> bool \<Rightarrow> bool" | |
| 277 | CSR122_1_bnd_likes_THFTYPE_IiioI :: | |
| 278 | "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool" | |
| 279 | ||
| 280 | lemma "\<forall>SV2. (CSR122_1_bnd_holdsDuring_THFTYPE_IiooI | |
| 281 | (CSR122_1_bnd_lYearFn_THFTYPE_IiiI CSR122_1_bnd_n2009_THFTYPE_i) | |
| 282 | (\<not> (\<not> CSR122_1_bnd_likes_THFTYPE_IiioI | |
| 283 | CSR122_1_bnd_lMary_THFTYPE_i | |
| 284 | CSR122_1_bnd_lBill_THFTYPE_i \<or> | |
| 285 | \<not> CSR122_1_bnd_likes_THFTYPE_IiioI | |
| 286 | CSR122_1_bnd_lSue_THFTYPE_i | |
| 287 | CSR122_1_bnd_lBill_THFTYPE_i)) = | |
| 288 | CSR122_1_bnd_holdsDuring_THFTYPE_IiooI SV2 True) = | |
| 289 | False \<Longrightarrow> | |
| 290 | \<forall>SV2. (CSR122_1_bnd_lYearFn_THFTYPE_IiiI CSR122_1_bnd_n2009_THFTYPE_i = | |
| 291 | SV2) = | |
| 292 | False \<or> | |
| 293 | ((\<not> (\<not> CSR122_1_bnd_likes_THFTYPE_IiioI | |
| 294 | CSR122_1_bnd_lMary_THFTYPE_i CSR122_1_bnd_lBill_THFTYPE_i \<or> | |
| 295 | \<not> CSR122_1_bnd_likes_THFTYPE_IiioI CSR122_1_bnd_lSue_THFTYPE_i | |
| 296 | CSR122_1_bnd_lBill_THFTYPE_i)) = | |
| 297 | True) = | |
| 298 | False" | |
| 299 | apply (rule allI, erule_tac x = "SV2" in allE) | |
| 63167 | 300 | apply (tactic \<open>extuni_dec_tac @{context} 1\<close>)
 | 
| 55596 | 301 | done | 
| 302 | ||
| 303 | (*SEU882^5*) | |
| 304 | (* | |
| 305 | lemma | |
| 61076 | 306 | "\<forall>(SV2::TPTP_Interpret.ind) | 
| 307 | SV1::TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind. | |
| 55596 | 308 | (SV1 SV2 = bnd_sK1_Xy) = | 
| 309 | False | |
| 310 | \<Longrightarrow> | |
| 61076 | 311 | \<forall>SV2::TPTP_Interpret.ind. | 
| 55596 | 312 | (bnd_sK1_Xy = bnd_sK1_Xy) = | 
| 313 | False" | |
| 314 | ML_prf {*
 | |
| 315 | open TPTP_Syntax; | |
| 316 | open TPTP_Proof; | |
| 317 | ||
| 318 | val binds = | |
| 319 | [Bind ("SV1", Quant (Lambda, [("SX0", SOME (Fmla_type (Atom (THF_Atom_term (Term_Func (TypeSymbol Type_Ind, []))))))], Atom (THF_Atom_term (Term_Func (Uninterpreted "sK1_Xy", [])))))]
 | |
| 320 | ||
| 321 | val tec = bind_tac @{context} (hd prob_names) binds
 | |
| 322 | *} | |
| 323 | (* | |
| 324 | apply (tactic {*strip_qtfrs
 | |
| 325 | (* THEN tec *)*}) | |
| 326 | *) | |
| 327 | apply (tactic {*tec*})
 | |
| 328 | done | |
| 329 | *) | |
| 330 | ||
| 331 | lemma "A | B \<Longrightarrow> C1 | A | C2 | B | C3" | |
| 332 | apply (erule disjE) | |
| 63167 | 333 | apply (tactic \<open>clause_breaker 1\<close>) | 
| 334 | apply (tactic \<open>clause_breaker 1\<close>) | |
| 55596 | 335 | done | 
| 336 | ||
| 337 | lemma "A \<Longrightarrow> A" | |
| 63167 | 338 | apply (tactic \<open>clause_breaker 1\<close>) | 
| 55596 | 339 | done | 
| 340 | ||
| 341 | typedecl NUM667_1_bnd_nat | |
| 342 | consts | |
| 343 | NUM667_1_bnd_less :: "NUM667_1_bnd_nat \<Rightarrow> NUM667_1_bnd_nat \<Rightarrow> bool" | |
| 344 | NUM667_1_bnd_x :: NUM667_1_bnd_nat | |
| 345 | NUM667_1_bnd_y :: NUM667_1_bnd_nat | |
| 346 | ||
| 347 | (*NUM667^1 node 302 -- dec*) | |
| 348 | lemma "\<forall>SV12 SV13 SV14 SV9 SV10 SV11. | |
| 349 | ((((NUM667_1_bnd_less SV12 SV13 = NUM667_1_bnd_less SV11 SV10) = False \<or> | |
| 350 | (SV14 = SV13) = False) \<or> | |
| 351 | NUM667_1_bnd_less SV12 SV14 = False) \<or> | |
| 352 | NUM667_1_bnd_less SV9 SV10 = True) \<or> | |
| 353 | (SV9 = SV11) = | |
| 354 | False \<Longrightarrow> | |
| 355 | \<forall>SV9 SV14 SV10 SV13 SV11 SV12. | |
| 356 | (((((SV12 = SV11) = False \<or> (SV13 = SV10) = False) \<or> | |
| 357 | (SV14 = SV13) = False) \<or> | |
| 358 | NUM667_1_bnd_less SV12 SV14 = False) \<or> | |
| 359 | NUM667_1_bnd_less SV9 SV10 = True) \<or> | |
| 360 | (SV9 = SV11) = | |
| 361 | False" | |
| 63167 | 362 | apply (tactic \<open>strip_qtfrs_tac @{context}\<close>)
 | 
| 363 | apply (tactic \<open>break_hypotheses 1\<close>) | |
| 364 | apply (tactic \<open>ALLGOALS (TRY o clause_breaker)\<close>) | |
| 365 | apply (tactic \<open>extuni_dec_tac @{context} 1\<close>)
 | |
| 55596 | 366 | done | 
| 367 | ||
| 63167 | 368 | ML \<open> | 
| 55596 | 369 | extuni_dec_n @{context} 2;
 | 
| 63167 | 370 | \<close> | 
| 55596 | 371 | |
| 372 | (*NUM667^1, node 202*) | |
| 373 | lemma "\<forall>SV9 SV10 SV11. | |
| 374 | ((((SV9 = SV11) = (NUM667_1_bnd_x = NUM667_1_bnd_y)) = False \<or> | |
| 375 | NUM667_1_bnd_less SV11 SV10 = False) \<or> | |
| 376 | NUM667_1_bnd_less SV9 SV10 = True) \<or> | |
| 377 | NUM667_1_bnd_less NUM667_1_bnd_x NUM667_1_bnd_y = | |
| 378 | True \<Longrightarrow> | |
| 379 | \<forall>SV10 SV9 SV11. | |
| 380 | ((((SV11 = NUM667_1_bnd_x) = False \<or> (SV9 = NUM667_1_bnd_y) = False) \<or> | |
| 381 | NUM667_1_bnd_less SV11 SV10 = False) \<or> | |
| 382 | NUM667_1_bnd_less SV9 SV10 = True) \<or> | |
| 383 | NUM667_1_bnd_less NUM667_1_bnd_x NUM667_1_bnd_y = | |
| 384 | True" | |
| 63167 | 385 | apply (tactic \<open>strip_qtfrs_tac @{context}\<close>)
 | 
| 386 | apply (tactic \<open>break_hypotheses 1\<close>) | |
| 387 | apply (tactic \<open>ALLGOALS (TRY o clause_breaker)\<close>) | |
| 388 | apply (tactic \<open>extuni_dec_tac @{context} 1\<close>)
 | |
| 55596 | 389 | done | 
| 390 | ||
| 391 | (*NUM667^1 node 141*) | |
| 392 | (* | |
| 393 | lemma "((bnd_x = bnd_x) = False \<or> (bnd_y = bnd_z) = False) \<or> | |
| 394 | bnd_less bnd_x bnd_y = True \<Longrightarrow> | |
| 395 | (bnd_y = bnd_z) = False \<or> bnd_less bnd_x bnd_y = True" | |
| 396 | apply (tactic {*strip_qtfrs*})
 | |
| 397 | apply (tactic {*break_hypotheses 1*})
 | |
| 398 | apply (tactic {*ALLGOALS (TRY o clause_breaker)*})
 | |
| 399 | apply (erule extuni_triv) | |
| 400 | done | |
| 401 | *) | |
| 402 | ||
| 63167 | 403 | ML \<open> | 
| 55596 | 404 | fun full_extcnf_combined_tac ctxt = | 
| 405 | extcnf_combined_tac ctxt NONE | |
| 406 | [ConstsDiff, | |
| 407 | StripQuantifiers, | |
| 408 | Flip_Conclusion, | |
| 409 | Loop [ | |
| 410 | Close_Branch, | |
| 411 | ConjI, | |
| 412 | King_Cong, | |
| 413 | Break_Hypotheses, | |
| 414 | Existential_Free, | |
| 415 | Existential_Var, | |
| 416 | Universal, | |
| 417 | RemoveRedundantQuantifications], | |
| 418 | CleanUp [RemoveHypothesesFromSkolemDefs, RemoveDuplicates], | |
| 419 | AbsorbSkolemDefs] | |
| 420 | [] | |
| 63167 | 421 | \<close> | 
| 55596 | 422 | |
| 63167 | 423 | ML \<open> | 
| 55596 | 424 | fun nonfull_extcnf_combined_tac ctxt feats = | 
| 425 | extcnf_combined_tac ctxt NONE | |
| 426 | [ConstsDiff, | |
| 427 | StripQuantifiers, | |
| 428 | InnerLoopOnce (Break_Hypotheses :: feats), | |
| 429 | AbsorbSkolemDefs] | |
| 430 | [] | |
| 63167 | 431 | \<close> | 
| 55596 | 432 | |
| 433 | consts SEU882_5_bnd_sK1_Xy :: TPTP_Interpret.ind | |
| 434 | lemma | |
| 435 | "\<forall>SV2. (SEU882_5_bnd_sK1_Xy = SEU882_5_bnd_sK1_Xy) = False \<Longrightarrow> | |
| 436 | (SEU882_5_bnd_sK1_Xy = SEU882_5_bnd_sK1_Xy) = False" | |
| 437 | (* apply (erule_tac x = "(@X. False)" in allE) *) | |
| 438 | (* apply (tactic {*remove_redundant_quantification 1*}) *)
 | |
| 439 | (* apply assumption *) | |
| 63167 | 440 | by (tactic \<open>nonfull_extcnf_combined_tac @{context} [RemoveRedundantQuantifications, Extuni_FlexRigid]\<close>)
 | 
| 55596 | 441 | |
| 442 | (*NUM667^1*) | |
| 443 | (* | |
| 444 |      (* (Annotated_step ("153", "extuni_triv"), *)
 | |
| 445 | lemma "((bnd_y = bnd_x) = False \<or> (bnd_z = bnd_z) = False) \<or> | |
| 446 | (bnd_y = bnd_z) = True \<Longrightarrow> | |
| 447 | (bnd_y = bnd_x) = False \<or> (bnd_y = bnd_z) = True" | |
| 448 | apply (tactic {*nonfull_extcnf_combined_tac [Extuni_Triv]*})
 | |
| 449 | done | |
| 450 |      (* (Annotated_step ("162", "extuni_triv"), *)
 | |
| 451 | lemma "((bnd_y = bnd_x) = False \<or> (bnd_z = bnd_z) = False) \<or> | |
| 452 | bnd_less bnd_y bnd_z = True \<Longrightarrow> | |
| 453 | (bnd_y = bnd_x) = False \<or> bnd_less bnd_y bnd_z = True" | |
| 454 | apply (tactic {*nonfull_extcnf_combined_tac [Extuni_Triv]*})
 | |
| 455 | done | |
| 456 | *) | |
| 457 | ||
| 458 | (* SEU602^2 *) | |
| 459 | consts | |
| 460 | SEU602_2_bnd_sK7_E :: "(TPTP_Interpret.ind \<Rightarrow> bool) \<Rightarrow> TPTP_Interpret.ind" | |
| 461 | SEU602_2_bnd_sK2_SY17 :: TPTP_Interpret.ind | |
| 462 | SEU602_2_bnd_in :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool" | |
| 463 | ||
| 464 |      (* (Annotated_step ("113", "extuni_func"), *)
 | |
| 61076 | 465 | lemma "\<forall>SV49::TPTP_Interpret.ind \<Rightarrow> bool. | 
| 55596 | 466 | (SV49 = | 
| 61076 | 467 | (\<lambda>SY23::TPTP_Interpret.ind. | 
| 55596 | 468 | \<not> SEU602_2_bnd_in SY23 SEU602_2_bnd_sK2_SY17)) = | 
| 469 | False \<Longrightarrow> | |
| 61076 | 470 | \<forall>SV49::TPTP_Interpret.ind \<Rightarrow> bool. | 
| 55596 | 471 | (SV49 (SEU602_2_bnd_sK7_E SV49) = | 
| 472 | (\<not> SEU602_2_bnd_in (SEU602_2_bnd_sK7_E SV49) SEU602_2_bnd_sK2_SY17)) = | |
| 473 | False" | |
| 474 | (*FIXME this (and similar) tests are getting the "Bad background theory of goal state" error since upgrading to Isabelle2013-2.*) | |
| 63167 | 475 | by (tactic \<open>nonfull_extcnf_combined_tac @{context} [Extuni_Func, Existential_Var]\<close>)
 | 
| 55596 | 476 | |
| 477 | consts | |
| 478 | SEV405_5_bnd_sK1_SY2 :: "(TPTP_Interpret.ind \<Rightarrow> bool) \<Rightarrow> TPTP_Interpret.ind" | |
| 479 | SEV405_5_bnd_cA :: bool | |
| 480 | ||
| 61076 | 481 | lemma "\<forall>SV1::TPTP_Interpret.ind \<Rightarrow> bool. | 
| 482 | (\<forall>SY2::TPTP_Interpret.ind. | |
| 55596 | 483 | \<not> (\<not> (\<not> SV1 SY2 \<or> SEV405_5_bnd_cA) \<or> | 
| 484 | \<not> (\<not> SEV405_5_bnd_cA \<or> SV1 SY2))) = | |
| 485 | False \<Longrightarrow> | |
| 61076 | 486 | \<forall>SV1::TPTP_Interpret.ind \<Rightarrow> bool. | 
| 55596 | 487 | (\<not> (\<not> (\<not> SV1 (SEV405_5_bnd_sK1_SY2 SV1) \<or> SEV405_5_bnd_cA) \<or> | 
| 488 | \<not> (\<not> SEV405_5_bnd_cA \<or> SV1 (SEV405_5_bnd_sK1_SY2 SV1)))) = | |
| 489 | False" | |
| 63167 | 490 | by (tactic \<open>nonfull_extcnf_combined_tac @{context} [Existential_Var]\<close>)
 | 
| 55596 | 491 | (* | 
| 492 | strip quantifiers -- creating a space of permutations; from shallowest to deepest (iterative deepening) | |
| 493 | flip the conclusion -- giving us branch | |
| 494 | apply some collection of rules, in some order, until the space has been explored completely. advantage of not having extcnf_combined: search space is shallow -- particularly if the collection of rules is small. | |
| 495 | *) | |
| 496 | ||
| 497 | consts | |
| 498 | SEU581_2_bnd_sK1 :: "TPTP_Interpret.ind" | |
| 499 | SEU581_2_bnd_sK2 :: "TPTP_Interpret.ind \<Rightarrow> bool" | |
| 500 | SEU581_2_bnd_subset :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> HOL.bool" | |
| 501 | SEU581_2_bnd_dsetconstr :: "TPTP_Interpret.ind \<Rightarrow> (TPTP_Interpret.ind \<Rightarrow> HOL.bool) \<Rightarrow> TPTP_Interpret.ind" | |
| 502 | ||
| 503 | (*testing parameters*) | |
| 504 | lemma "! X :: TPTP_Interpret.ind . (\<forall>A B. SEU581_2_bnd_in B (SEU581_2_bnd_powerset A) \<longrightarrow> SEU581_2_bnd_subset B A) = True | |
| 505 | \<Longrightarrow> ! X :: TPTP_Interpret.ind . (\<forall>A B. \<not> SEU581_2_bnd_in B (SEU581_2_bnd_powerset A) \<or> SEU581_2_bnd_subset B A) = True" | |
| 63167 | 506 | by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
 | 
| 55596 | 507 | |
| 508 | lemma "(A & B) = True ==> (A | B) = True" | |
| 63167 | 509 | by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
 | 
| 55596 | 510 | |
| 511 | lemma "(\<not> bnd_subset (bnd_dsetconstr bnd_sK1 bnd_sK2) bnd_sK1) = True \<Longrightarrow> (bnd_subset (bnd_dsetconstr bnd_sK1 bnd_sK2) bnd_sK1) = False" | |
| 63167 | 512 | by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
 | 
| 55596 | 513 | |
| 514 | (*testing goals with parameters*) | |
| 515 | lemma "(\<not> bnd_subset (bnd_dsetconstr bnd_sK1 bnd_sK2) bnd_sK1) = True \<Longrightarrow> ! X. (bnd_subset (bnd_dsetconstr bnd_sK1 bnd_sK2) bnd_sK1) = False" | |
| 63167 | 516 | by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
 | 
| 55596 | 517 | |
| 518 | lemma "(A & B) = True ==> (B & A) = True" | |
| 63167 | 519 | by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
 | 
| 55596 | 520 | |
| 521 | (*appreciating differences between THEN, REPEAT, and APPEND*) | |
| 522 | lemma "A & B ==> B & A" | |
| 63167 | 523 | apply (tactic \<open> | 
| 55596 | 524 |   TRY (etac @{thm conjE} 1)
 | 
| 63167 | 525 |   THEN TRY (rtac @{thm conjI} 1)\<close>)
 | 
| 55596 | 526 | by assumption+ | 
| 527 | ||
| 528 | lemma "A & B ==> B & A" | |
| 63167 | 529 | by (tactic \<open> | 
| 55596 | 530 |   etac @{thm conjE} 1
 | 
| 531 |   THEN rtac @{thm conjI} 1
 | |
| 63167 | 532 | THEN REPEAT (atac 1)\<close>) | 
| 55596 | 533 | |
| 534 | lemma "A & B ==> B & A" | |
| 63167 | 535 | apply (tactic \<open> | 
| 55596 | 536 |   rtac @{thm conjI} 1
 | 
| 63167 | 537 |   APPEND etac @{thm conjE} 1\<close>)+
 | 
| 55596 | 538 | back | 
| 539 | by assumption+ | |
| 540 | ||
| 541 | consts | |
| 542 | SEU581_2_bnd_sK3 :: "TPTP_Interpret.ind" | |
| 543 | SEU581_2_bnd_sK4 :: "TPTP_Interpret.ind" | |
| 544 | SEU581_2_bnd_sK5 :: "(TPTP_Interpret.ind \<Rightarrow> bool) \<Rightarrow> TPTP_Interpret.ind" | |
| 545 | SEU581_2_bnd_powerset :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind" | |
| 546 | SEU581_2_bnd_in :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool" | |
| 547 | ||
| 548 | consts | |
| 549 | bnd_c1 :: TPTP_Interpret.ind | |
| 550 | bnd_s :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind" | |
| 551 | ||
| 552 | lemma "(\<forall>SX0. (\<not> (\<not> SX0 (PUZ114_5_bnd_sK4 SX0) (PUZ114_5_bnd_sK5 SX0) \<or> | |
| 553 | \<not> (\<not> SX0 (bnd_s (bnd_s (PUZ114_5_bnd_sK4 SX0))) | |
| 554 | (PUZ114_5_bnd_sK5 SX0) \<or> | |
| 555 | \<not> SX0 (bnd_s (PUZ114_5_bnd_sK4 SX0)) | |
| 556 | (bnd_s (PUZ114_5_bnd_sK5 SX0)))) \<or> | |
| 557 | \<not> SX0 bnd_c1 bnd_c1) \<or> | |
| 558 | SX0 PUZ114_5_bnd_sK1 PUZ114_5_bnd_sK2) = | |
| 559 | True ==> \<forall>SV1. ((\<not> (\<not> SV1 (PUZ114_5_bnd_sK4 SV1) (PUZ114_5_bnd_sK5 SV1) \<or> | |
| 560 | \<not> (\<not> SV1 (bnd_s (bnd_s (PUZ114_5_bnd_sK4 SV1))) | |
| 561 | (PUZ114_5_bnd_sK5 SV1) \<or> | |
| 562 | \<not> SV1 (bnd_s (PUZ114_5_bnd_sK4 SV1)) | |
| 563 | (bnd_s (PUZ114_5_bnd_sK5 SV1)))) \<or> | |
| 564 | \<not> SV1 bnd_c1 bnd_c1) \<or> | |
| 565 | SV1 PUZ114_5_bnd_sK1 PUZ114_5_bnd_sK2) = | |
| 566 | True" | |
| 63167 | 567 | by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
 | 
| 55596 | 568 | |
| 569 | lemma "(\<not> SEU581_2_bnd_subset (SEU581_2_bnd_dsetconstr SEU581_2_bnd_sK1 SEU581_2_bnd_sK2) SEU581_2_bnd_sK1) = True \<Longrightarrow> (SEU581_2_bnd_subset (SEU581_2_bnd_dsetconstr SEU581_2_bnd_sK1 SEU581_2_bnd_sK2) SEU581_2_bnd_sK1) = False" | |
| 63167 | 570 | by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
 | 
| 55596 | 571 | |
| 572 | (*testing repeated application of simulator*) | |
| 573 | lemma "(\<not> \<not> False) = True \<Longrightarrow> | |
| 574 | SEU581_2_bnd_subset (SEU581_2_bnd_dsetconstr SEU581_2_bnd_sK1 SEU581_2_bnd_sK2) SEU581_2_bnd_sK1 = True \<or> | |
| 575 | False = True \<or> False = True \<or> False = True" | |
| 63167 | 576 | by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
 | 
| 55596 | 577 | |
| 578 | (*Testing non-normal conclusion. Ideally we should be able to apply | |
| 579 | the tactic to arbitrary chains of extcnf steps -- where it's not | |
| 580 | generally the case that the conclusions are normal.*) | |
| 581 | lemma "(\<not> \<not> False) = True \<Longrightarrow> | |
| 582 | SEU581_2_bnd_subset (SEU581_2_bnd_dsetconstr SEU581_2_bnd_sK1 SEU581_2_bnd_sK2) SEU581_2_bnd_sK1 = True \<or> | |
| 583 | (\<not> False) = False" | |
| 63167 | 584 | by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
 | 
| 55596 | 585 | |
| 586 | (*testing repeated application of simulator, involving different extcnf rules*) | |
| 587 | lemma "(\<not> \<not> (False | False)) = True \<Longrightarrow> | |
| 588 | SEU581_2_bnd_subset (SEU581_2_bnd_dsetconstr SEU581_2_bnd_sK1 SEU581_2_bnd_sK2) SEU581_2_bnd_sK1 = True \<or> | |
| 589 | False = True \<or> False = True \<or> False = True" | |
| 63167 | 590 | by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
 | 
| 55596 | 591 | |
| 592 | (*testing logical expansion*) | |
| 593 | lemma "(\<forall>A B. SEU581_2_bnd_in B (SEU581_2_bnd_powerset A) \<longrightarrow> SEU581_2_bnd_subset B A) = True | |
| 594 | \<Longrightarrow> (\<forall>A B. \<not> SEU581_2_bnd_in B (SEU581_2_bnd_powerset A) \<or> SEU581_2_bnd_subset B A) = True" | |
| 63167 | 595 | by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
 | 
| 55596 | 596 | |
| 597 | (*testing extcnf_forall_pos*) | |
| 598 | lemma "(\<forall>A Xphi. SEU581_2_bnd_in (SEU581_2_bnd_dsetconstr A Xphi) (SEU581_2_bnd_powerset A)) = True \<Longrightarrow> \<forall>SV1. (\<forall>SY14. | |
| 599 | SEU581_2_bnd_in (SEU581_2_bnd_dsetconstr SV1 SY14) | |
| 600 | (SEU581_2_bnd_powerset SV1)) = True" | |
| 63167 | 601 | by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
 | 
| 55596 | 602 | |
| 603 | lemma "((\<forall>A Xphi. SEU581_2_bnd_in (SEU581_2_bnd_dsetconstr A Xphi) (SEU581_2_bnd_powerset A)) = True) | ((~ False) = False) \<Longrightarrow> | |
| 604 | \<forall>SV1. ((\<forall>SY14. SEU581_2_bnd_in (SEU581_2_bnd_dsetconstr SV1 SY14) (SEU581_2_bnd_powerset SV1)) = True) | ((~ False) = False)" | |
| 63167 | 605 | by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
 | 
| 55596 | 606 | |
| 607 | (*testing parameters*) | |
| 608 | lemma "(\<forall>A B. SEU581_2_bnd_in B (SEU581_2_bnd_powerset A) \<longrightarrow> SEU581_2_bnd_subset B A) = True | |
| 609 | \<Longrightarrow> ! X. (\<forall>A B. \<not> SEU581_2_bnd_in B (SEU581_2_bnd_powerset A) \<or> SEU581_2_bnd_subset B A) = True" | |
| 63167 | 610 | by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
 | 
| 55596 | 611 | |
| 612 | lemma "((? A .P1 A) = False) | P2 = True \<Longrightarrow> !X. ((P1 X) = False | P2 = True)" | |
| 63167 | 613 | by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
 | 
| 55596 | 614 | |
| 615 | lemma "((!A . (P1a A | P1b A)) = True) | (P2 = True) \<Longrightarrow> !X. (P1a X = True | P1b X = True | P2 = True)" | |
| 63167 | 616 | by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
 | 
| 55596 | 617 | |
| 618 | lemma "! Y. (((!A .(P1a A | P1b A)) = True) | P2 = True) \<Longrightarrow> ! Y. (!X. (P1a X = True | P1b X = True | P2 = True))" | |
| 63167 | 619 | by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
 | 
| 55596 | 620 | |
| 621 | lemma "! Y. (((!A .(P1a A | P1b A)) = True) | P2 = True) \<Longrightarrow> ! Y. (!X. (P1a X = True | P1b X = True | P2 = True))" | |
| 63167 | 622 | by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
 | 
| 55596 | 623 | |
| 624 | lemma "! Y. (((!A .(P1a A | P1b A)) = True) | P2 = True) \<Longrightarrow> ! Y. (!X. (P1a X = True | P1b X = True | P2 = True))" | |
| 63167 | 625 | by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
 | 
| 55596 | 626 | |
| 627 | consts dud_bnd_s :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind" | |
| 628 | ||
| 629 | (*this lemma kills blast*) | |
| 630 | lemma "(\<not> (\<forall>SX0 SX1. | |
| 631 | \<not> PUZ114_5_bnd_sK3 SX0 SX1 \<or> PUZ114_5_bnd_sK3 (dud_bnd_s (dud_bnd_s SX0)) SX1) \<or> | |
| 632 | \<not> (\<forall>SX0 SX1. | |
| 633 | \<not> PUZ114_5_bnd_sK3 SX0 SX1 \<or> | |
| 634 | PUZ114_5_bnd_sK3 (dud_bnd_s SX0) (dud_bnd_s SX1))) = | |
| 635 | False \<Longrightarrow> (\<not> (\<forall>SX0 SX1. | |
| 636 | \<not> PUZ114_5_bnd_sK3 SX0 SX1 \<or> | |
| 637 | PUZ114_5_bnd_sK3 (dud_bnd_s SX0) (dud_bnd_s SX1))) = | |
| 638 | False" | |
| 63167 | 639 | by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
 | 
| 55596 | 640 | |
| 641 | (*testing logical expansion -- this should be done by blast*) | |
| 642 | lemma "(\<forall>A B. bnd_in B (bnd_powerset A) \<longrightarrow> SEU581_2_bnd_subset B A) = True | |
| 643 | \<Longrightarrow> (\<forall>A B. \<not> bnd_in B (bnd_powerset A) \<or> SEU581_2_bnd_subset B A) = True" | |
| 63167 | 644 | by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
 | 
| 55596 | 645 | |
| 646 | (*testing related to PUZ114^5.p.out*) | |
| 647 | lemma "\<forall>SV1. ((\<not> (\<not> SV1 (PUZ114_5_bnd_sK4 SV1) (PUZ114_5_bnd_sK5 SV1) \<or> | |
| 648 | \<not> (\<not> SV1 (bnd_s (bnd_s (PUZ114_5_bnd_sK4 SV1))) | |
| 649 | (PUZ114_5_bnd_sK5 SV1) \<or> | |
| 650 | \<not> SV1 (bnd_s (PUZ114_5_bnd_sK4 SV1)) | |
| 651 | (bnd_s (PUZ114_5_bnd_sK5 SV1))))) = | |
| 652 | True \<or> | |
| 653 | (\<not> SV1 bnd_c1 bnd_c1) = True) \<or> | |
| 654 | SV1 PUZ114_5_bnd_sK1 PUZ114_5_bnd_sK2 = True \<Longrightarrow> | |
| 655 | \<forall>SV1. (SV1 bnd_c1 bnd_c1 = False \<or> | |
| 656 | (\<not> (\<not> SV1 (PUZ114_5_bnd_sK4 SV1) (PUZ114_5_bnd_sK5 SV1) \<or> | |
| 657 | \<not> (\<not> SV1 (bnd_s (bnd_s (PUZ114_5_bnd_sK4 SV1))) | |
| 658 | (PUZ114_5_bnd_sK5 SV1) \<or> | |
| 659 | \<not> SV1 (bnd_s (PUZ114_5_bnd_sK4 SV1)) | |
| 660 | (bnd_s (PUZ114_5_bnd_sK5 SV1))))) = | |
| 661 | True) \<or> | |
| 662 | SV1 PUZ114_5_bnd_sK1 PUZ114_5_bnd_sK2 = True" | |
| 63167 | 663 | by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
 | 
| 55596 | 664 | |
| 665 | lemma "\<forall>SV2. (\<forall>SY41. | |
| 666 | \<not> PUZ114_5_bnd_sK3 SV2 SY41 \<or> | |
| 667 | PUZ114_5_bnd_sK3 (dud_bnd_s (dud_bnd_s SV2)) SY41) = | |
| 668 | True \<Longrightarrow> | |
| 669 | \<forall>SV4 SV2. | |
| 670 | (\<not> PUZ114_5_bnd_sK3 SV2 SV4 \<or> | |
| 671 | PUZ114_5_bnd_sK3 (dud_bnd_s (dud_bnd_s SV2)) SV4) = | |
| 672 | True" | |
| 63167 | 673 | by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
 | 
| 55596 | 674 | |
| 675 | lemma "\<forall>SV3. (\<forall>SY42. | |
| 676 | \<not> PUZ114_5_bnd_sK3 SV3 SY42 \<or> | |
| 677 | PUZ114_5_bnd_sK3 (dud_bnd_s SV3) (dud_bnd_s SY42)) = | |
| 678 | True \<Longrightarrow> | |
| 679 | \<forall>SV5 SV3. | |
| 680 | (\<not> PUZ114_5_bnd_sK3 SV3 SV5 \<or> | |
| 681 | PUZ114_5_bnd_sK3 (dud_bnd_s SV3) (dud_bnd_s SV5)) = | |
| 682 | True" | |
| 63167 | 683 | by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
 | 
| 55596 | 684 | |
| 685 | ||
| 686 | subsection "unfold_def" | |
| 687 |      (* (Annotated_step ("9", "unfold_def"), *)
 | |
| 688 | lemma "bnd_kpairiskpair = | |
| 689 | (ALL Xx Xy. | |
| 690 | bnd_iskpair | |
| 691 | (bnd_setadjoin (bnd_setadjoin Xx bnd_emptyset) | |
| 692 | (bnd_setadjoin (bnd_setadjoin Xx (bnd_setadjoin Xy bnd_emptyset)) | |
| 693 | bnd_emptyset))) & | |
| 694 | bnd_kpair = | |
| 695 | (%Xx Xy. | |
| 696 | bnd_setadjoin (bnd_setadjoin Xx bnd_emptyset) | |
| 697 | (bnd_setadjoin (bnd_setadjoin Xx (bnd_setadjoin Xy bnd_emptyset)) | |
| 698 | bnd_emptyset)) & | |
| 699 | bnd_iskpair = | |
| 700 | (%A. EX Xx. bnd_in Xx (bnd_setunion A) & | |
| 701 | (EX Xy. bnd_in Xy (bnd_setunion A) & | |
| 702 | A = bnd_setadjoin (bnd_setadjoin Xx bnd_emptyset) | |
| 703 | (bnd_setadjoin | |
| 704 | (bnd_setadjoin Xx | |
| 705 | (bnd_setadjoin Xy bnd_emptyset)) | |
| 706 | bnd_emptyset))) & | |
| 707 | (~ (ALL SY0 SY1. | |
| 708 | EX SY3. | |
| 709 | bnd_in SY3 | |
| 710 | (bnd_setunion | |
| 711 | (bnd_setadjoin (bnd_setadjoin SY0 bnd_emptyset) | |
| 712 | (bnd_setadjoin | |
| 713 | (bnd_setadjoin SY0 (bnd_setadjoin SY1 bnd_emptyset)) | |
| 714 | bnd_emptyset))) & | |
| 715 | (EX SY4. | |
| 716 | bnd_in SY4 | |
| 717 | (bnd_setunion | |
| 718 | (bnd_setadjoin (bnd_setadjoin SY0 bnd_emptyset) | |
| 719 | (bnd_setadjoin | |
| 720 | (bnd_setadjoin SY0 | |
| 721 | (bnd_setadjoin SY1 bnd_emptyset)) | |
| 722 | bnd_emptyset))) & | |
| 723 | bnd_setadjoin (bnd_setadjoin SY0 bnd_emptyset) | |
| 724 | (bnd_setadjoin | |
| 725 | (bnd_setadjoin SY0 (bnd_setadjoin SY1 bnd_emptyset)) | |
| 726 | bnd_emptyset) = | |
| 727 | bnd_setadjoin (bnd_setadjoin SY3 bnd_emptyset) | |
| 728 | (bnd_setadjoin | |
| 729 | (bnd_setadjoin SY3 (bnd_setadjoin SY4 bnd_emptyset)) | |
| 730 | bnd_emptyset)))) = | |
| 731 | True | |
| 732 | ==> (~ (ALL SX0 SX1. | |
| 733 | ~ (ALL SX2. | |
| 734 | ~ ~ (~ bnd_in SX2 | |
| 735 | (bnd_setunion | |
| 736 | (bnd_setadjoin | |
| 737 | (bnd_setadjoin SX0 bnd_emptyset) | |
| 738 | (bnd_setadjoin (bnd_setadjoin SX0 (bnd_setadjoin SX1 bnd_emptyset)) bnd_emptyset))) | | |
| 739 | ~ ~ (ALL SX3. | |
| 740 | ~ ~ (~ bnd_in SX3 | |
| 741 | (bnd_setunion | |
| 742 | (bnd_setadjoin (bnd_setadjoin SX0 bnd_emptyset) | |
| 743 | (bnd_setadjoin (bnd_setadjoin SX0 (bnd_setadjoin SX1 bnd_emptyset)) | |
| 744 | bnd_emptyset))) | | |
| 745 | bnd_setadjoin (bnd_setadjoin SX0 bnd_emptyset) | |
| 746 | (bnd_setadjoin (bnd_setadjoin SX0 (bnd_setadjoin SX1 bnd_emptyset)) | |
| 747 | bnd_emptyset) ~= | |
| 748 | bnd_setadjoin (bnd_setadjoin SX2 bnd_emptyset) | |
| 749 | (bnd_setadjoin (bnd_setadjoin SX2 (bnd_setadjoin SX3 bnd_emptyset)) | |
| 750 | bnd_emptyset))))))) = | |
| 751 | True" | |
| 63167 | 752 | by (tactic \<open>unfold_def_tac @{context} []\<close>)
 | 
| 55596 | 753 | |
| 754 |      (* (Annotated_step ("10", "unfold_def"), *)
 | |
| 755 | lemma "bnd_kpairiskpair = | |
| 756 | (ALL Xx Xy. | |
| 757 | bnd_iskpair | |
| 758 | (bnd_setadjoin (bnd_setadjoin Xx bnd_emptyset) | |
| 759 | (bnd_setadjoin (bnd_setadjoin Xx (bnd_setadjoin Xy bnd_emptyset)) | |
| 760 | bnd_emptyset))) & | |
| 761 | bnd_kpair = | |
| 762 | (%Xx Xy. | |
| 763 | bnd_setadjoin (bnd_setadjoin Xx bnd_emptyset) | |
| 764 | (bnd_setadjoin (bnd_setadjoin Xx (bnd_setadjoin Xy bnd_emptyset)) | |
| 765 | bnd_emptyset)) & | |
| 766 | bnd_iskpair = | |
| 767 | (%A. EX Xx. bnd_in Xx (bnd_setunion A) & | |
| 768 | (EX Xy. bnd_in Xy (bnd_setunion A) & | |
| 769 | A = bnd_setadjoin (bnd_setadjoin Xx bnd_emptyset) | |
| 770 | (bnd_setadjoin | |
| 771 | (bnd_setadjoin Xx | |
| 772 | (bnd_setadjoin Xy bnd_emptyset)) | |
| 773 | bnd_emptyset))) & | |
| 774 | (ALL SY5 SY6. | |
| 775 | EX SY7. | |
| 776 | bnd_in SY7 | |
| 777 | (bnd_setunion | |
| 778 | (bnd_setadjoin (bnd_setadjoin SY5 bnd_emptyset) | |
| 779 | (bnd_setadjoin | |
| 780 | (bnd_setadjoin SY5 (bnd_setadjoin SY6 bnd_emptyset)) | |
| 781 | bnd_emptyset))) & | |
| 782 | (EX SY8. | |
| 783 | bnd_in SY8 | |
| 784 | (bnd_setunion | |
| 785 | (bnd_setadjoin (bnd_setadjoin SY5 bnd_emptyset) | |
| 786 | (bnd_setadjoin | |
| 787 | (bnd_setadjoin SY5 (bnd_setadjoin SY6 bnd_emptyset)) | |
| 788 | bnd_emptyset))) & | |
| 789 | bnd_setadjoin (bnd_setadjoin SY5 bnd_emptyset) | |
| 790 | (bnd_setadjoin | |
| 791 | (bnd_setadjoin SY5 (bnd_setadjoin SY6 bnd_emptyset)) | |
| 792 | bnd_emptyset) = | |
| 793 | bnd_setadjoin (bnd_setadjoin SY7 bnd_emptyset) | |
| 794 | (bnd_setadjoin | |
| 795 | (bnd_setadjoin SY7 (bnd_setadjoin SY8 bnd_emptyset)) | |
| 796 | bnd_emptyset))) = | |
| 797 | True | |
| 798 | ==> (ALL SX0 SX1. | |
| 799 | ~ (ALL SX2. | |
| 800 | ~ ~ (~ bnd_in SX2 | |
| 801 | (bnd_setunion | |
| 802 | (bnd_setadjoin (bnd_setadjoin SX0 bnd_emptyset) | |
| 803 | (bnd_setadjoin | |
| 804 | (bnd_setadjoin SX0 | |
| 805 | (bnd_setadjoin SX1 bnd_emptyset)) | |
| 806 | bnd_emptyset))) | | |
| 807 | ~ ~ (ALL SX3. | |
| 808 | ~ ~ (~ bnd_in SX3 | |
| 809 | (bnd_setunion | |
| 810 | (bnd_setadjoin (bnd_setadjoin SX0 bnd_emptyset) | |
| 811 | (bnd_setadjoin (bnd_setadjoin SX0 (bnd_setadjoin SX1 bnd_emptyset)) | |
| 812 | bnd_emptyset))) | | |
| 813 | bnd_setadjoin (bnd_setadjoin SX0 bnd_emptyset) | |
| 814 | (bnd_setadjoin (bnd_setadjoin SX0 (bnd_setadjoin SX1 bnd_emptyset)) | |
| 815 | bnd_emptyset) ~= | |
| 816 | bnd_setadjoin (bnd_setadjoin SX2 bnd_emptyset) | |
| 817 | (bnd_setadjoin (bnd_setadjoin SX2 (bnd_setadjoin SX3 bnd_emptyset)) | |
| 818 | bnd_emptyset)))))) = | |
| 819 | True" | |
| 63167 | 820 | by (tactic \<open>unfold_def_tac @{context} []\<close>)
 | 
| 55596 | 821 | |
| 822 |      (* (Annotated_step ("12", "unfold_def"), *)
 | |
| 823 | lemma "bnd_cCKB6_BLACK = | |
| 824 | (\<lambda>Xu Xv. | |
| 825 | \<forall>Xw. Xw bnd_c1 bnd_c1 \<and> | |
| 826 | (\<forall>Xj Xk. | |
| 827 | Xw Xj Xk \<longrightarrow> | |
| 828 | Xw (bnd_s (bnd_s Xj)) Xk \<and> | |
| 829 | Xw (bnd_s Xj) (bnd_s Xk)) \<longrightarrow> | |
| 830 | Xw Xu Xv) \<and> | |
| 831 | ((((\<forall>SY36 SY37. | |
| 832 | \<not> PUZ114_5_bnd_sK3 SY36 SY37 \<or> | |
| 833 | PUZ114_5_bnd_sK3 (bnd_s (bnd_s SY36)) SY37) \<and> | |
| 834 | (\<forall>SY38 SY39. | |
| 835 | \<not> PUZ114_5_bnd_sK3 SY38 SY39 \<or> | |
| 836 | PUZ114_5_bnd_sK3 (bnd_s SY38) (bnd_s SY39))) \<and> | |
| 837 | PUZ114_5_bnd_sK3 bnd_c1 bnd_c1) \<and> | |
| 838 | \<not> PUZ114_5_bnd_sK3 (bnd_s (bnd_s (bnd_s PUZ114_5_bnd_sK1))) | |
| 839 | (bnd_s PUZ114_5_bnd_sK2)) = | |
| 840 | True \<Longrightarrow> | |
| 841 | (\<not> (\<not> \<not> (\<not> \<not> (\<not> (\<forall>SX0 SX1. | |
| 842 | \<not> PUZ114_5_bnd_sK3 SX0 SX1 \<or> | |
| 843 | PUZ114_5_bnd_sK3 (bnd_s (bnd_s SX0)) SX1) \<or> | |
| 844 | \<not> (\<forall>SX0 SX1. | |
| 845 | \<not> PUZ114_5_bnd_sK3 SX0 SX1 \<or> | |
| 846 | PUZ114_5_bnd_sK3 (bnd_s SX0) (bnd_s SX1))) \<or> | |
| 847 | \<not> PUZ114_5_bnd_sK3 bnd_c1 bnd_c1) \<or> | |
| 848 | \<not> \<not> PUZ114_5_bnd_sK3 (bnd_s (bnd_s (bnd_s PUZ114_5_bnd_sK1))) | |
| 849 | (bnd_s PUZ114_5_bnd_sK2))) = | |
| 850 | True" | |
| 851 | (* | |
| 852 | apply (erule conjE)+ | |
| 853 | apply (erule subst)+ | |
| 854 | apply (tactic {*log_expander 1*})+
 | |
| 855 | apply (rule refl) | |
| 856 | *) | |
| 63167 | 857 | by (tactic \<open>unfold_def_tac @{context} []\<close>)
 | 
| 55596 | 858 | |
| 859 |      (* (Annotated_step ("13", "unfold_def"), *)
 | |
| 860 | lemma "bnd_cCKB6_BLACK = | |
| 861 | (\<lambda>Xu Xv. | |
| 862 | \<forall>Xw. Xw bnd_c1 bnd_c1 \<and> | |
| 863 | (\<forall>Xj Xk. | |
| 864 | Xw Xj Xk \<longrightarrow> | |
| 865 | Xw (bnd_s (bnd_s Xj)) Xk \<and> | |
| 866 | Xw (bnd_s Xj) (bnd_s Xk)) \<longrightarrow> | |
| 867 | Xw Xu Xv) \<and> | |
| 868 | (\<forall>SY30. | |
| 869 | (SY30 (PUZ114_5_bnd_sK4 SY30) (PUZ114_5_bnd_sK5 SY30) \<and> | |
| 870 | (\<not> SY30 (bnd_s (bnd_s (PUZ114_5_bnd_sK4 SY30))) | |
| 871 | (PUZ114_5_bnd_sK5 SY30) \<or> | |
| 872 | \<not> SY30 (bnd_s (PUZ114_5_bnd_sK4 SY30)) | |
| 873 | (bnd_s (PUZ114_5_bnd_sK5 SY30))) \<or> | |
| 874 | \<not> SY30 bnd_c1 bnd_c1) \<or> | |
| 875 | SY30 PUZ114_5_bnd_sK1 PUZ114_5_bnd_sK2) = | |
| 876 | True \<Longrightarrow> | |
| 877 | (\<forall>SX0. (\<not> (\<not> SX0 (PUZ114_5_bnd_sK4 SX0) (PUZ114_5_bnd_sK5 SX0) \<or> | |
| 878 | \<not> (\<not> SX0 (bnd_s (bnd_s (PUZ114_5_bnd_sK4 SX0))) | |
| 879 | (PUZ114_5_bnd_sK5 SX0) \<or> | |
| 880 | \<not> SX0 (bnd_s (PUZ114_5_bnd_sK4 SX0)) | |
| 881 | (bnd_s (PUZ114_5_bnd_sK5 SX0)))) \<or> | |
| 882 | \<not> SX0 bnd_c1 bnd_c1) \<or> | |
| 883 | SX0 PUZ114_5_bnd_sK1 PUZ114_5_bnd_sK2) = | |
| 884 | True" | |
| 885 | (* | |
| 886 | apply (erule conjE)+ | |
| 887 | apply (tactic {*expander_animal 1*})+
 | |
| 888 | apply assumption | |
| 889 | *) | |
| 63167 | 890 | by (tactic \<open>unfold_def_tac @{context} []\<close>)
 | 
| 55596 | 891 | |
| 892 | (*FIXME move this heuristic elsewhere*) | |
| 63167 | 893 | ML \<open> | 
| 55596 | 894 | (*Other than the list (which must not be empty) this function | 
| 895 | expects a parameter indicating the smallest integer. | |
| 896 | (Using Int.minInt isn't always viable).*) | |
| 897 | fun max_int_floored min l = | |
| 898 | if null l then raise List.Empty | |
| 899 | else fold (curry Int.max) l min; | |
| 900 | ||
| 901 | val _ = @{assert} (max_int_floored ~101002 [1]  = 1)
 | |
| 902 | val _ = @{assert} (max_int_floored 0 [1, 3, 5] = 5)
 | |
| 903 | ||
| 904 | fun max_index_floored min l = | |
| 905 | let | |
| 906 | val max = max_int_floored min l | |
| 67399 | 907 | in find_index (pair max #> (=)) l end | 
| 63167 | 908 | \<close> | 
| 55596 | 909 | |
| 63167 | 910 | ML \<open> | 
| 55596 | 911 | max_index_floored 0 [1, 3, 5] | 
| 63167 | 912 | \<close> | 
| 55596 | 913 | |
| 63167 | 914 | ML \<open> | 
| 55596 | 915 | (* | 
| 916 | Given argument ([h_1, ..., h_n], conc), | |
| 917 | obtained from term of form | |
| 918 | h_1 ==> ... ==> h_n ==> conclusion, | |
| 919 | this function indicates which h_i is biggest, | |
| 920 | or NONE if h_n = 0. | |
| 921 | *) | |
| 922 | fun biggest_hypothesis (hypos, _) = | |
| 923 | if null hypos then NONE | |
| 924 | else | |
| 925 | map size_of_term hypos | |
| 926 | |> max_index_floored 0 | |
| 927 | |> SOME | |
| 63167 | 928 | \<close> | 
| 55596 | 929 | |
| 63167 | 930 | ML \<open> | 
| 55596 | 931 | fun biggest_hyp_first_tac i = fn st => | 
| 932 | let | |
| 933 | val results = TERMFUN biggest_hypothesis (SOME i) st | |
| 934 | in | |
| 935 | if null results then no_tac st | |
| 936 | else | |
| 937 | let | |
| 938 | val result = the_single results | |
| 939 | in | |
| 940 | case result of | |
| 941 | NONE => no_tac st | |
| 942 | | SOME n => | |
| 943 | if n > 0 then rotate_tac n i st else no_tac st | |
| 944 | end | |
| 945 | end | |
| 63167 | 946 | \<close> | 
| 55596 | 947 | |
| 948 |      (* (Annotated_step ("6", "unfold_def"), *)
 | |
| 949 | lemma "(\<not> (\<exists>U :: TPTP_Interpret.ind \<Rightarrow> bool. \<forall>V. U V = SEV405_5_bnd_cA)) = True \<Longrightarrow> | |
| 950 | (\<not> \<not> (\<forall>SX0 :: TPTP_Interpret.ind \<Rightarrow> bool. \<not> (\<forall>SX1. \<not> (\<not> (\<not> SX0 SX1 \<or> SEV405_5_bnd_cA) \<or> | |
| 951 | \<not> (\<not> SEV405_5_bnd_cA \<or> SX0 SX1))))) = | |
| 952 | True" | |
| 953 | (* by (tactic {*unfold_def_tac []*}) *)
 | |
| 954 | oops | |
| 955 | ||
| 956 | subsection "Using leo2_tac" | |
| 957 | (*these require PUZ114^5's proof to be loaded | |
| 958 | ||
| 959 | ML {*leo2_tac @{context} (hd prob_names) "50"*}
 | |
| 960 | ||
| 961 | ML {*leo2_tac @{context} (hd prob_names) "4"*}
 | |
| 962 | ||
| 963 | ML {*leo2_tac @{context} (hd prob_names) "9"*}
 | |
| 964 | ||
| 965 |      (* (Annotated_step ("9", "extcnf_combined"), *)
 | |
| 966 | lemma "(\<forall>SY30. | |
| 967 | SY30 bnd_c1 bnd_c1 \<and> | |
| 968 | (\<forall>Xj Xk. | |
| 969 | SY30 Xj Xk \<longrightarrow> | |
| 970 | SY30 (bnd_s (bnd_s Xj)) Xk \<and> | |
| 971 | SY30 (bnd_s Xj) (bnd_s Xk)) \<longrightarrow> | |
| 972 | SY30 bnd_sK1 bnd_sK2) = | |
| 973 | True \<Longrightarrow> | |
| 974 | (\<forall>SY30. | |
| 975 | (SY30 (bnd_sK4 SY30) (bnd_sK5 SY30) \<and> | |
| 976 | (\<not> SY30 (bnd_s (bnd_s (bnd_sK4 SY30))) | |
| 977 | (bnd_sK5 SY30) \<or> | |
| 978 | \<not> SY30 (bnd_s (bnd_sK4 SY30)) | |
| 979 | (bnd_s (bnd_sK5 SY30))) \<or> | |
| 980 | \<not> SY30 bnd_c1 bnd_c1) \<or> | |
| 981 | SY30 bnd_sK1 bnd_sK2) = | |
| 982 | True" | |
| 983 | by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "9") 1*})
 | |
| 984 | *) | |
| 985 | ||
| 986 | ||
| 987 | ||
| 988 | typedecl GEG007_1_bnd_reg | |
| 989 | consts | |
| 990 | GEG007_1_bnd_sK7_SX2 :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> GEG007_1_bnd_reg" | |
| 991 | GEG007_1_bnd_sK6_SX2 :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> GEG007_1_bnd_reg" | |
| 992 | GEG007_1_bnd_a :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool" | |
| 993 | GEG007_1_bnd_catalunya :: "GEG007_1_bnd_reg" | |
| 994 | GEG007_1_bnd_spain :: "GEG007_1_bnd_reg" | |
| 995 | GEG007_1_bnd_c :: "GEG007_1_bnd_reg \<Rightarrow> GEG007_1_bnd_reg \<Rightarrow> bool" | |
| 996 | ||
| 997 |      (* (Annotated_step ("147", "extcnf_forall_neg"), *)
 | |
| 998 | lemma "\<forall>SV13 SV6. | |
| 999 | (\<forall>SX2. \<not> GEG007_1_bnd_c SX2 GEG007_1_bnd_spain \<or> | |
| 1000 | GEG007_1_bnd_c SX2 GEG007_1_bnd_catalunya) = | |
| 1001 | False \<or> | |
| 1002 | GEG007_1_bnd_a SV6 SV13 = False \<Longrightarrow> | |
| 1003 | \<forall>SV6 SV13. | |
| 1004 | (\<not> GEG007_1_bnd_c (GEG007_1_bnd_sK7_SX2 SV13 SV6) GEG007_1_bnd_spain \<or> | |
| 1005 | GEG007_1_bnd_c (GEG007_1_bnd_sK7_SX2 SV13 SV6) GEG007_1_bnd_catalunya) = | |
| 1006 | False \<or> | |
| 1007 | GEG007_1_bnd_a SV6 SV13 = False" | |
| 63167 | 1008 | by (tactic \<open>nonfull_extcnf_combined_tac @{context} [Existential_Var]\<close>)
 | 
| 55596 | 1009 | |
| 1010 |      (* (Annotated_step ("116", "extcnf_forall_neg"), *)
 | |
| 1011 | lemma "\<forall>SV13 SV6. | |
| 1012 | (\<forall>SX2. \<not> \<not> (\<not> \<not> (\<not> GEG007_1_bnd_c SX2 GEG007_1_bnd_catalunya \<or> | |
| 1013 | \<not> \<not> \<not> (\<forall>SX3. | |
| 1014 | \<not> \<not> (\<not> (\<forall>SX4. \<not> GEG007_1_bnd_c SX4 SX3 \<or> GEG007_1_bnd_c SX4 SX2) \<or> | |
| 1015 | \<not> (\<forall>SX4. \<not> GEG007_1_bnd_c SX4 SX3 \<or> | |
| 1016 | GEG007_1_bnd_c SX4 GEG007_1_bnd_catalunya)))) \<or> | |
| 1017 | \<not> \<not> (\<not> GEG007_1_bnd_c SX2 GEG007_1_bnd_spain \<or> | |
| 1018 | \<not> \<not> \<not> (\<forall>SX3. | |
| 1019 | \<not> \<not> (\<not> (\<forall>SX4. \<not> GEG007_1_bnd_c SX4 SX3 \<or> GEG007_1_bnd_c SX4 SX2) \<or> | |
| 1020 | \<not> (\<forall>SX4. \<not> GEG007_1_bnd_c SX4 SX3 \<or> | |
| 1021 | GEG007_1_bnd_c SX4 GEG007_1_bnd_spain)))))) = | |
| 1022 | False \<or> | |
| 1023 | GEG007_1_bnd_a SV6 SV13 = False \<Longrightarrow> | |
| 1024 | \<forall>SV6 SV13. | |
| 1025 | (\<not> \<not> (\<not> \<not> (\<not> GEG007_1_bnd_c (GEG007_1_bnd_sK6_SX2 SV13 SV6) | |
| 1026 | GEG007_1_bnd_catalunya \<or> | |
| 1027 | \<not> \<not> \<not> (\<forall>SY68. | |
| 1028 | \<not> \<not> (\<not> (\<forall>SY69. | |
| 1029 | \<not> GEG007_1_bnd_c SY69 SY68 \<or> | |
| 1030 | GEG007_1_bnd_c SY69 (GEG007_1_bnd_sK6_SX2 SV13 SV6)) \<or> | |
| 1031 | \<not> (\<forall>SX4. \<not> GEG007_1_bnd_c SX4 SY68 \<or> GEG007_1_bnd_c SX4 GEG007_1_bnd_catalunya)))) \<or> | |
| 1032 | \<not> \<not> (\<not> GEG007_1_bnd_c (GEG007_1_bnd_sK6_SX2 SV13 SV6) | |
| 1033 | GEG007_1_bnd_spain \<or> | |
| 1034 | \<not> \<not> \<not> (\<forall>SY71. | |
| 1035 | \<not> \<not> (\<not> (\<forall>SY72. | |
| 1036 | \<not> GEG007_1_bnd_c SY72 SY71 \<or> | |
| 1037 | GEG007_1_bnd_c SY72 (GEG007_1_bnd_sK6_SX2 SV13 SV6)) \<or> | |
| 1038 | \<not> (\<forall>SX4. \<not> GEG007_1_bnd_c SX4 SY71 \<or> GEG007_1_bnd_c SX4 GEG007_1_bnd_spain)))))) = | |
| 1039 | False \<or> | |
| 1040 | GEG007_1_bnd_a SV6 SV13 = False" | |
| 63167 | 1041 | by (tactic \<open>nonfull_extcnf_combined_tac @{context} [Existential_Var]\<close>)
 | 
| 55596 | 1042 | |
| 1043 | consts PUZ107_5_bnd_sK1_SX0 :: | |
| 1044 | "TPTP_Interpret.ind | |
| 1045 | \<Rightarrow> TPTP_Interpret.ind | |
| 1046 | \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool" | |
| 1047 | ||
| 61076 | 1048 | lemma "\<forall>(SV4::TPTP_Interpret.ind) (SV8::TPTP_Interpret.ind) | 
| 1049 | (SV6::TPTP_Interpret.ind) (SV2::TPTP_Interpret.ind) | |
| 1050 | (SV3::TPTP_Interpret.ind) SV1::TPTP_Interpret.ind. | |
| 55596 | 1051 | ((SV1 \<noteq> SV3) = False \<or> PUZ107_5_bnd_sK1_SX0 SV1 SV2 SV6 SV8 = False) \<or> | 
| 1052 | PUZ107_5_bnd_sK1_SX0 SV3 SV4 SV6 SV8 = False \<Longrightarrow> | |
| 61076 | 1053 | \<forall>(SV4::TPTP_Interpret.ind) (SV8::TPTP_Interpret.ind) | 
| 1054 | (SV6::TPTP_Interpret.ind) (SV2::TPTP_Interpret.ind) | |
| 1055 | (SV3::TPTP_Interpret.ind) SV1::TPTP_Interpret.ind. | |
| 55596 | 1056 | ((SV1 = SV3) = True \<or> PUZ107_5_bnd_sK1_SX0 SV1 SV2 SV6 SV8 = False) \<or> | 
| 1057 | PUZ107_5_bnd_sK1_SX0 SV3 SV4 SV6 SV8 = False" | |
| 63167 | 1058 | by (tactic \<open>nonfull_extcnf_combined_tac @{context} [Not_neg]\<close>)
 | 
| 55596 | 1059 | |
| 1060 | lemma " | |
| 61076 | 1061 | \<forall>(SV8::TPTP_Interpret.ind) (SV6::TPTP_Interpret.ind) | 
| 1062 | (SV4::TPTP_Interpret.ind) (SV2::TPTP_Interpret.ind) | |
| 1063 | (SV3::TPTP_Interpret.ind) SV1::TPTP_Interpret.ind. | |
| 55596 | 1064 | ((SV1 \<noteq> SV3 \<or> SV2 \<noteq> SV4) = False \<or> | 
| 1065 | PUZ107_5_bnd_sK1_SX0 SV1 SV2 SV6 SV8 = False) \<or> | |
| 1066 | PUZ107_5_bnd_sK1_SX0 SV3 SV4 SV6 SV8 = False \<Longrightarrow> | |
| 61076 | 1067 | \<forall>(SV4::TPTP_Interpret.ind) (SV8::TPTP_Interpret.ind) | 
| 1068 | (SV6::TPTP_Interpret.ind) (SV2::TPTP_Interpret.ind) | |
| 1069 | (SV3::TPTP_Interpret.ind) SV1::TPTP_Interpret.ind. | |
| 55596 | 1070 | ((SV1 \<noteq> SV3) = False \<or> PUZ107_5_bnd_sK1_SX0 SV1 SV2 SV6 SV8 = False) \<or> | 
| 1071 | PUZ107_5_bnd_sK1_SX0 SV3 SV4 SV6 SV8 = False" | |
| 63167 | 1072 | by (tactic \<open>nonfull_extcnf_combined_tac @{context} [Or_neg]\<close>)
 | 
| 55596 | 1073 | |
| 1074 | consts | |
| 1075 | NUM016_5_bnd_a :: TPTP_Interpret.ind | |
| 1076 | NUM016_5_bnd_prime :: "TPTP_Interpret.ind \<Rightarrow> bool" | |
| 1077 | NUM016_5_bnd_factorial_plus_one :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind" | |
| 1078 | NUM016_5_bnd_prime_divisor :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind" | |
| 1079 | NUM016_5_bnd_divides :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool" | |
| 1080 | NUM016_5_bnd_less :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool" | |
| 1081 | ||
| 1082 |      (* (Annotated_step ("6", "unfold_def"), *)
 | |
| 61076 | 1083 | lemma "((((((((((((\<forall>X::TPTP_Interpret.ind. \<not> NUM016_5_bnd_less X X) \<and> | 
| 1084 | (\<forall>(X::TPTP_Interpret.ind) | |
| 1085 | Y::TPTP_Interpret.ind. | |
| 55596 | 1086 | \<not> NUM016_5_bnd_less X Y \<or> \<not> NUM016_5_bnd_less Y X)) \<and> | 
| 61076 | 1087 | (\<forall>X::TPTP_Interpret.ind. NUM016_5_bnd_divides X X)) \<and> | 
| 1088 | (\<forall>(X::TPTP_Interpret.ind) | |
| 1089 | (Y::TPTP_Interpret.ind) | |
| 1090 | Z::TPTP_Interpret.ind. | |
| 55596 | 1091 | (\<not> NUM016_5_bnd_divides X Y \<or> \<not> NUM016_5_bnd_divides Y Z) \<or> | 
| 1092 | NUM016_5_bnd_divides X Z)) \<and> | |
| 61076 | 1093 | (\<forall>(X::TPTP_Interpret.ind) Y::TPTP_Interpret.ind. | 
| 55596 | 1094 | \<not> NUM016_5_bnd_divides X Y \<or> \<not> NUM016_5_bnd_less Y X)) \<and> | 
| 61076 | 1095 | (\<forall>X::TPTP_Interpret.ind. | 
| 55596 | 1096 | NUM016_5_bnd_less X (NUM016_5_bnd_factorial_plus_one X))) \<and> | 
| 61076 | 1097 | (\<forall>(X::TPTP_Interpret.ind) Y::TPTP_Interpret.ind. | 
| 55596 | 1098 | \<not> NUM016_5_bnd_divides X (NUM016_5_bnd_factorial_plus_one Y) \<or> | 
| 1099 | NUM016_5_bnd_less Y X)) \<and> | |
| 61076 | 1100 | (\<forall>X::TPTP_Interpret.ind. | 
| 55596 | 1101 | NUM016_5_bnd_prime X \<or> | 
| 1102 | NUM016_5_bnd_divides (NUM016_5_bnd_prime_divisor X) X)) \<and> | |
| 61076 | 1103 | (\<forall>X::TPTP_Interpret.ind. | 
| 55596 | 1104 | NUM016_5_bnd_prime X \<or> | 
| 1105 | NUM016_5_bnd_prime (NUM016_5_bnd_prime_divisor X))) \<and> | |
| 61076 | 1106 | (\<forall>X::TPTP_Interpret.ind. | 
| 55596 | 1107 | NUM016_5_bnd_prime X \<or> | 
| 1108 | NUM016_5_bnd_less (NUM016_5_bnd_prime_divisor X) X)) \<and> | |
| 1109 | NUM016_5_bnd_prime NUM016_5_bnd_a) \<and> | |
| 61076 | 1110 | (\<forall>X::TPTP_Interpret.ind. | 
| 55596 | 1111 | (\<not> NUM016_5_bnd_prime X \<or> \<not> NUM016_5_bnd_less NUM016_5_bnd_a X) \<or> | 
| 1112 | NUM016_5_bnd_less (NUM016_5_bnd_factorial_plus_one NUM016_5_bnd_a) X)) = | |
| 1113 | True \<Longrightarrow> | |
| 61076 | 1114 | (\<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> (\<forall>SX0::TPTP_Interpret.ind. | 
| 55596 | 1115 | \<not> NUM016_5_bnd_less SX0 SX0) \<or> | 
| 61076 | 1116 | \<not> (\<forall>(SX0::TPTP_Interpret.ind) | 
| 1117 | SX1::TPTP_Interpret.ind. | |
| 55596 | 1118 | \<not> NUM016_5_bnd_less SX0 SX1 \<or> \<not> NUM016_5_bnd_less SX1 SX0)) \<or> | 
| 61076 | 1119 | \<not> (\<forall>SX0::TPTP_Interpret.ind. | 
| 55596 | 1120 | NUM016_5_bnd_divides SX0 SX0)) \<or> | 
| 61076 | 1121 | \<not> (\<forall>(SX0::TPTP_Interpret.ind) | 
| 1122 | (SX1::TPTP_Interpret.ind) | |
| 1123 | SX2::TPTP_Interpret.ind. | |
| 55596 | 1124 | (\<not> NUM016_5_bnd_divides SX0 SX1 \<or> | 
| 1125 | \<not> NUM016_5_bnd_divides SX1 SX2) \<or> | |
| 1126 | NUM016_5_bnd_divides SX0 SX2)) \<or> | |
| 61076 | 1127 | \<not> (\<forall>(SX0::TPTP_Interpret.ind) | 
| 1128 | SX1::TPTP_Interpret.ind. | |
| 55596 | 1129 | \<not> NUM016_5_bnd_divides SX0 SX1 \<or> | 
| 1130 | \<not> NUM016_5_bnd_less SX1 SX0)) \<or> | |
| 61076 | 1131 | \<not> (\<forall>SX0::TPTP_Interpret.ind. | 
| 55596 | 1132 | NUM016_5_bnd_less SX0 (NUM016_5_bnd_factorial_plus_one SX0))) \<or> | 
| 61076 | 1133 | \<not> (\<forall>(SX0::TPTP_Interpret.ind) SX1::TPTP_Interpret.ind. | 
| 55596 | 1134 | \<not> NUM016_5_bnd_divides SX0 (NUM016_5_bnd_factorial_plus_one SX1) \<or> | 
| 1135 | NUM016_5_bnd_less SX1 SX0)) \<or> | |
| 61076 | 1136 | \<not> (\<forall>SX0::TPTP_Interpret.ind. | 
| 55596 | 1137 | NUM016_5_bnd_prime SX0 \<or> | 
| 1138 | NUM016_5_bnd_divides (NUM016_5_bnd_prime_divisor SX0) SX0)) \<or> | |
| 61076 | 1139 | \<not> (\<forall>SX0::TPTP_Interpret.ind. | 
| 55596 | 1140 | NUM016_5_bnd_prime SX0 \<or> NUM016_5_bnd_prime (NUM016_5_bnd_prime_divisor SX0))) \<or> | 
| 61076 | 1141 | \<not> (\<forall>SX0::TPTP_Interpret.ind. | 
| 55596 | 1142 | NUM016_5_bnd_prime SX0 \<or> | 
| 1143 | NUM016_5_bnd_less (NUM016_5_bnd_prime_divisor SX0) | |
| 1144 | SX0)) \<or> | |
| 1145 | \<not> NUM016_5_bnd_prime NUM016_5_bnd_a) \<or> | |
| 61076 | 1146 | \<not> (\<forall>SX0::TPTP_Interpret.ind. | 
| 55596 | 1147 | (\<not> NUM016_5_bnd_prime SX0 \<or> \<not> NUM016_5_bnd_less NUM016_5_bnd_a SX0) \<or> | 
| 1148 | NUM016_5_bnd_less (NUM016_5_bnd_factorial_plus_one NUM016_5_bnd_a) | |
| 1149 | SX0))) = | |
| 1150 | True" | |
| 63167 | 1151 | by (tactic \<open>unfold_def_tac @{context} []\<close>)
 | 
| 55596 | 1152 | |
| 1153 |      (* (Annotated_step ("3", "unfold_def"), *)
 | |
| 1154 | lemma "(~ ((((((((((((ALL X. ~ bnd_less X X) & | |
| 1155 | (ALL X Y. ~ bnd_less X Y | ~ bnd_less Y X)) & | |
| 1156 | (ALL X. bnd_divides X X)) & | |
| 1157 | (ALL X Y Z. | |
| 1158 | (~ bnd_divides X Y | ~ bnd_divides Y Z) | | |
| 1159 | bnd_divides X Z)) & | |
| 1160 | (ALL X Y. ~ bnd_divides X Y | ~ bnd_less Y X)) & | |
| 1161 | (ALL X. bnd_less X (bnd_factorial_plus_one X))) & | |
| 1162 | (ALL X Y. | |
| 1163 | ~ bnd_divides X (bnd_factorial_plus_one Y) | | |
| 1164 | bnd_less Y X)) & | |
| 1165 | (ALL X. bnd_prime X | bnd_divides (bnd_prime_divisor X) X)) & | |
| 1166 | (ALL X. bnd_prime X | bnd_prime (bnd_prime_divisor X))) & | |
| 1167 | (ALL X. bnd_prime X | bnd_less (bnd_prime_divisor X) X)) & | |
| 1168 | bnd_prime bnd_a) & | |
| 1169 | (ALL X. (~ bnd_prime X | ~ bnd_less bnd_a X) | | |
| 1170 | bnd_less (bnd_factorial_plus_one bnd_a) X))) = | |
| 1171 | False | |
| 1172 | ==> (~ ((((((((((((ALL X. ~ bnd_less X X) & | |
| 1173 | (ALL X Y. ~ bnd_less X Y | ~ bnd_less Y X)) & | |
| 1174 | (ALL X. bnd_divides X X)) & | |
| 1175 | (ALL X Y Z. | |
| 1176 | (~ bnd_divides X Y | ~ bnd_divides Y Z) | | |
| 1177 | bnd_divides X Z)) & | |
| 1178 | (ALL X Y. ~ bnd_divides X Y | ~ bnd_less Y X)) & | |
| 1179 | (ALL X. bnd_less X (bnd_factorial_plus_one X))) & | |
| 1180 | (ALL X Y. | |
| 1181 | ~ bnd_divides X (bnd_factorial_plus_one Y) | | |
| 1182 | bnd_less Y X)) & | |
| 1183 | (ALL X. bnd_prime X | | |
| 1184 | bnd_divides (bnd_prime_divisor X) X)) & | |
| 1185 | (ALL X. bnd_prime X | bnd_prime (bnd_prime_divisor X))) & | |
| 1186 | (ALL X. bnd_prime X | bnd_less (bnd_prime_divisor X) X)) & | |
| 1187 | bnd_prime bnd_a) & | |
| 1188 | (ALL X. (~ bnd_prime X | ~ bnd_less bnd_a X) | | |
| 1189 | bnd_less (bnd_factorial_plus_one bnd_a) X))) = | |
| 1190 | False" | |
| 63167 | 1191 | by (tactic \<open>unfold_def_tac @{context} []\<close>)
 | 
| 55596 | 1192 | |
| 1193 | (* SET062^6.p.out | |
| 1194 |       [[(Annotated_step ("3", "unfold_def"), *)
 | |
| 1195 | lemma "(\<forall>Z3. False \<longrightarrow> bnd_cA Z3) = False \<Longrightarrow> | |
| 1196 | (\<forall>Z3. False \<longrightarrow> bnd_cA Z3) = False" | |
| 63167 | 1197 | by (tactic \<open>unfold_def_tac @{context} []\<close>)
 | 
| 55596 | 1198 | |
| 1199 | (* | |
| 1200 | (* SEU559^2.p.out *) | |
| 1201 |    (* [[(Annotated_step ("3", "unfold_def"), *)
 | |
| 1202 | lemma "bnd_subset = (\<lambda>A B. \<forall>Xx. bnd_in Xx A \<longrightarrow> bnd_in Xx B) \<and> | |
| 1203 | (\<forall>A B. (\<forall>Xx. bnd_in Xx A \<longrightarrow> bnd_in Xx B) \<longrightarrow> | |
| 1204 | bnd_subset A B) = | |
| 1205 | False \<Longrightarrow> | |
| 1206 | (\<forall>SY0 SY1. | |
| 1207 | (\<forall>Xx. bnd_in Xx SY0 \<longrightarrow> bnd_in Xx SY1) \<longrightarrow> | |
| 1208 | (\<forall>SY5. bnd_in SY5 SY0 \<longrightarrow> bnd_in SY5 SY1)) = | |
| 1209 | False" | |
| 1210 | by (tactic {*unfold_def_tac [@{thm bnd_subset_def}]*})
 | |
| 1211 | ||
| 1212 | (* SEU559^2.p.out | |
| 1213 |     [[(Annotated_step ("6", "unfold_def"), *)
 | |
| 1214 | lemma "(\<not> (\<exists>Xx. \<forall>Xy. Xx \<longrightarrow> Xy)) = True \<Longrightarrow> | |
| 1215 | (\<not> \<not> (\<forall>SX0. \<not> (\<forall>SX1. \<not> SX0 \<or> SX1))) = True" | |
| 1216 | by (tactic {*unfold_def_tac []*})
 | |
| 1217 | ||
| 1218 | (* SEU502^2.p.out | |
| 1219 |     [[(Annotated_step ("3", "unfold_def"), *)
 | |
| 1220 | lemma "bnd_emptysetE = | |
| 1221 | (\<forall>Xx. bnd_in Xx bnd_emptyset \<longrightarrow> (\<forall>Xphi. Xphi)) \<and> | |
| 1222 | (bnd_emptysetE \<longrightarrow> | |
| 1223 | (\<forall>Xx. bnd_in Xx bnd_emptyset \<longrightarrow> False)) = | |
| 1224 | False \<Longrightarrow> | |
| 1225 | ((\<forall>Xx. bnd_in Xx bnd_emptyset \<longrightarrow> (\<forall>Xphi. Xphi)) \<longrightarrow> | |
| 1226 | (\<forall>Xx. bnd_in Xx bnd_emptyset \<longrightarrow> False)) = | |
| 1227 | False" | |
| 1228 | by (tactic {*unfold_def_tac [@{thm bnd_emptysetE_def}]*})
 | |
| 1229 | *) | |
| 1230 | ||
| 1231 | typedecl AGT037_2_bnd_mu | |
| 1232 | consts | |
| 1233 | AGT037_2_bnd_sK1_SX0 :: TPTP_Interpret.ind | |
| 1234 | AGT037_2_bnd_cola :: AGT037_2_bnd_mu | |
| 1235 | AGT037_2_bnd_jan :: AGT037_2_bnd_mu | |
| 1236 | AGT037_2_bnd_possibly_likes :: "AGT037_2_bnd_mu \<Rightarrow> AGT037_2_bnd_mu \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool" | |
| 1237 | AGT037_2_bnd_sK5_SY68 :: | |
| 1238 | "TPTP_Interpret.ind | |
| 1239 | \<Rightarrow> AGT037_2_bnd_mu | |
| 1240 | \<Rightarrow> AGT037_2_bnd_mu | |
| 1241 | \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind" | |
| 1242 | AGT037_2_bnd_likes :: "AGT037_2_bnd_mu \<Rightarrow> AGT037_2_bnd_mu \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool" | |
| 1243 | AGT037_2_bnd_very_much_likes :: "AGT037_2_bnd_mu \<Rightarrow> AGT037_2_bnd_mu \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool" | |
| 1244 | AGT037_2_bnd_a1 :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool" | |
| 1245 | AGT037_2_bnd_a2 :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool" | |
| 1246 | AGT037_2_bnd_a3 :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool" | |
| 1247 | ||
| 1248 | (*test that nullary skolem terms are OK*) | |
| 1249 |      (* (Annotated_step ("79", "extcnf_forall_neg"), *)
 | |
| 61076 | 1250 | lemma "(\<forall>SX0::TPTP_Interpret.ind. | 
| 55596 | 1251 | AGT037_2_bnd_possibly_likes AGT037_2_bnd_jan AGT037_2_bnd_cola SX0) = | 
| 1252 | False \<Longrightarrow> | |
| 1253 | AGT037_2_bnd_possibly_likes AGT037_2_bnd_jan AGT037_2_bnd_cola AGT037_2_bnd_sK1_SX0 = | |
| 1254 | False" | |
| 63167 | 1255 | by (tactic \<open>nonfull_extcnf_combined_tac @{context} [Existential_Var]\<close>)
 | 
| 55596 | 1256 | |
| 1257 |      (* (Annotated_step ("202", "extcnf_forall_neg"), *)
 | |
| 61076 | 1258 | lemma "\<forall>(SV13::TPTP_Interpret.ind) (SV39::AGT037_2_bnd_mu) (SV29::AGT037_2_bnd_mu) | 
| 1259 | SV45::TPTP_Interpret.ind. | |
| 1260 | ((((\<forall>SY68::TPTP_Interpret.ind. | |
| 55596 | 1261 | \<not> AGT037_2_bnd_a1 SV45 SY68 \<or> | 
| 1262 | AGT037_2_bnd_likes SV29 SV39 SY68) = | |
| 1263 | False \<or> | |
| 61076 | 1264 | (\<not> (\<forall>SY69::TPTP_Interpret.ind. | 
| 55596 | 1265 | \<not> AGT037_2_bnd_a2 SV45 SY69 \<or> | 
| 1266 | AGT037_2_bnd_likes SV29 SV39 SY69)) = | |
| 1267 | True) \<or> | |
| 1268 | AGT037_2_bnd_likes SV29 SV39 SV45 = False) \<or> | |
| 1269 | AGT037_2_bnd_very_much_likes SV29 SV39 SV45 = True) \<or> | |
| 1270 | AGT037_2_bnd_a3 SV13 SV45 = False \<Longrightarrow> | |
| 61076 | 1271 | \<forall>(SV29::AGT037_2_bnd_mu) (SV39::AGT037_2_bnd_mu) (SV13::TPTP_Interpret.ind) | 
| 1272 | SV45::TPTP_Interpret.ind. | |
| 55596 | 1273 | ((((\<not> AGT037_2_bnd_a1 SV45 | 
| 1274 | (AGT037_2_bnd_sK5_SY68 SV13 SV39 SV29 SV45) \<or> | |
| 1275 | AGT037_2_bnd_likes SV29 SV39 | |
| 1276 | (AGT037_2_bnd_sK5_SY68 SV13 SV39 SV29 SV45)) = | |
| 1277 | False \<or> | |
| 61076 | 1278 | (\<not> (\<forall>SY69::TPTP_Interpret.ind. | 
| 55596 | 1279 | \<not> AGT037_2_bnd_a2 SV45 SY69 \<or> | 
| 1280 | AGT037_2_bnd_likes SV29 SV39 SY69)) = | |
| 1281 | True) \<or> | |
| 1282 | AGT037_2_bnd_likes SV29 SV39 SV45 = False) \<or> | |
| 1283 | AGT037_2_bnd_very_much_likes SV29 SV39 SV45 = True) \<or> | |
| 1284 | AGT037_2_bnd_a3 SV13 SV45 = False" | |
| 1285 | (* | |
| 1286 | apply (rule allI)+ | |
| 1287 | apply (erule_tac x = "SV13" in allE) | |
| 1288 | apply (erule_tac x = "SV39" in allE) | |
| 1289 | apply (erule_tac x = "SV29" in allE) | |
| 1290 | apply (erule_tac x = "SV45" in allE) | |
| 1291 | apply (erule disjE)+ | |
| 1292 | defer | |
| 1293 | apply (tactic {*clause_breaker 1*})+
 | |
| 1294 | apply (drule_tac sk = "bnd_sK5_SY68 SV13 SV39 SV29 SV45" in leo2_skolemise) | |
| 1295 | defer | |
| 1296 | apply (tactic {*clause_breaker 1*})
 | |
| 1297 | apply (tactic {*nonfull_extcnf_combined_tac []*})
 | |
| 1298 | *) | |
| 63167 | 1299 | by (tactic \<open>nonfull_extcnf_combined_tac @{context} [Existential_Var]\<close>)
 | 
| 55596 | 1300 | |
| 1301 | (*(*NUM667^1*) | |
| 1302 | lemma "\<forall>SV12 SV13 SV14 SV9 SV10 SV11. | |
| 1303 | ((((bnd_less SV12 SV13 = bnd_less SV11 SV10) = False \<or> | |
| 1304 | (SV14 = SV13) = False) \<or> | |
| 1305 | bnd_less SV12 SV14 = False) \<or> | |
| 1306 | bnd_less SV9 SV10 = True) \<or> | |
| 1307 | (SV9 = SV11) = False \<Longrightarrow> | |
| 1308 | \<forall>SV9 SV14 SV10 SV11 SV13 SV12. | |
| 1309 | ((((bnd_less SV12 SV13 = False \<or> | |
| 1310 | bnd_less SV11 SV10 = False) \<or> | |
| 1311 | (SV14 = SV13) = False) \<or> | |
| 1312 | bnd_less SV12 SV14 = False) \<or> | |
| 1313 | bnd_less SV9 SV10 = True) \<or> | |
| 1314 | (SV9 = SV11) = False" | |
| 1315 | (* | |
| 1316 | apply (tactic {*
 | |
| 1317 | extcnf_combined_tac NONE | |
| 1318 | [ConstsDiff, | |
| 1319 | StripQuantifiers] | |
| 1320 | []*}) | |
| 1321 | *) | |
| 1322 | (* | |
| 1323 | apply (rule allI)+ | |
| 1324 | apply (erule_tac x = "SV12" in allE) | |
| 1325 | apply (erule_tac x = "SV13" in allE) | |
| 1326 | apply (erule_tac x = "SV14" in allE) | |
| 1327 | apply (erule_tac x = "SV9" in allE) | |
| 1328 | apply (erule_tac x = "SV10" in allE) | |
| 1329 | apply (erule_tac x = "SV11" in allE) | |
| 1330 | *) | |
| 1331 | by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "300") 1*})
 | |
| 1332 | ||
| 1333 | ||
| 1334 | (*NUM667^1 node 302 -- dec*) | |
| 1335 | lemma "\<forall>SV12 SV13 SV14 SV9 SV10 SV11. | |
| 1336 | ((((bnd_less SV12 SV13 = bnd_less SV11 SV10) = False \<or> | |
| 1337 | (SV14 = SV13) = False) \<or> | |
| 1338 | bnd_less SV12 SV14 = False) \<or> | |
| 1339 | bnd_less SV9 SV10 = True) \<or> | |
| 1340 | (SV9 = SV11) = | |
| 1341 | False \<Longrightarrow> | |
| 1342 | \<forall>SV9 SV14 SV10 SV13 SV11 SV12. | |
| 1343 | (((((SV12 = SV11) = False \<or> (SV13 = SV10) = False) \<or> | |
| 1344 | (SV14 = SV13) = False) \<or> | |
| 1345 | bnd_less SV12 SV14 = False) \<or> | |
| 1346 | bnd_less SV9 SV10 = True) \<or> | |
| 1347 | (SV9 = SV11) = | |
| 1348 | False" | |
| 1349 | by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "302") 1*})
 | |
| 1350 | *) | |
| 1351 | ||
| 1352 | ||
| 1353 | (* | |
| 1354 | (*CSR122^2*) | |
| 1355 |      (* (Annotated_step ("23", "extuni_bool2"), *)
 | |
| 1356 | lemma "(bnd_holdsDuring_THFTYPE_IiooI | |
| 1357 | (bnd_lYearFn_THFTYPE_IiiI bnd_n2009_THFTYPE_i) | |
| 1358 | (\<not> (\<not> bnd_likes_THFTYPE_IiioI bnd_lMary_THFTYPE_i | |
| 1359 | bnd_lBill_THFTYPE_i \<or> | |
| 1360 | \<not> bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i | |
| 1361 | bnd_lBill_THFTYPE_i)) = | |
| 1362 | bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i | |
| 1363 | bnd_lBill_THFTYPE_i) = | |
| 1364 | False \<Longrightarrow> | |
| 1365 | bnd_holdsDuring_THFTYPE_IiooI | |
| 1366 | (bnd_lYearFn_THFTYPE_IiiI bnd_n2009_THFTYPE_i) | |
| 1367 | (\<not> (\<not> bnd_likes_THFTYPE_IiioI bnd_lMary_THFTYPE_i | |
| 1368 | bnd_lBill_THFTYPE_i \<or> | |
| 1369 | \<not> bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i | |
| 1370 | bnd_lBill_THFTYPE_i)) = | |
| 1371 | True \<or> | |
| 1372 | bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i | |
| 1373 | bnd_lBill_THFTYPE_i = | |
| 1374 | True" | |
| 1375 | (* apply (erule extuni_bool2) *) | |
| 1376 | (* done *) | |
| 1377 | by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "23") 1*})
 | |
| 1378 | ||
| 1379 |      (* (Annotated_step ("24", "extuni_bool1"), *)
 | |
| 1380 | lemma "(bnd_holdsDuring_THFTYPE_IiooI | |
| 1381 | (bnd_lYearFn_THFTYPE_IiiI bnd_n2009_THFTYPE_i) | |
| 1382 | (\<not> (\<not> bnd_likes_THFTYPE_IiioI bnd_lMary_THFTYPE_i | |
| 1383 | bnd_lBill_THFTYPE_i \<or> | |
| 1384 | \<not> bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i | |
| 1385 | bnd_lBill_THFTYPE_i)) = | |
| 1386 | bnd_likes_THFTYPE_IiioI bnd_lMary_THFTYPE_i | |
| 1387 | bnd_lBill_THFTYPE_i) = | |
| 1388 | False \<Longrightarrow> | |
| 1389 | bnd_holdsDuring_THFTYPE_IiooI | |
| 1390 | (bnd_lYearFn_THFTYPE_IiiI bnd_n2009_THFTYPE_i) | |
| 1391 | (\<not> (\<not> bnd_likes_THFTYPE_IiioI bnd_lMary_THFTYPE_i | |
| 1392 | bnd_lBill_THFTYPE_i \<or> | |
| 1393 | \<not> bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i | |
| 1394 | bnd_lBill_THFTYPE_i)) = | |
| 1395 | False \<or> | |
| 1396 | bnd_likes_THFTYPE_IiioI bnd_lMary_THFTYPE_i | |
| 1397 | bnd_lBill_THFTYPE_i = | |
| 1398 | False" | |
| 1399 | (* apply (erule extuni_bool1) *) | |
| 1400 | (* done *) | |
| 1401 | by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "24") 1*})
 | |
| 1402 | ||
| 1403 |      (* (Annotated_step ("25", "extuni_bool2"), *)
 | |
| 1404 | lemma "(bnd_holdsDuring_THFTYPE_IiooI | |
| 1405 | (bnd_lYearFn_THFTYPE_IiiI bnd_n2009_THFTYPE_i) | |
| 1406 | (\<not> (\<not> bnd_likes_THFTYPE_IiioI bnd_lMary_THFTYPE_i | |
| 1407 | bnd_lBill_THFTYPE_i \<or> | |
| 1408 | \<not> bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i | |
| 1409 | bnd_lBill_THFTYPE_i)) = | |
| 1410 | bnd_likes_THFTYPE_IiioI bnd_lMary_THFTYPE_i | |
| 1411 | bnd_lBill_THFTYPE_i) = | |
| 1412 | False \<Longrightarrow> | |
| 1413 | bnd_holdsDuring_THFTYPE_IiooI | |
| 1414 | (bnd_lYearFn_THFTYPE_IiiI bnd_n2009_THFTYPE_i) | |
| 1415 | (\<not> (\<not> bnd_likes_THFTYPE_IiioI bnd_lMary_THFTYPE_i | |
| 1416 | bnd_lBill_THFTYPE_i \<or> | |
| 1417 | \<not> bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i | |
| 1418 | bnd_lBill_THFTYPE_i)) = | |
| 1419 | True \<or> | |
| 1420 | bnd_likes_THFTYPE_IiioI bnd_lMary_THFTYPE_i | |
| 1421 | bnd_lBill_THFTYPE_i = | |
| 1422 | True" | |
| 1423 | (* apply (erule extuni_bool2) *) | |
| 1424 | (* done *) | |
| 1425 | by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "25") 1*})
 | |
| 1426 | ||
| 1427 |      (* (Annotated_step ("26", "extuni_bool1"), *)
 | |
| 1428 | lemma "\<forall>SV2. (bnd_holdsDuring_THFTYPE_IiooI | |
| 1429 | (bnd_lYearFn_THFTYPE_IiiI bnd_n2009_THFTYPE_i) | |
| 1430 | (\<not> (\<not> bnd_likes_THFTYPE_IiioI | |
| 1431 | bnd_lMary_THFTYPE_i | |
| 1432 | bnd_lBill_THFTYPE_i \<or> | |
| 1433 | \<not> bnd_likes_THFTYPE_IiioI | |
| 1434 | bnd_lSue_THFTYPE_i | |
| 1435 | bnd_lBill_THFTYPE_i)) = | |
| 1436 | bnd_holdsDuring_THFTYPE_IiooI SV2 True) = | |
| 1437 | False \<Longrightarrow> | |
| 1438 | \<forall>SV2. bnd_holdsDuring_THFTYPE_IiooI | |
| 1439 | (bnd_lYearFn_THFTYPE_IiiI bnd_n2009_THFTYPE_i) | |
| 1440 | (\<not> (\<not> bnd_likes_THFTYPE_IiioI | |
| 1441 | bnd_lMary_THFTYPE_i bnd_lBill_THFTYPE_i \<or> | |
| 1442 | \<not> bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i | |
| 1443 | bnd_lBill_THFTYPE_i)) = | |
| 1444 | False \<or> | |
| 1445 | bnd_holdsDuring_THFTYPE_IiooI SV2 True = False" | |
| 1446 | (* apply (rule allI, erule allE) *) | |
| 1447 | (* apply (erule extuni_bool1) *) | |
| 1448 | (* done *) | |
| 1449 | by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "26") 1*})
 | |
| 1450 | ||
| 1451 |      (* (Annotated_step ("27", "extuni_bool2"), *)
 | |
| 1452 | lemma "\<forall>SV2. (bnd_holdsDuring_THFTYPE_IiooI | |
| 1453 | (bnd_lYearFn_THFTYPE_IiiI bnd_n2009_THFTYPE_i) | |
| 1454 | (\<not> (\<not> bnd_likes_THFTYPE_IiioI | |
| 1455 | bnd_lMary_THFTYPE_i | |
| 1456 | bnd_lBill_THFTYPE_i \<or> | |
| 1457 | \<not> bnd_likes_THFTYPE_IiioI | |
| 1458 | bnd_lSue_THFTYPE_i | |
| 1459 | bnd_lBill_THFTYPE_i)) = | |
| 1460 | bnd_holdsDuring_THFTYPE_IiooI SV2 True) = | |
| 1461 | False \<Longrightarrow> | |
| 1462 | \<forall>SV2. bnd_holdsDuring_THFTYPE_IiooI | |
| 1463 | (bnd_lYearFn_THFTYPE_IiiI bnd_n2009_THFTYPE_i) | |
| 1464 | (\<not> (\<not> bnd_likes_THFTYPE_IiioI | |
| 1465 | bnd_lMary_THFTYPE_i bnd_lBill_THFTYPE_i \<or> | |
| 1466 | \<not> bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i | |
| 1467 | bnd_lBill_THFTYPE_i)) = | |
| 1468 | True \<or> | |
| 1469 | bnd_holdsDuring_THFTYPE_IiooI SV2 True = True" | |
| 1470 | (* apply (rule allI, erule allE) *) | |
| 1471 | (* apply (erule extuni_bool2) *) | |
| 1472 | (* done *) | |
| 1473 | by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "27") 1*})
 | |
| 1474 | ||
| 1475 |      (* (Annotated_step ("30", "extuni_bool1"), *)
 | |
| 1476 | lemma "((\<not> (\<not> bnd_likes_THFTYPE_IiioI bnd_lMary_THFTYPE_i | |
| 1477 | bnd_lBill_THFTYPE_i \<or> | |
| 1478 | \<not> bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i | |
| 1479 | bnd_lBill_THFTYPE_i)) = | |
| 1480 | True) = | |
| 1481 | False \<Longrightarrow> | |
| 1482 | (\<not> (\<not> bnd_likes_THFTYPE_IiioI bnd_lMary_THFTYPE_i | |
| 1483 | bnd_lBill_THFTYPE_i \<or> | |
| 1484 | \<not> bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i | |
| 1485 | bnd_lBill_THFTYPE_i)) = | |
| 1486 | False \<or> | |
| 1487 | True = False" | |
| 1488 | (* apply (erule extuni_bool1) *) | |
| 1489 | (* done *) | |
| 1490 | by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "30") 1*})
 | |
| 1491 | ||
| 1492 |      (* (Annotated_step ("29", "extuni_bind"), *)
 | |
| 1493 | lemma "(bnd_lYearFn_THFTYPE_IiiI bnd_n2009_THFTYPE_i = | |
| 1494 | bnd_lYearFn_THFTYPE_IiiI bnd_n2009_THFTYPE_i) = | |
| 1495 | False \<or> | |
| 1496 | ((\<not> (\<not> bnd_likes_THFTYPE_IiioI bnd_lMary_THFTYPE_i | |
| 1497 | bnd_lBill_THFTYPE_i \<or> | |
| 1498 | \<not> bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i | |
| 1499 | bnd_lBill_THFTYPE_i)) = | |
| 1500 | True) = | |
| 1501 | False \<Longrightarrow> | |
| 1502 | ((\<not> (\<not> bnd_likes_THFTYPE_IiioI bnd_lMary_THFTYPE_i | |
| 1503 | bnd_lBill_THFTYPE_i \<or> | |
| 1504 | \<not> bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i | |
| 1505 | bnd_lBill_THFTYPE_i)) = | |
| 1506 | True) = | |
| 1507 | False" | |
| 1508 | (* apply (tactic {*break_hypotheses 1*}) *)
 | |
| 1509 | (* apply (erule extuni_bind) *) | |
| 1510 | (* apply (tactic {*clause_breaker 1*}) *)
 | |
| 1511 | (* done *) | |
| 1512 | by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "29") 1*})
 | |
| 1513 | ||
| 1514 |      (* (Annotated_step ("28", "extuni_dec"), *)
 | |
| 1515 | lemma "\<forall>SV2. (bnd_holdsDuring_THFTYPE_IiooI | |
| 1516 | (bnd_lYearFn_THFTYPE_IiiI bnd_n2009_THFTYPE_i) | |
| 1517 | (\<not> (\<not> bnd_likes_THFTYPE_IiioI | |
| 1518 | bnd_lMary_THFTYPE_i | |
| 1519 | bnd_lBill_THFTYPE_i \<or> | |
| 1520 | \<not> bnd_likes_THFTYPE_IiioI | |
| 1521 | bnd_lSue_THFTYPE_i | |
| 1522 | bnd_lBill_THFTYPE_i)) = | |
| 1523 | bnd_holdsDuring_THFTYPE_IiooI SV2 True) = | |
| 1524 | False \<Longrightarrow> | |
| 1525 | \<forall>SV2. (bnd_lYearFn_THFTYPE_IiiI bnd_n2009_THFTYPE_i = | |
| 1526 | SV2) = | |
| 1527 | False \<or> | |
| 1528 | ((\<not> (\<not> bnd_likes_THFTYPE_IiioI | |
| 1529 | bnd_lMary_THFTYPE_i bnd_lBill_THFTYPE_i \<or> | |
| 1530 | \<not> bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i | |
| 1531 | bnd_lBill_THFTYPE_i)) = | |
| 1532 | True) = | |
| 1533 | False" | |
| 1534 | (* apply (rule allI) *) | |
| 1535 | (* apply (erule_tac x = "SV2" in allE) *) | |
| 1536 | (* apply (erule extuni_dec_2) *) | |
| 1537 | (* done *) | |
| 1538 | by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "28") 1*})
 | |
| 1539 | *) | |
| 1540 | ||
| 1541 | (* QUA002^1 | |
| 1542 |    (* [[(Annotated_step ("49", "extuni_dec"), *)
 | |
| 1543 | lemma "((bnd_sK3_E = bnd_sK1_X1 \<or> bnd_sK3_E = bnd_sK2_X2) = | |
| 1544 | (bnd_sK3_E = bnd_sK2_X2 \<or> bnd_sK3_E = bnd_sK1_X1)) = | |
| 1545 | False \<Longrightarrow> | |
| 1546 | ((bnd_sK3_E = bnd_sK2_X2) = (bnd_sK3_E = bnd_sK2_X2)) = | |
| 1547 | False \<or> | |
| 1548 | ((bnd_sK3_E = bnd_sK1_X1) = (bnd_sK3_E = bnd_sK1_X1)) = | |
| 1549 | False" | |
| 1550 | by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "49") 1*})
 | |
| 1551 | ||
| 1552 |      (* (Annotated_step ("20", "unfold_def"), *)
 | |
| 1553 | lemma "(bnd_addition bnd_sK1_X1 bnd_sK2_X2 \<noteq> | |
| 1554 | bnd_addition bnd_sK2_X2 bnd_sK1_X1) = | |
| 1555 | True \<Longrightarrow> | |
| 1556 | (bnd_sup | |
| 61076 | 1557 | (\<lambda>SX0::TPTP_Interpret.ind. | 
| 55596 | 1558 | SX0 = bnd_sK1_X1 \<or> SX0 = bnd_sK2_X2) \<noteq> | 
| 1559 | bnd_sup | |
| 61076 | 1560 | (\<lambda>SX0::TPTP_Interpret.ind. | 
| 55596 | 1561 | SX0 = bnd_sK2_X2 \<or> SX0 = bnd_sK1_X1)) = | 
| 1562 | True" | |
| 1563 | by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "20") 1*})
 | |
| 1564 | *) | |
| 1565 | ||
| 1566 | (* | |
| 1567 | (*SEU620^2*) | |
| 1568 |      (* (Annotated_step ("11", "unfold_def"), *)
 | |
| 1569 | lemma "bnd_kpairiskpair = | |
| 1570 | (\<forall>Xx Xy. | |
| 1571 | bnd_iskpair | |
| 1572 | (bnd_setadjoin (bnd_setadjoin Xx bnd_emptyset) | |
| 1573 | (bnd_setadjoin | |
| 1574 | (bnd_setadjoin Xx | |
| 1575 | (bnd_setadjoin Xy bnd_emptyset)) | |
| 1576 | bnd_emptyset))) \<and> | |
| 1577 | bnd_kpair = | |
| 1578 | (\<lambda>Xx Xy. | |
| 1579 | bnd_setadjoin (bnd_setadjoin Xx bnd_emptyset) | |
| 1580 | (bnd_setadjoin | |
| 1581 | (bnd_setadjoin Xx | |
| 1582 | (bnd_setadjoin Xy bnd_emptyset)) | |
| 1583 | bnd_emptyset)) \<and> | |
| 1584 | bnd_iskpair = | |
| 1585 | (\<lambda>A. \<exists>Xx. bnd_in Xx (bnd_setunion A) \<and> | |
| 1586 | (\<exists>Xy. bnd_in Xy (bnd_setunion A) \<and> | |
| 1587 | A = | |
| 1588 | bnd_setadjoin | |
| 1589 | (bnd_setadjoin Xx bnd_emptyset) | |
| 1590 | (bnd_setadjoin | |
| 1591 | (bnd_setadjoin Xx | |
| 1592 | (bnd_setadjoin Xy bnd_emptyset)) | |
| 1593 | bnd_emptyset))) \<and> | |
| 1594 | (\<forall>SY5 SY6. | |
| 1595 | (bnd_setadjoin (bnd_setadjoin SY5 bnd_emptyset) | |
| 1596 | (bnd_setadjoin | |
| 1597 | (bnd_setadjoin SY5 | |
| 1598 | (bnd_setadjoin SY6 bnd_emptyset)) | |
| 1599 | bnd_emptyset) = | |
| 1600 | bnd_setadjoin | |
| 1601 | (bnd_setadjoin (bnd_sK3 SY6 SY5) bnd_emptyset) | |
| 1602 | (bnd_setadjoin | |
| 1603 | (bnd_setadjoin (bnd_sK3 SY6 SY5) | |
| 1604 | (bnd_setadjoin (bnd_sK4 SY6 SY5) | |
| 1605 | bnd_emptyset)) | |
| 1606 | bnd_emptyset) \<and> | |
| 1607 | bnd_in (bnd_sK4 SY6 SY5) | |
| 1608 | (bnd_setunion | |
| 1609 | (bnd_setadjoin (bnd_setadjoin SY5 bnd_emptyset) | |
| 1610 | (bnd_setadjoin | |
| 1611 | (bnd_setadjoin SY5 | |
| 1612 | (bnd_setadjoin SY6 bnd_emptyset)) | |
| 1613 | bnd_emptyset)))) \<and> | |
| 1614 | bnd_in (bnd_sK3 SY6 SY5) | |
| 1615 | (bnd_setunion | |
| 1616 | (bnd_setadjoin (bnd_setadjoin SY5 bnd_emptyset) | |
| 1617 | (bnd_setadjoin | |
| 1618 | (bnd_setadjoin SY5 | |
| 1619 | (bnd_setadjoin SY6 bnd_emptyset)) | |
| 1620 | bnd_emptyset)))) = | |
| 1621 | True \<Longrightarrow> | |
| 1622 | (\<forall>SX0 SX1. | |
| 1623 | \<not> (\<not> \<not> (bnd_setadjoin | |
| 1624 | (bnd_setadjoin SX0 bnd_emptyset) | |
| 1625 | (bnd_setadjoin | |
| 1626 | (bnd_setadjoin SX0 | |
| 1627 | (bnd_setadjoin SX1 bnd_emptyset)) | |
| 1628 | bnd_emptyset) \<noteq> | |
| 1629 | bnd_setadjoin | |
| 1630 | (bnd_setadjoin (bnd_sK3 SX1 SX0) | |
| 1631 | bnd_emptyset) | |
| 1632 | (bnd_setadjoin | |
| 1633 | (bnd_setadjoin (bnd_sK3 SX1 SX0) | |
| 1634 | (bnd_setadjoin (bnd_sK4 SX1 SX0) | |
| 1635 | bnd_emptyset)) | |
| 1636 | bnd_emptyset) \<or> | |
| 1637 | \<not> bnd_in (bnd_sK4 SX1 SX0) | |
| 1638 | (bnd_setunion | |
| 1639 | (bnd_setadjoin | |
| 1640 | (bnd_setadjoin SX0 bnd_emptyset) | |
| 1641 | (bnd_setadjoin | |
| 1642 | (bnd_setadjoin SX0 | |
| 1643 | (bnd_setadjoin SX1 bnd_emptyset)) | |
| 1644 | bnd_emptyset)))) \<or> | |
| 1645 | \<not> bnd_in (bnd_sK3 SX1 SX0) | |
| 1646 | (bnd_setunion | |
| 1647 | (bnd_setadjoin | |
| 1648 | (bnd_setadjoin SX0 bnd_emptyset) | |
| 1649 | (bnd_setadjoin | |
| 1650 | (bnd_setadjoin SX0 | |
| 1651 | (bnd_setadjoin SX1 bnd_emptyset)) | |
| 1652 | bnd_emptyset))))) = | |
| 1653 | True" | |
| 1654 | by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "11") 1*})
 | |
| 1655 | ||
| 1656 |      (* (Annotated_step ("3", "unfold_def"), *)
 | |
| 1657 | lemma "bnd_kpairiskpair = | |
| 1658 | (\<forall>Xx Xy. | |
| 1659 | bnd_iskpair | |
| 1660 | (bnd_setadjoin (bnd_setadjoin Xx bnd_emptyset) | |
| 1661 | (bnd_setadjoin | |
| 1662 | (bnd_setadjoin Xx | |
| 1663 | (bnd_setadjoin Xy bnd_emptyset)) | |
| 1664 | bnd_emptyset))) \<and> | |
| 1665 | bnd_kpair = | |
| 1666 | (\<lambda>Xx Xy. | |
| 1667 | bnd_setadjoin (bnd_setadjoin Xx bnd_emptyset) | |
| 1668 | (bnd_setadjoin | |
| 1669 | (bnd_setadjoin Xx | |
| 1670 | (bnd_setadjoin Xy bnd_emptyset)) | |
| 1671 | bnd_emptyset)) \<and> | |
| 1672 | bnd_iskpair = | |
| 1673 | (\<lambda>A. \<exists>Xx. bnd_in Xx (bnd_setunion A) \<and> | |
| 1674 | (\<exists>Xy. bnd_in Xy (bnd_setunion A) \<and> | |
| 1675 | A = | |
| 1676 | bnd_setadjoin | |
| 1677 | (bnd_setadjoin Xx bnd_emptyset) | |
| 1678 | (bnd_setadjoin | |
| 1679 | (bnd_setadjoin Xx | |
| 1680 | (bnd_setadjoin Xy bnd_emptyset)) | |
| 1681 | bnd_emptyset))) \<and> | |
| 1682 | (bnd_kpairiskpair \<longrightarrow> | |
| 1683 | (\<forall>Xx Xy. bnd_iskpair (bnd_kpair Xx Xy))) = | |
| 1684 | False \<Longrightarrow> | |
| 1685 | ((\<forall>SY5 SY6. | |
| 1686 | \<exists>SY7. bnd_in SY7 | |
| 1687 | (bnd_setunion | |
| 1688 | (bnd_setadjoin | |
| 1689 | (bnd_setadjoin SY5 bnd_emptyset) | |
| 1690 | (bnd_setadjoin | |
| 1691 | (bnd_setadjoin SY5 | |
| 1692 | (bnd_setadjoin SY6 bnd_emptyset)) | |
| 1693 | bnd_emptyset))) \<and> | |
| 1694 | (\<exists>SY8. bnd_in SY8 | |
| 1695 | (bnd_setunion | |
| 1696 | (bnd_setadjoin | |
| 1697 | (bnd_setadjoin SY5 bnd_emptyset) | |
| 1698 | (bnd_setadjoin | |
| 1699 | (bnd_setadjoin SY5 (bnd_setadjoin SY6 bnd_emptyset)) | |
| 1700 | bnd_emptyset))) \<and> | |
| 1701 | bnd_setadjoin | |
| 1702 | (bnd_setadjoin SY5 bnd_emptyset) | |
| 1703 | (bnd_setadjoin | |
| 1704 | (bnd_setadjoin SY5 | |
| 1705 | (bnd_setadjoin SY6 bnd_emptyset)) | |
| 1706 | bnd_emptyset) = | |
| 1707 | bnd_setadjoin | |
| 1708 | (bnd_setadjoin SY7 bnd_emptyset) | |
| 1709 | (bnd_setadjoin | |
| 1710 | (bnd_setadjoin SY7 | |
| 1711 | (bnd_setadjoin SY8 bnd_emptyset)) | |
| 1712 | bnd_emptyset))) \<longrightarrow> | |
| 1713 | (\<forall>SY0 SY1. | |
| 1714 | \<exists>SY3. bnd_in SY3 | |
| 1715 | (bnd_setunion | |
| 1716 | (bnd_setadjoin | |
| 1717 | (bnd_setadjoin SY0 bnd_emptyset) | |
| 1718 | (bnd_setadjoin | |
| 1719 | (bnd_setadjoin SY0 | |
| 1720 | (bnd_setadjoin SY1 bnd_emptyset)) | |
| 1721 | bnd_emptyset))) \<and> | |
| 1722 | (\<exists>SY4. bnd_in SY4 | |
| 1723 | (bnd_setunion | |
| 1724 | (bnd_setadjoin | |
| 1725 | (bnd_setadjoin SY0 bnd_emptyset) | |
| 1726 | (bnd_setadjoin | |
| 1727 | (bnd_setadjoin SY0 (bnd_setadjoin SY1 bnd_emptyset)) | |
| 1728 | bnd_emptyset))) \<and> | |
| 1729 | bnd_setadjoin | |
| 1730 | (bnd_setadjoin SY0 bnd_emptyset) | |
| 1731 | (bnd_setadjoin | |
| 1732 | (bnd_setadjoin SY0 | |
| 1733 | (bnd_setadjoin SY1 bnd_emptyset)) | |
| 1734 | bnd_emptyset) = | |
| 1735 | bnd_setadjoin | |
| 1736 | (bnd_setadjoin SY3 bnd_emptyset) | |
| 1737 | (bnd_setadjoin | |
| 1738 | (bnd_setadjoin SY3 | |
| 1739 | (bnd_setadjoin SY4 bnd_emptyset)) | |
| 1740 | bnd_emptyset)))) = | |
| 1741 | False" | |
| 1742 | by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "3") 1*})
 | |
| 1743 | ||
| 1744 |      (* (Annotated_step ("8", "extcnf_combined"), *)
 | |
| 1745 | lemma "(\<forall>SY5 SY6. | |
| 1746 | \<exists>SY7. bnd_in SY7 | |
| 1747 | (bnd_setunion | |
| 1748 | (bnd_setadjoin | |
| 1749 | (bnd_setadjoin SY5 bnd_emptyset) | |
| 1750 | (bnd_setadjoin | |
| 1751 | (bnd_setadjoin SY5 | |
| 1752 | (bnd_setadjoin SY6 bnd_emptyset)) | |
| 1753 | bnd_emptyset))) \<and> | |
| 1754 | (\<exists>SY8. bnd_in SY8 | |
| 1755 | (bnd_setunion | |
| 1756 | (bnd_setadjoin | |
| 1757 | (bnd_setadjoin SY5 bnd_emptyset) | |
| 1758 | (bnd_setadjoin | |
| 1759 | (bnd_setadjoin SY5 (bnd_setadjoin SY6 bnd_emptyset)) | |
| 1760 | bnd_emptyset))) \<and> | |
| 1761 | bnd_setadjoin | |
| 1762 | (bnd_setadjoin SY5 bnd_emptyset) | |
| 1763 | (bnd_setadjoin | |
| 1764 | (bnd_setadjoin SY5 | |
| 1765 | (bnd_setadjoin SY6 bnd_emptyset)) | |
| 1766 | bnd_emptyset) = | |
| 1767 | bnd_setadjoin | |
| 1768 | (bnd_setadjoin SY7 bnd_emptyset) | |
| 1769 | (bnd_setadjoin | |
| 1770 | (bnd_setadjoin SY7 | |
| 1771 | (bnd_setadjoin SY8 bnd_emptyset)) | |
| 1772 | bnd_emptyset))) = | |
| 1773 | True \<Longrightarrow> | |
| 1774 | (\<forall>SY5 SY6. | |
| 1775 | (bnd_setadjoin (bnd_setadjoin SY5 bnd_emptyset) | |
| 1776 | (bnd_setadjoin | |
| 1777 | (bnd_setadjoin SY5 | |
| 1778 | (bnd_setadjoin SY6 bnd_emptyset)) | |
| 1779 | bnd_emptyset) = | |
| 1780 | bnd_setadjoin | |
| 1781 | (bnd_setadjoin (bnd_sK3 SY6 SY5) bnd_emptyset) | |
| 1782 | (bnd_setadjoin | |
| 1783 | (bnd_setadjoin (bnd_sK3 SY6 SY5) | |
| 1784 | (bnd_setadjoin (bnd_sK4 SY6 SY5) | |
| 1785 | bnd_emptyset)) | |
| 1786 | bnd_emptyset) \<and> | |
| 1787 | bnd_in (bnd_sK4 SY6 SY5) | |
| 1788 | (bnd_setunion | |
| 1789 | (bnd_setadjoin (bnd_setadjoin SY5 bnd_emptyset) | |
| 1790 | (bnd_setadjoin | |
| 1791 | (bnd_setadjoin SY5 | |
| 1792 | (bnd_setadjoin SY6 bnd_emptyset)) | |
| 1793 | bnd_emptyset)))) \<and> | |
| 1794 | bnd_in (bnd_sK3 SY6 SY5) | |
| 1795 | (bnd_setunion | |
| 1796 | (bnd_setadjoin (bnd_setadjoin SY5 bnd_emptyset) | |
| 1797 | (bnd_setadjoin | |
| 1798 | (bnd_setadjoin SY5 | |
| 1799 | (bnd_setadjoin SY6 bnd_emptyset)) | |
| 1800 | bnd_emptyset)))) = | |
| 1801 | True" | |
| 1802 | by (tactic {*
 | |
| 1803 | HEADGOAL (extcnf_combined_tac Full false (hd prob_names)) | |
| 1804 | *}) | |
| 1805 | ||
| 1806 |      (* (Annotated_step ("7", "extcnf_combined"), *)
 | |
| 1807 | lemma "(\<not> (\<forall>SY0 SY1. | |
| 1808 | \<exists>SY3. bnd_in SY3 | |
| 1809 | (bnd_setunion | |
| 1810 | (bnd_setadjoin | |
| 1811 | (bnd_setadjoin SY0 bnd_emptyset) | |
| 1812 | (bnd_setadjoin | |
| 1813 | (bnd_setadjoin SY0 | |
| 1814 | (bnd_setadjoin SY1 bnd_emptyset)) | |
| 1815 | bnd_emptyset))) \<and> | |
| 1816 | (\<exists>SY4. bnd_in SY4 | |
| 1817 | (bnd_setunion | |
| 1818 | (bnd_setadjoin (bnd_setadjoin SY0 bnd_emptyset) | |
| 1819 | (bnd_setadjoin | |
| 1820 | (bnd_setadjoin SY0 (bnd_setadjoin SY1 bnd_emptyset)) | |
| 1821 | bnd_emptyset))) \<and> | |
| 1822 | bnd_setadjoin | |
| 1823 | (bnd_setadjoin SY0 bnd_emptyset) | |
| 1824 | (bnd_setadjoin | |
| 1825 | (bnd_setadjoin SY0 (bnd_setadjoin SY1 bnd_emptyset)) | |
| 1826 | bnd_emptyset) = | |
| 1827 | bnd_setadjoin | |
| 1828 | (bnd_setadjoin SY3 bnd_emptyset) | |
| 1829 | (bnd_setadjoin | |
| 1830 | (bnd_setadjoin SY3 (bnd_setadjoin SY4 bnd_emptyset)) | |
| 1831 | bnd_emptyset)))) = | |
| 1832 | True \<Longrightarrow> | |
| 1833 | (\<forall>SY24. | |
| 1834 | (\<forall>SY25. | |
| 1835 | bnd_setadjoin | |
| 1836 | (bnd_setadjoin bnd_sK1 bnd_emptyset) | |
| 1837 | (bnd_setadjoin | |
| 1838 | (bnd_setadjoin bnd_sK1 | |
| 1839 | (bnd_setadjoin bnd_sK2 bnd_emptyset)) | |
| 1840 | bnd_emptyset) \<noteq> | |
| 1841 | bnd_setadjoin (bnd_setadjoin SY24 bnd_emptyset) | |
| 1842 | (bnd_setadjoin | |
| 1843 | (bnd_setadjoin SY24 | |
| 1844 | (bnd_setadjoin SY25 bnd_emptyset)) | |
| 1845 | bnd_emptyset) \<or> | |
| 1846 | \<not> bnd_in SY25 | |
| 1847 | (bnd_setunion | |
| 1848 | (bnd_setadjoin | |
| 1849 | (bnd_setadjoin bnd_sK1 bnd_emptyset) | |
| 1850 | (bnd_setadjoin | |
| 1851 | (bnd_setadjoin bnd_sK1 | |
| 1852 | (bnd_setadjoin bnd_sK2 | |
| 1853 | bnd_emptyset)) | |
| 1854 | bnd_emptyset)))) \<or> | |
| 1855 | \<not> bnd_in SY24 | |
| 1856 | (bnd_setunion | |
| 1857 | (bnd_setadjoin | |
| 1858 | (bnd_setadjoin bnd_sK1 bnd_emptyset) | |
| 1859 | (bnd_setadjoin | |
| 1860 | (bnd_setadjoin bnd_sK1 | |
| 1861 | (bnd_setadjoin bnd_sK2 bnd_emptyset)) | |
| 1862 | bnd_emptyset)))) = | |
| 1863 | True" | |
| 1864 | by (tactic {*HEADGOAL (extcnf_combined_tac Full false (hd prob_names))*})
 | |
| 1865 | *) | |
| 1866 | ||
| 1867 | (*PUZ081^2*) | |
| 1868 | (* | |
| 1869 |      (* (Annotated_step ("9", "unfold_def"), *)
 | |
| 1870 | lemma "bnd_says bnd_mel | |
| 1871 | (\<not> bnd_knave bnd_zoey \<and> \<not> bnd_knave bnd_mel) \<Longrightarrow> | |
| 1872 | bnd_says bnd_mel | |
| 1873 | (\<not> bnd_knave bnd_zoey \<and> \<not> bnd_knave bnd_mel) = | |
| 1874 | True" | |
| 1875 | by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "9") 1*})
 | |
| 1876 | ||
| 1877 |      (* (Annotated_step ("10", "unfold_def"), *)
 | |
| 1878 | lemma "bnd_says bnd_zoey (bnd_knave bnd_mel) \<Longrightarrow> | |
| 1879 | bnd_says bnd_zoey (bnd_knave bnd_mel) = True" | |
| 1880 | by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "10") 1*})
 | |
| 1881 | ||
| 1882 |      (* (Annotated_step ("11", "unfold_def"), *)
 | |
| 1883 | lemma "\<forall>P S. bnd_knave P \<and> bnd_says P S \<longrightarrow> \<not> S \<Longrightarrow> | |
| 1884 | (\<forall>P S. bnd_knave P \<and> bnd_says P S \<longrightarrow> \<not> S) = True" | |
| 1885 | by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "11") 1*})
 | |
| 1886 | ||
| 1887 |      (* (Annotated_step ("12", "unfold_def"), *)
 | |
| 1888 | lemma "\<forall>P S. bnd_knight P \<and> bnd_says P S \<longrightarrow> S \<Longrightarrow> | |
| 1889 | (\<forall>P S. bnd_knight P \<and> bnd_says P S \<longrightarrow> S) = True" | |
| 1890 | by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "12") 1*})
 | |
| 1891 | ||
| 1892 |      (* (Annotated_step ("13", "unfold_def"), *)
 | |
| 1893 | lemma "\<forall>P. bnd_knight P \<noteq> bnd_knave P \<Longrightarrow> | |
| 1894 | (\<forall>P. bnd_knight P \<noteq> bnd_knave P) = True" | |
| 1895 | by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "13") 1*})
 | |
| 1896 | ||
| 1897 |      (* (Annotated_step ("15", "extcnf_combined"), *)
 | |
| 1898 | lemma "(\<not> (\<exists>TZ TM. TZ bnd_zoey \<and> TM bnd_mel)) = True \<Longrightarrow> | |
| 1899 | ((\<forall>TM. \<not> TM bnd_mel) \<or> (\<forall>TZ. \<not> TZ bnd_zoey)) = True" | |
| 1900 | by (tactic {*extcnf_combined_tac Full false (hd prob_names) 1*})
 | |
| 1901 | ||
| 1902 |      (* (Annotated_step ("18", "extcnf_combined"), *)
 | |
| 1903 | lemma "(\<forall>P. bnd_knight P \<noteq> bnd_knave P) = True \<Longrightarrow> | |
| 1904 | ((\<forall>P. \<not> bnd_knave P \<or> \<not> bnd_knight P) \<and> | |
| 1905 | (\<forall>P. bnd_knave P \<or> bnd_knight P)) = | |
| 1906 | True" | |
| 1907 | by (tactic {*extcnf_combined_tac Full false (hd prob_names) 1*})
 | |
| 1908 | *) | |
| 1909 | ||
| 1910 | (* | |
| 1911 | (*from SEU667^2.p.out*) | |
| 1912 |      (* (Annotated_step ("10", "unfold_def"), *)
 | |
| 1913 | lemma "bnd_dpsetconstrSub = | |
| 1914 | (\<forall>A B Xphi. | |
| 1915 | bnd_subset (bnd_dpsetconstr A B Xphi) | |
| 1916 | (bnd_cartprod A B)) \<and> | |
| 1917 | bnd_dpsetconstr = | |
| 1918 | (\<lambda>A B Xphi. | |
| 1919 | bnd_dsetconstr (bnd_cartprod A B) | |
| 1920 | (\<lambda>Xu. \<exists>Xx. bnd_in Xx A \<and> | |
| 1921 | (\<exists>Xy. (bnd_in Xy B \<and> Xphi Xx Xy) \<and> | |
| 1922 | Xu = bnd_kpair Xx Xy))) \<and> | |
| 1923 | bnd_breln = (\<lambda>A B C. bnd_subset C (bnd_cartprod A B)) \<and> | |
| 1924 | (\<not> bnd_subset | |
| 1925 | (bnd_dsetconstr (bnd_cartprod bnd_sK1 bnd_sK2) | |
| 1926 | (\<lambda>SY66. | |
| 1927 | \<exists>SY67. | |
| 1928 | bnd_in SY67 bnd_sK1 \<and> | |
| 1929 | (\<exists>SY68. | |
| 1930 | (bnd_in SY68 bnd_sK2 \<and> | |
| 1931 | bnd_sK3 SY67 SY68) \<and> | |
| 1932 | SY66 = bnd_kpair SY67 SY68))) | |
| 1933 | (bnd_cartprod bnd_sK1 bnd_sK2)) = | |
| 1934 | True \<Longrightarrow> | |
| 1935 | (\<not> bnd_subset | |
| 1936 | (bnd_dsetconstr (bnd_cartprod bnd_sK1 bnd_sK2) | |
| 1937 | (\<lambda>SX0. \<not> (\<forall>SX1. \<not> \<not> (\<not> bnd_in SX1 bnd_sK1 \<or> | |
| 1938 | \<not> \<not> (\<forall>SX2. \<not> \<not> (\<not> \<not> (\<not> bnd_in SX2 bnd_sK2 \<or> | |
| 1939 | \<not> bnd_sK3 SX1 SX2) \<or> | |
| 1940 | SX0 \<noteq> bnd_kpair SX1 SX2)))))) | |
| 1941 | (bnd_cartprod bnd_sK1 bnd_sK2)) = | |
| 1942 | True" | |
| 1943 | by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "10") 1*})
 | |
| 1944 | ||
| 1945 | ||
| 1946 |      (* (Annotated_step ("11", "unfold_def"), *)
 | |
| 1947 | lemma "bnd_dpsetconstrSub = | |
| 1948 | (\<forall>A B Xphi. | |
| 1949 | bnd_subset (bnd_dpsetconstr A B Xphi) | |
| 1950 | (bnd_cartprod A B)) \<and> | |
| 1951 | bnd_dpsetconstr = | |
| 1952 | (\<lambda>A B Xphi. | |
| 1953 | bnd_dsetconstr (bnd_cartprod A B) | |
| 1954 | (\<lambda>Xu. \<exists>Xx. bnd_in Xx A \<and> | |
| 1955 | (\<exists>Xy. (bnd_in Xy B \<and> Xphi Xx Xy) \<and> | |
| 1956 | Xu = bnd_kpair Xx Xy))) \<and> | |
| 1957 | bnd_breln = (\<lambda>A B C. bnd_subset C (bnd_cartprod A B)) \<and> | |
| 1958 | (\<forall>SY21 SY22 SY23. | |
| 1959 | bnd_subset | |
| 1960 | (bnd_dsetconstr (bnd_cartprod SY21 SY22) | |
| 1961 | (\<lambda>SY35. | |
| 1962 | \<exists>SY36. | |
| 1963 | bnd_in SY36 SY21 \<and> | |
| 1964 | (\<exists>SY37. | |
| 1965 | (bnd_in SY37 SY22 \<and> SY23 SY36 SY37) \<and> | |
| 1966 | SY35 = bnd_kpair SY36 SY37))) | |
| 1967 | (bnd_cartprod SY21 SY22)) = | |
| 1968 | True \<Longrightarrow> | |
| 1969 | (\<forall>SX0 SX1 SX2. | |
| 1970 | bnd_subset | |
| 1971 | (bnd_dsetconstr (bnd_cartprod SX0 SX1) | |
| 1972 | (\<lambda>SX3. \<not> (\<forall>SX4. \<not> \<not> (\<not> bnd_in SX4 SX0 \<or> | |
| 1973 | \<not> \<not> (\<forall>SX5. \<not> \<not> (\<not> \<not> (\<not> bnd_in SX5 SX1 \<or> \<not> SX2 SX4 SX5) \<or> | |
| 1974 | SX3 \<noteq> bnd_kpair SX4 SX5)))))) | |
| 1975 | (bnd_cartprod SX0 SX1)) = | |
| 1976 | True" | |
| 1977 | by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "11") 1*})
 | |
| 1978 | *) | |
| 1979 | ||
| 1980 | (* | |
| 1981 | (*from ALG001^5*) | |
| 1982 |       (* (Annotated_step ("4", "extcnf_forall_neg"), *)
 | |
| 1983 | lemma "(\<forall>(Xh1 :: bnd_g \<Rightarrow> bnd_b) (Xh2 :: bnd_b \<Rightarrow> bnd_a) (Xf1 :: bnd_g \<Rightarrow> bnd_g \<Rightarrow> bnd_g) (Xf2 :: bnd_b \<Rightarrow> bnd_b \<Rightarrow> bnd_b) (Xf3 :: bnd_a \<Rightarrow> bnd_a \<Rightarrow> bnd_a). | |
| 1984 | (\<forall>Xx Xy. Xh1 (Xf1 Xx Xy) = Xf2 (Xh1 Xx) (Xh1 Xy)) \<and> | |
| 1985 | (\<forall>Xx Xy. | |
| 1986 | Xh2 (Xf2 Xx Xy) = Xf3 (Xh2 Xx) (Xh2 Xy)) \<longrightarrow> | |
| 1987 | (\<forall>Xx Xy. | |
| 1988 | Xh2 (Xh1 (Xf1 Xx Xy)) = | |
| 1989 | Xf3 (Xh2 (Xh1 Xx)) (Xh2 (Xh1 Xy)))) = | |
| 1990 | False \<Longrightarrow> | |
| 1991 | ((\<forall>SY35 SY36. | |
| 1992 | bnd_sK1 (bnd_sK3 SY35 SY36) = | |
| 1993 | bnd_sK4 (bnd_sK1 SY35) (bnd_sK1 SY36)) \<and> | |
| 1994 | (\<forall>SY31 SY32. | |
| 1995 | bnd_sK2 (bnd_sK4 SY31 SY32) = | |
| 1996 | bnd_sK5 (bnd_sK2 SY31) (bnd_sK2 SY32)) \<longrightarrow> | |
| 1997 | (\<forall>SY39 SY40. | |
| 1998 | bnd_sK2 (bnd_sK1 (bnd_sK3 SY39 SY40)) = | |
| 1999 | bnd_sK5 (bnd_sK2 (bnd_sK1 SY39)) | |
| 2000 | (bnd_sK2 (bnd_sK1 SY40)))) = | |
| 2001 | False" | |
| 2002 | by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "4") 1*})
 | |
| 2003 | *) | |
| 2004 | ||
| 2005 | (*SYN044^4*) | |
| 2006 | (* | |
| 56281 | 2007 | declare [[ML_print_depth = 1400]] | 
| 55596 | 2008 | (* the_tactics *) | 
| 2009 | *} | |
| 2010 | ||
| 2011 |      (* (Annotated_step ("12", "unfold_def"), *)
 | |
| 2012 | lemma "bnd_mor = | |
| 61076 | 2013 | (\<lambda>(X::TPTP_Interpret.ind \<Rightarrow> bool) | 
| 2014 | (Y::TPTP_Interpret.ind \<Rightarrow> bool) U::TPTP_Interpret.ind. | |
| 55596 | 2015 | X U \<or> Y U) \<and> | 
| 2016 | bnd_mnot = | |
| 61076 | 2017 | (\<lambda>(X::TPTP_Interpret.ind \<Rightarrow> bool) U::TPTP_Interpret.ind. | 
| 55596 | 2018 | \<not> X U) \<and> | 
| 2019 | bnd_mimplies = | |
| 61076 | 2020 | (\<lambda>U::TPTP_Interpret.ind \<Rightarrow> bool. bnd_mor (bnd_mnot U)) \<and> | 
| 55596 | 2021 | bnd_mbox_s4 = | 
| 61076 | 2022 | (\<lambda>(P::TPTP_Interpret.ind \<Rightarrow> bool) X::TPTP_Interpret.ind. | 
| 2023 | \<forall>Y::TPTP_Interpret.ind. bnd_irel X Y \<longrightarrow> P Y) \<and> | |
| 55596 | 2024 | bnd_mand = | 
| 61076 | 2025 | (\<lambda>(X::TPTP_Interpret.ind \<Rightarrow> bool) | 
| 2026 | (Y::TPTP_Interpret.ind \<Rightarrow> bool) U::TPTP_Interpret.ind. | |
| 55596 | 2027 | X U \<and> Y U) \<and> | 
| 2028 | bnd_ixor = | |
| 61076 | 2029 | (\<lambda>(P::TPTP_Interpret.ind \<Rightarrow> bool) | 
| 2030 | Q::TPTP_Interpret.ind \<Rightarrow> bool. | |
| 55596 | 2031 | bnd_inot (bnd_iequiv P Q)) \<and> | 
| 2032 | bnd_ivalid = All \<and> | |
| 61076 | 2033 | bnd_itrue = (\<lambda>W::TPTP_Interpret.ind. True) \<and> | 
| 55596 | 2034 | bnd_isatisfiable = Ex \<and> | 
| 2035 | bnd_ior = | |
| 61076 | 2036 | (\<lambda>(P::TPTP_Interpret.ind \<Rightarrow> bool) | 
| 2037 | Q::TPTP_Interpret.ind \<Rightarrow> bool. | |
| 55596 | 2038 | bnd_mor (bnd_mbox_s4 P) (bnd_mbox_s4 Q)) \<and> | 
| 2039 | bnd_inot = | |
| 61076 | 2040 | (\<lambda>P::TPTP_Interpret.ind \<Rightarrow> bool. | 
| 55596 | 2041 | bnd_mnot (bnd_mbox_s4 P)) \<and> | 
| 2042 | bnd_iinvalid = | |
| 61076 | 2043 | (\<lambda>Phi::TPTP_Interpret.ind \<Rightarrow> bool. | 
| 2044 | \<forall>W::TPTP_Interpret.ind. \<not> Phi W) \<and> | |
| 55596 | 2045 | bnd_iimplies = | 
| 61076 | 2046 | (\<lambda>(P::TPTP_Interpret.ind \<Rightarrow> bool) | 
| 2047 | Q::TPTP_Interpret.ind \<Rightarrow> bool. | |
| 55596 | 2048 | bnd_mimplies (bnd_mbox_s4 P) (bnd_mbox_s4 Q)) \<and> | 
| 2049 | bnd_iimplied = | |
| 61076 | 2050 | (\<lambda>(P::TPTP_Interpret.ind \<Rightarrow> bool) | 
| 2051 | Q::TPTP_Interpret.ind \<Rightarrow> bool. bnd_iimplies Q P) \<and> | |
| 55596 | 2052 | bnd_ifalse = bnd_inot bnd_itrue \<and> | 
| 2053 | bnd_iequiv = | |
| 61076 | 2054 | (\<lambda>(P::TPTP_Interpret.ind \<Rightarrow> bool) | 
| 2055 | Q::TPTP_Interpret.ind \<Rightarrow> bool. | |
| 55596 | 2056 | bnd_iand (bnd_iimplies P Q) (bnd_iimplies Q P)) \<and> | 
| 2057 | bnd_icountersatisfiable = | |
| 61076 | 2058 | (\<lambda>Phi::TPTP_Interpret.ind \<Rightarrow> bool. | 
| 2059 | \<exists>W::TPTP_Interpret.ind. \<not> Phi W) \<and> | |
| 2060 | bnd_iatom = (\<lambda>P::TPTP_Interpret.ind \<Rightarrow> bool. P) \<and> | |
| 55596 | 2061 | bnd_iand = bnd_mand \<and> | 
| 61076 | 2062 | (\<forall>(X::TPTP_Interpret.ind) (Y::TPTP_Interpret.ind) | 
| 2063 | Z::TPTP_Interpret.ind. | |
| 55596 | 2064 | bnd_irel X Y \<and> bnd_irel Y Z \<longrightarrow> bnd_irel X Z) \<Longrightarrow> | 
| 61076 | 2065 | (\<forall>(X::TPTP_Interpret.ind) (Y::TPTP_Interpret.ind) | 
| 2066 | Z::TPTP_Interpret.ind. | |
| 55596 | 2067 | bnd_irel X Y \<and> bnd_irel Y Z \<longrightarrow> bnd_irel X Z) = | 
| 2068 | True" | |
| 2069 | (* by (tactic {*tectoc @{context}*}) *)
 | |
| 2070 | by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "12") 1*})
 | |
| 2071 | ||
| 2072 |      (* (Annotated_step ("11", "unfold_def"), *)
 | |
| 2073 | lemma "bnd_mor = | |
| 61076 | 2074 | (\<lambda>(X::TPTP_Interpret.ind \<Rightarrow> bool) | 
| 2075 | (Y::TPTP_Interpret.ind \<Rightarrow> bool) U::TPTP_Interpret.ind. | |
| 55596 | 2076 | X U \<or> Y U) \<and> | 
| 2077 | bnd_mnot = | |
| 61076 | 2078 | (\<lambda>(X::TPTP_Interpret.ind \<Rightarrow> bool) U::TPTP_Interpret.ind. | 
| 55596 | 2079 | \<not> X U) \<and> | 
| 2080 | bnd_mimplies = | |
| 61076 | 2081 | (\<lambda>U::TPTP_Interpret.ind \<Rightarrow> bool. bnd_mor (bnd_mnot U)) \<and> | 
| 55596 | 2082 | bnd_mbox_s4 = | 
| 61076 | 2083 | (\<lambda>(P::TPTP_Interpret.ind \<Rightarrow> bool) X::TPTP_Interpret.ind. | 
| 2084 | \<forall>Y::TPTP_Interpret.ind. bnd_irel X Y \<longrightarrow> P Y) \<and> | |
| 55596 | 2085 | bnd_mand = | 
| 61076 | 2086 | (\<lambda>(X::TPTP_Interpret.ind \<Rightarrow> bool) | 
| 2087 | (Y::TPTP_Interpret.ind \<Rightarrow> bool) U::TPTP_Interpret.ind. | |
| 55596 | 2088 | X U \<and> Y U) \<and> | 
| 2089 | bnd_ixor = | |
| 61076 | 2090 | (\<lambda>(P::TPTP_Interpret.ind \<Rightarrow> bool) | 
| 2091 | Q::TPTP_Interpret.ind \<Rightarrow> bool. | |
| 55596 | 2092 | bnd_inot (bnd_iequiv P Q)) \<and> | 
| 2093 | bnd_ivalid = All \<and> | |
| 61076 | 2094 | bnd_itrue = (\<lambda>W::TPTP_Interpret.ind. True) \<and> | 
| 55596 | 2095 | bnd_isatisfiable = Ex \<and> | 
| 2096 | bnd_ior = | |
| 61076 | 2097 | (\<lambda>(P::TPTP_Interpret.ind \<Rightarrow> bool) | 
| 2098 | Q::TPTP_Interpret.ind \<Rightarrow> bool. | |
| 55596 | 2099 | bnd_mor (bnd_mbox_s4 P) (bnd_mbox_s4 Q)) \<and> | 
| 2100 | bnd_inot = | |
| 61076 | 2101 | (\<lambda>P::TPTP_Interpret.ind \<Rightarrow> bool. | 
| 55596 | 2102 | bnd_mnot (bnd_mbox_s4 P)) \<and> | 
| 2103 | bnd_iinvalid = | |
| 61076 | 2104 | (\<lambda>Phi::TPTP_Interpret.ind \<Rightarrow> bool. | 
| 2105 | \<forall>W::TPTP_Interpret.ind. \<not> Phi W) \<and> | |
| 55596 | 2106 | bnd_iimplies = | 
| 61076 | 2107 | (\<lambda>(P::TPTP_Interpret.ind \<Rightarrow> bool) | 
| 2108 | Q::TPTP_Interpret.ind \<Rightarrow> bool. | |
| 55596 | 2109 | bnd_mimplies (bnd_mbox_s4 P) (bnd_mbox_s4 Q)) \<and> | 
| 2110 | bnd_iimplied = | |
| 61076 | 2111 | (\<lambda>(P::TPTP_Interpret.ind \<Rightarrow> bool) | 
| 2112 | Q::TPTP_Interpret.ind \<Rightarrow> bool. bnd_iimplies Q P) \<and> | |
| 55596 | 2113 | bnd_ifalse = bnd_inot bnd_itrue \<and> | 
| 2114 | bnd_iequiv = | |
| 61076 | 2115 | (\<lambda>(P::TPTP_Interpret.ind \<Rightarrow> bool) | 
| 2116 | Q::TPTP_Interpret.ind \<Rightarrow> bool. | |
| 55596 | 2117 | bnd_iand (bnd_iimplies P Q) (bnd_iimplies Q P)) \<and> | 
| 2118 | bnd_icountersatisfiable = | |
| 61076 | 2119 | (\<lambda>Phi::TPTP_Interpret.ind \<Rightarrow> bool. | 
| 2120 | \<exists>W::TPTP_Interpret.ind. \<not> Phi W) \<and> | |
| 2121 | bnd_iatom = (\<lambda>P::TPTP_Interpret.ind \<Rightarrow> bool. P) \<and> | |
| 55596 | 2122 | bnd_iand = bnd_mand \<and> | 
| 2123 | bnd_ivalid | |
| 2124 | (bnd_iimplies (bnd_iatom bnd_q) (bnd_iatom bnd_r)) \<Longrightarrow> | |
| 61076 | 2125 | (\<forall>SY161::TPTP_Interpret.ind. | 
| 2126 | \<not> (\<forall>SY162::TPTP_Interpret.ind. | |
| 55596 | 2127 | bnd_irel SY161 SY162 \<longrightarrow> bnd_q SY162) \<or> | 
| 61076 | 2128 | (\<forall>SY163::TPTP_Interpret.ind. | 
| 55596 | 2129 | bnd_irel SY161 SY163 \<longrightarrow> bnd_r SY163)) = | 
| 2130 | True" | |
| 2131 | (* by (tactic {*tectoc @{context}*}) *)
 | |
| 2132 | by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "11") 1*})
 | |
| 2133 | ||
| 2134 | lemma " | |
| 2135 | (\<forall>SY136. | |
| 2136 | \<not> (\<forall>SY137. bnd_irel SY136 SY137 \<longrightarrow> bnd_r SY137) \<or> | |
| 2137 | (\<forall>SY138. | |
| 2138 | bnd_irel SY136 SY138 \<longrightarrow> bnd_p SY138 \<and> bnd_q SY138)) = | |
| 2139 | True \<Longrightarrow> | |
| 2140 | (\<forall>SY136. | |
| 2141 | bnd_irel SY136 (bnd_sK5 SY136) \<and> \<not> bnd_r (bnd_sK5 SY136) \<or> | |
| 2142 | (\<forall>SY138. \<not> bnd_irel SY136 SY138 \<or> bnd_p SY138) \<and> | |
| 2143 | (\<forall>SY138. \<not> bnd_irel SY136 SY138 \<or> bnd_q SY138)) = | |
| 2144 | True" | |
| 2145 | by (tactic {*HEADGOAL (extcnf_combined_tac Full false (hd prob_names))*})
 | |
| 2146 | *) | |
| 2147 | ||
| 2148 |      (* (Annotated_step ("11", "extcnf_forall_neg"), *)
 | |
| 61076 | 2149 | lemma "\<forall>SV1::TPTP_Interpret.ind \<Rightarrow> bool. | 
| 2150 | (\<forall>SY2::TPTP_Interpret.ind. | |
| 55596 | 2151 | \<not> (\<not> (\<not> SV1 SY2 \<or> SEV405_5_bnd_cA) \<or> | 
| 2152 | \<not> (\<not> SEV405_5_bnd_cA \<or> SV1 SY2))) = | |
| 2153 | False \<Longrightarrow> | |
| 61076 | 2154 | \<forall>SV1::TPTP_Interpret.ind \<Rightarrow> bool. | 
| 55596 | 2155 | (\<not> (\<not> (\<not> SV1 (SEV405_5_bnd_sK1_SY2 SV1) \<or> SEV405_5_bnd_cA) \<or> | 
| 2156 | \<not> (\<not> SEV405_5_bnd_cA \<or> SV1 (SEV405_5_bnd_sK1_SY2 SV1)))) = | |
| 2157 | False" | |
| 2158 | (* apply (tactic {*full_extcnf_combined_tac*}) *)
 | |
| 63167 | 2159 | by (tactic \<open>nonfull_extcnf_combined_tac @{context} [Existential_Var]\<close>)
 | 
| 55596 | 2160 | |
| 2161 |      (* (Annotated_step ("13", "extcnf_forall_pos"), *)
 | |
| 2162 | lemma "(\<forall>SY31 SY32. | |
| 2163 | bnd_sK2 (bnd_sK4 SY31 SY32) = | |
| 2164 | bnd_sK5 (bnd_sK2 SY31) (bnd_sK2 SY32)) = | |
| 2165 | True \<Longrightarrow> | |
| 2166 | \<forall>SV1. (\<forall>SY42. | |
| 2167 | bnd_sK2 (bnd_sK4 SV1 SY42) = | |
| 2168 | bnd_sK5 (bnd_sK2 SV1) (bnd_sK2 SY42)) = | |
| 2169 | True" | |
| 63167 | 2170 | by (tactic \<open>nonfull_extcnf_combined_tac @{context} [Universal]\<close>)
 | 
| 55596 | 2171 | |
| 2172 |      (* (Annotated_step ("14", "extcnf_forall_pos"), *)
 | |
| 2173 | lemma "(\<forall>SY35 SY36. | |
| 2174 | bnd_sK1 (bnd_sK3 SY35 SY36) = | |
| 2175 | bnd_sK4 (bnd_sK1 SY35) (bnd_sK1 SY36)) = | |
| 2176 | True \<Longrightarrow> | |
| 2177 | \<forall>SV2. (\<forall>SY43. | |
| 2178 | bnd_sK1 (bnd_sK3 SV2 SY43) = | |
| 2179 | bnd_sK4 (bnd_sK1 SV2) (bnd_sK1 SY43)) = | |
| 2180 | True" | |
| 63167 | 2181 | by (tactic \<open>nonfull_extcnf_combined_tac @{context} [Universal]\<close>)
 | 
| 55596 | 2182 | |
| 2183 | ||
| 2184 | (*from SYO198^5.p.out*) | |
| 2185 |    (* [[(Annotated_step ("11", "extcnf_forall_special_pos"), *)
 | |
| 61076 | 2186 | lemma "(\<forall>SX0::bool \<Rightarrow> bool. | 
| 55596 | 2187 | \<not> \<not> (\<not> SX0 bnd_sK1_Xx \<or> \<not> SX0 bnd_sK2_Xy)) = | 
| 2188 | True \<Longrightarrow> | |
| 2189 | (\<not> \<not> (\<not> True \<or> \<not> True)) = True" | |
| 63167 | 2190 | apply (tactic \<open>extcnf_forall_special_pos_tac 1\<close>) | 
| 55596 | 2191 | done | 
| 2192 | ||
| 2193 |      (* (Annotated_step ("13", "extcnf_forall_special_pos"), *)
 | |
| 61076 | 2194 | lemma "(\<forall>SX0::bool \<Rightarrow> bool. | 
| 55596 | 2195 | \<not> \<not> (\<not> SX0 bnd_sK1_Xx \<or> \<not> SX0 bnd_sK2_Xy)) = | 
| 2196 | True \<Longrightarrow> | |
| 2197 | (\<not> \<not> (\<not> bnd_sK1_Xx \<or> \<not> bnd_sK2_Xy)) = True" | |
| 63167 | 2198 | apply (tactic \<open>extcnf_forall_special_pos_tac 1\<close>) | 
| 55596 | 2199 | done | 
| 2200 | ||
| 2201 |    (* [[(Annotated_step ("8", "polarity_switch"), *)
 | |
| 61076 | 2202 | lemma "(\<forall>(Xx::bool) (Xy::bool) Xz::bool. True \<and> True \<longrightarrow> True) = | 
| 55596 | 2203 | False \<Longrightarrow> | 
| 61076 | 2204 | (\<not> (\<forall>(Xx::bool) (Xy::bool) Xz::bool. | 
| 55596 | 2205 | True \<and> True \<longrightarrow> True)) = | 
| 2206 | True" | |
| 63167 | 2207 | apply (tactic \<open>nonfull_extcnf_combined_tac @{context} [Polarity_switch]\<close>)
 | 
| 55596 | 2208 | done | 
| 2209 | ||
| 2210 | lemma "((\<forall>SY22 SY23 SY24. | |
| 2211 | bnd_sK1_RRR SY22 SY23 \<and> bnd_sK1_RRR SY23 SY24 \<longrightarrow> | |
| 2212 | bnd_sK1_RRR SY22 SY24) \<and> | |
| 2213 | (\<forall>SY25. | |
| 2214 | (\<forall>SY26. SY25 SY26 \<longrightarrow> bnd_sK1_RRR SY26 (bnd_sK2_U SY25)) \<and> | |
| 2215 | (\<forall>SY27. | |
| 2216 | (\<forall>SY28. SY25 SY28 \<longrightarrow> bnd_sK1_RRR SY28 SY27) \<longrightarrow> | |
| 2217 | bnd_sK1_RRR (bnd_sK2_U SY25) SY27)) \<longrightarrow> | |
| 2218 | (\<forall>SY29. \<exists>SY30. \<forall>SY31. SY29 SY31 \<longrightarrow> bnd_sK1_RRR SY31 SY30)) = | |
| 2219 | False \<Longrightarrow> | |
| 2220 | (\<forall>SY25. | |
| 2221 | (\<forall>SY26. SY25 SY26 \<longrightarrow> bnd_sK1_RRR SY26 (bnd_sK2_U SY25)) \<and> | |
| 2222 | (\<forall>SY27. | |
| 2223 | (\<forall>SY28. SY25 SY28 \<longrightarrow> bnd_sK1_RRR SY28 SY27) \<longrightarrow> | |
| 2224 | bnd_sK1_RRR (bnd_sK2_U SY25) SY27)) = | |
| 2225 | True" | |
| 63167 | 2226 | apply (tactic \<open>standard_cnf_tac @{context} 1\<close>)
 | 
| 55596 | 2227 | done | 
| 2228 | ||
| 2229 | lemma "((\<forall>Xx. bnd_in Xx bnd_emptyset \<longrightarrow> (\<forall>Xphi. Xphi)) \<longrightarrow> | |
| 2230 | (\<forall>Xx Xy. bnd_in Xx (bnd_setadjoin Xx Xy)) \<longrightarrow> | |
| 2231 | (\<forall>A B. A = B \<longrightarrow> | |
| 2232 | (\<forall>Xx Xy. Xx = Xy \<longrightarrow> bnd_in Xx A = bnd_in Xy B)) \<longrightarrow> | |
| 2233 | (\<forall>SY0. bnd_in SY0 bnd_omega \<longrightarrow> | |
| 2234 | bnd_setadjoin SY0 SY0 \<noteq> bnd_emptyset)) = | |
| 2235 | False \<Longrightarrow> | |
| 2236 | (\<forall>Xx. bnd_in Xx bnd_emptyset \<longrightarrow> (\<forall>Xphi. Xphi)) = | |
| 2237 | True" | |
| 63167 | 2238 | apply (tactic \<open>standard_cnf_tac @{context} 1\<close>)
 | 
| 55596 | 2239 | done | 
| 2240 | ||
| 2241 | lemma "((\<forall>Xx. bnd_in Xx bnd_emptyset \<longrightarrow> (\<forall>Xphi. Xphi)) \<longrightarrow> | |
| 2242 | (\<forall>Xx Xy. bnd_in Xx (bnd_setadjoin Xx Xy)) \<longrightarrow> | |
| 2243 | (\<forall>A B. A = B \<longrightarrow> | |
| 2244 | (\<forall>Xx Xy. Xx = Xy \<longrightarrow> bnd_in Xx A = bnd_in Xy B)) \<longrightarrow> | |
| 2245 | (\<forall>SY0. bnd_in SY0 bnd_omega \<longrightarrow> | |
| 2246 | bnd_setadjoin SY0 SY0 \<noteq> bnd_emptyset)) = | |
| 2247 | False \<Longrightarrow> | |
| 2248 | (\<forall>Xx Xy. bnd_in Xx (bnd_setadjoin Xx Xy)) = | |
| 2249 | True" | |
| 63167 | 2250 | apply (tactic \<open>standard_cnf_tac @{context} 1\<close>)
 | 
| 55596 | 2251 | done | 
| 2252 | ||
| 2253 | lemma "((\<forall>Xx. bnd_in Xx bnd_emptyset \<longrightarrow> (\<forall>Xphi. Xphi)) \<longrightarrow> | |
| 2254 | (\<forall>Xx Xy. bnd_in Xx (bnd_setadjoin Xx Xy)) \<longrightarrow> | |
| 2255 | (\<forall>A B. A = B \<longrightarrow> | |
| 2256 | (\<forall>Xx Xy. Xx = Xy \<longrightarrow> bnd_in Xx A = bnd_in Xy B)) \<longrightarrow> | |
| 2257 | (\<forall>SY0. bnd_in SY0 bnd_omega \<longrightarrow> | |
| 2258 | bnd_setadjoin SY0 SY0 \<noteq> bnd_emptyset)) = | |
| 2259 | False \<Longrightarrow> | |
| 2260 | (\<forall>A B. A = B \<longrightarrow> | |
| 2261 | (\<forall>Xx Xy. Xx = Xy \<longrightarrow> bnd_in Xx A = bnd_in Xy B)) = | |
| 2262 | True" | |
| 63167 | 2263 | apply (tactic \<open>standard_cnf_tac @{context} 1\<close>)
 | 
| 55596 | 2264 | done | 
| 2265 | ||
| 2266 | lemma "((\<forall>Xx. bnd_in Xx bnd_emptyset \<longrightarrow> (\<forall>Xphi. Xphi)) \<longrightarrow> | |
| 2267 | (\<forall>Xx Xy. bnd_in Xx (bnd_setadjoin Xx Xy)) \<longrightarrow> | |
| 2268 | (\<forall>A B. A = B \<longrightarrow> | |
| 2269 | (\<forall>Xx Xy. Xx = Xy \<longrightarrow> bnd_in Xx A = bnd_in Xy B)) \<longrightarrow> | |
| 2270 | (\<forall>SY0. bnd_in SY0 bnd_omega \<longrightarrow> | |
| 2271 | bnd_setadjoin SY0 SY0 \<noteq> bnd_emptyset)) = | |
| 2272 | False \<Longrightarrow> | |
| 2273 | (\<forall>SY0. bnd_in SY0 bnd_omega \<longrightarrow> | |
| 2274 | bnd_setadjoin SY0 SY0 \<noteq> bnd_emptyset) = | |
| 2275 | False" | |
| 63167 | 2276 | apply (tactic \<open>standard_cnf_tac @{context} 1\<close>)
 | 
| 55596 | 2277 | done | 
| 2278 | ||
| 2279 | (*nested conjunctions*) | |
| 2280 | lemma "((((\<forall>Xx. bnd_cP bnd_e Xx = Xx) \<and> | |
| 2281 | (\<forall>Xy. bnd_cP Xy bnd_e = Xy)) \<and> | |
| 2282 | (\<forall>Xz. bnd_cP Xz Xz = bnd_e)) \<and> | |
| 2283 | (\<forall>Xx Xy Xz. | |
| 2284 | bnd_cP (bnd_cP Xx Xy) Xz = bnd_cP Xx (bnd_cP Xy Xz)) \<longrightarrow> | |
| 2285 | (\<forall>Xa Xb. bnd_cP Xa Xb = bnd_cP Xb Xa)) = | |
| 2286 | False \<Longrightarrow> | |
| 2287 | (\<forall>Xx. bnd_cP bnd_e Xx = Xx) = | |
| 2288 | True" | |
| 63167 | 2289 | apply (tactic \<open>standard_cnf_tac @{context} 1\<close>)
 | 
| 55596 | 2290 | done | 
| 2291 | ||
| 62390 | 2292 | end |