src/HOL/ATP.thy
author nipkow
Mon, 30 May 2022 20:58:45 +0200
changeset 75496 99b37c391433
parent 74899 b4beb55c574e
permissions -rw-r--r--
Added lemmas
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
39951
ff60a6e4edfe factor out "ATP" from "Sledgehammer" (cf. "SAT" vs. "Refute", etc.) -- the theories now reflect the directory structure
blanchet
parents:
diff changeset
     1
(*  Title:      HOL/ATP.thy
ff60a6e4edfe factor out "ATP" from "Sledgehammer" (cf. "SAT" vs. "Refute", etc.) -- the theories now reflect the directory structure
blanchet
parents:
diff changeset
     2
    Author:     Fabian Immler, TU Muenchen
ff60a6e4edfe factor out "ATP" from "Sledgehammer" (cf. "SAT" vs. "Refute", etc.) -- the theories now reflect the directory structure
blanchet
parents:
diff changeset
     3
    Author:     Jasmin Blanchette, TU Muenchen
72922
d78bd4432f05 tuned name generation in tptp to not depend on shadowing
desharna
parents: 72593
diff changeset
     4
    Author:     Martin Desharnais, UniBw Muenchen
39951
ff60a6e4edfe factor out "ATP" from "Sledgehammer" (cf. "SAT" vs. "Refute", etc.) -- the theories now reflect the directory structure
blanchet
parents:
diff changeset
     5
*)
ff60a6e4edfe factor out "ATP" from "Sledgehammer" (cf. "SAT" vs. "Refute", etc.) -- the theories now reflect the directory structure
blanchet
parents:
diff changeset
     6
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
     7
section \<open>Automatic Theorem Provers (ATPs)\<close>
39951
ff60a6e4edfe factor out "ATP" from "Sledgehammer" (cf. "SAT" vs. "Refute", etc.) -- the theories now reflect the directory structure
blanchet
parents:
diff changeset
     8
ff60a6e4edfe factor out "ATP" from "Sledgehammer" (cf. "SAT" vs. "Refute", etc.) -- the theories now reflect the directory structure
blanchet
parents:
diff changeset
     9
theory ATP
74896
f9908452b282 proper proxy for Hilbert choice in TPTP output
desharna
parents: 72922
diff changeset
    10
  imports Meson Hilbert_Choice
39951
ff60a6e4edfe factor out "ATP" from "Sledgehammer" (cf. "SAT" vs. "Refute", etc.) -- the theories now reflect the directory structure
blanchet
parents:
diff changeset
    11
begin
ff60a6e4edfe factor out "ATP" from "Sledgehammer" (cf. "SAT" vs. "Refute", etc.) -- the theories now reflect the directory structure
blanchet
parents:
diff changeset
    12
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    13
subsection \<open>ATP problems and proofs\<close>
57263
2b6a96cc64c9 simplified code
blanchet
parents: 57262
diff changeset
    14
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 66364
diff changeset
    15
ML_file \<open>Tools/ATP/atp_util.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 66364
diff changeset
    16
ML_file \<open>Tools/ATP/atp_problem.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 66364
diff changeset
    17
ML_file \<open>Tools/ATP/atp_proof.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 66364
diff changeset
    18
ML_file \<open>Tools/ATP/atp_proof_redirect.ML\<close>
57707
0242e9578828 imported patch satallax_proof_support_Sledgehammer
fleury
parents: 57263
diff changeset
    19
48891
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 47946
diff changeset
    20
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    21
subsection \<open>Higher-order reasoning helpers\<close>
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 40178
diff changeset
    22
54148
c8cc5ab4a863 killed more "no_atp"s
blanchet
parents: 53479
diff changeset
    23
definition fFalse :: bool where
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 40178
diff changeset
    24
"fFalse \<longleftrightarrow> False"
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 40178
diff changeset
    25
54148
c8cc5ab4a863 killed more "no_atp"s
blanchet
parents: 53479
diff changeset
    26
definition fTrue :: bool where
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 40178
diff changeset
    27
"fTrue \<longleftrightarrow> True"
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 40178
diff changeset
    28
54148
c8cc5ab4a863 killed more "no_atp"s
blanchet
parents: 53479
diff changeset
    29
definition fNot :: "bool \<Rightarrow> bool" where
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 40178
diff changeset
    30
"fNot P \<longleftrightarrow> \<not> P"
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 40178
diff changeset
    31
54148
c8cc5ab4a863 killed more "no_atp"s
blanchet
parents: 53479
diff changeset
    32
definition fComp :: "('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool" where
47946
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    33
"fComp P = (\<lambda>x. \<not> P x)"
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    34
54148
c8cc5ab4a863 killed more "no_atp"s
blanchet
parents: 53479
diff changeset
    35
definition fconj :: "bool \<Rightarrow> bool \<Rightarrow> bool" where
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 40178
diff changeset
    36
"fconj P Q \<longleftrightarrow> P \<and> Q"
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 40178
diff changeset
    37
54148
c8cc5ab4a863 killed more "no_atp"s
blanchet
parents: 53479
diff changeset
    38
definition fdisj :: "bool \<Rightarrow> bool \<Rightarrow> bool" where
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 40178
diff changeset
    39
"fdisj P Q \<longleftrightarrow> P \<or> Q"
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 40178
diff changeset
    40
54148
c8cc5ab4a863 killed more "no_atp"s
blanchet
parents: 53479
diff changeset
    41
definition fimplies :: "bool \<Rightarrow> bool \<Rightarrow> bool" where
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 40178
diff changeset
    42
"fimplies P Q \<longleftrightarrow> (P \<longrightarrow> Q)"
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 40178
diff changeset
    43
54148
c8cc5ab4a863 killed more "no_atp"s
blanchet
parents: 53479
diff changeset
    44
definition fAll :: "('a \<Rightarrow> bool) \<Rightarrow> bool" where
43678
56d352659500 improved translation of lambdas in THF
nik
parents: 43108
diff changeset
    45
"fAll P \<longleftrightarrow> All P"
56d352659500 improved translation of lambdas in THF
nik
parents: 43108
diff changeset
    46
54148
c8cc5ab4a863 killed more "no_atp"s
blanchet
parents: 53479
diff changeset
    47
definition fEx :: "('a \<Rightarrow> bool) \<Rightarrow> bool" where
43678
56d352659500 improved translation of lambdas in THF
nik
parents: 43108
diff changeset
    48
"fEx P \<longleftrightarrow> Ex P"
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 40178
diff changeset
    49
56946
10d9bd4ea94f hide more internal names
blanchet
parents: 54148
diff changeset
    50
definition fequal :: "'a \<Rightarrow> 'a \<Rightarrow> bool" where
10d9bd4ea94f hide more internal names
blanchet
parents: 54148
diff changeset
    51
"fequal x y \<longleftrightarrow> (x = y)"
10d9bd4ea94f hide more internal names
blanchet
parents: 54148
diff changeset
    52
74896
f9908452b282 proper proxy for Hilbert choice in TPTP output
desharna
parents: 72922
diff changeset
    53
definition fChoice :: "('a \<Rightarrow> bool) \<Rightarrow> 'a" where
f9908452b282 proper proxy for Hilbert choice in TPTP output
desharna
parents: 72922
diff changeset
    54
  "fChoice \<equiv> Hilbert_Choice.Eps"
f9908452b282 proper proxy for Hilbert choice in TPTP output
desharna
parents: 72922
diff changeset
    55
47946
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    56
lemma fTrue_ne_fFalse: "fFalse \<noteq> fTrue"
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    57
unfolding fFalse_def fTrue_def by simp
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    58
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    59
lemma fNot_table:
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    60
"fNot fFalse = fTrue"
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    61
"fNot fTrue = fFalse"
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    62
unfolding fFalse_def fTrue_def fNot_def by auto
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    63
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    64
lemma fconj_table:
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    65
"fconj fFalse P = fFalse"
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    66
"fconj P fFalse = fFalse"
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    67
"fconj fTrue fTrue = fTrue"
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    68
unfolding fFalse_def fTrue_def fconj_def by auto
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    69
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    70
lemma fdisj_table:
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    71
"fdisj fTrue P = fTrue"
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    72
"fdisj P fTrue = fTrue"
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    73
"fdisj fFalse fFalse = fFalse"
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    74
unfolding fFalse_def fTrue_def fdisj_def by auto
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    75
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    76
lemma fimplies_table:
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    77
"fimplies P fTrue = fTrue"
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    78
"fimplies fFalse P = fTrue"
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    79
"fimplies fTrue fFalse = fFalse"
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    80
unfolding fFalse_def fTrue_def fimplies_def by auto
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    81
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    82
lemma fAll_table:
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    83
"Ex (fComp P) \<or> fAll P = fTrue"
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    84
"All P \<or> fAll P = fFalse"
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    85
unfolding fFalse_def fTrue_def fComp_def fAll_def by auto
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    86
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    87
lemma fEx_table:
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    88
"All (fComp P) \<or> fEx P = fTrue"
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    89
"Ex P \<or> fEx P = fFalse"
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    90
unfolding fFalse_def fTrue_def fComp_def fEx_def by auto
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    91
56946
10d9bd4ea94f hide more internal names
blanchet
parents: 54148
diff changeset
    92
lemma fequal_table:
10d9bd4ea94f hide more internal names
blanchet
parents: 54148
diff changeset
    93
"fequal x x = fTrue"
10d9bd4ea94f hide more internal names
blanchet
parents: 54148
diff changeset
    94
"x = y \<or> fequal x y = fFalse"
10d9bd4ea94f hide more internal names
blanchet
parents: 54148
diff changeset
    95
unfolding fFalse_def fTrue_def fequal_def by auto
10d9bd4ea94f hide more internal names
blanchet
parents: 54148
diff changeset
    96
47946
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    97
lemma fNot_law:
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    98
"fNot P \<noteq> P"
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
    99
unfolding fNot_def by auto
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   100
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   101
lemma fComp_law:
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   102
"fComp P x \<longleftrightarrow> \<not> P x"
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   103
unfolding fComp_def ..
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   104
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   105
lemma fconj_laws:
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   106
"fconj P P \<longleftrightarrow> P"
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   107
"fconj P Q \<longleftrightarrow> fconj Q P"
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   108
"fNot (fconj P Q) \<longleftrightarrow> fdisj (fNot P) (fNot Q)"
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   109
unfolding fNot_def fconj_def fdisj_def by auto
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   110
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   111
lemma fdisj_laws:
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   112
"fdisj P P \<longleftrightarrow> P"
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   113
"fdisj P Q \<longleftrightarrow> fdisj Q P"
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   114
"fNot (fdisj P Q) \<longleftrightarrow> fconj (fNot P) (fNot Q)"
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   115
unfolding fNot_def fconj_def fdisj_def by auto
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   116
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   117
lemma fimplies_laws:
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   118
"fimplies P Q \<longleftrightarrow> fdisj (\<not> P) Q"
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   119
"fNot (fimplies P Q) \<longleftrightarrow> fconj P (fNot Q)"
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   120
unfolding fNot_def fconj_def fdisj_def fimplies_def by auto
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   121
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   122
lemma fAll_law:
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   123
"fNot (fAll R) \<longleftrightarrow> fEx (fComp R)"
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   124
unfolding fNot_def fComp_def fAll_def fEx_def by auto
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   125
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   126
lemma fEx_law:
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   127
"fNot (fEx R) \<longleftrightarrow> fAll (fComp R)"
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   128
unfolding fNot_def fComp_def fAll_def fEx_def by auto
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 46320
diff changeset
   129
56946
10d9bd4ea94f hide more internal names
blanchet
parents: 54148
diff changeset
   130
lemma fequal_laws:
10d9bd4ea94f hide more internal names
blanchet
parents: 54148
diff changeset
   131
"fequal x y = fequal y x"
10d9bd4ea94f hide more internal names
blanchet
parents: 54148
diff changeset
   132
"fequal x y = fFalse \<or> fequal y z = fFalse \<or> fequal x z = fTrue"
10d9bd4ea94f hide more internal names
blanchet
parents: 54148
diff changeset
   133
"fequal x y = fFalse \<or> fequal (f x) (f y) = fTrue"
10d9bd4ea94f hide more internal names
blanchet
parents: 54148
diff changeset
   134
unfolding fFalse_def fTrue_def fequal_def by auto
10d9bd4ea94f hide more internal names
blanchet
parents: 54148
diff changeset
   135
74899
b4beb55c574e proper handling of Hilbert choice in TFX logics
desharna
parents: 74896
diff changeset
   136
lemma fChoice_iff_Ex: "P (fChoice P) \<longleftrightarrow> HOL.Ex P"
b4beb55c574e proper handling of Hilbert choice in TFX logics
desharna
parents: 74896
diff changeset
   137
  unfolding fChoice_def
74896
f9908452b282 proper proxy for Hilbert choice in TPTP output
desharna
parents: 72922
diff changeset
   138
  by (fact some_eq_ex)
56946
10d9bd4ea94f hide more internal names
blanchet
parents: 54148
diff changeset
   139
74899
b4beb55c574e proper handling of Hilbert choice in TFX logics
desharna
parents: 74896
diff changeset
   140
text \<open>
b4beb55c574e proper handling of Hilbert choice in TFX logics
desharna
parents: 74896
diff changeset
   141
We use the @{const HOL.Ex} constant on the right-hand side of @{thm [source] fChoice_iff_Ex} because
b4beb55c574e proper handling of Hilbert choice in TFX logics
desharna
parents: 74896
diff changeset
   142
we want to use the TPTP-native version if fChoice is introduced in a logic that supports FOOL.
b4beb55c574e proper handling of Hilbert choice in TFX logics
desharna
parents: 74896
diff changeset
   143
In logics that don't support it, it gets replaced by fEx during processing.
b4beb55c574e proper handling of Hilbert choice in TFX logics
desharna
parents: 74896
diff changeset
   144
Notice that we cannot use @{term "\<exists>x. P x"}, as existentials are not skolimized by the metis proof
b4beb55c574e proper handling of Hilbert choice in TFX logics
desharna
parents: 74896
diff changeset
   145
method but @{term "Ex P"} is eta-expanded if FOOL is supported.\<close>
b4beb55c574e proper handling of Hilbert choice in TFX logics
desharna
parents: 74896
diff changeset
   146
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   147
subsection \<open>Basic connection between ATPs and HOL\<close>
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 40178
diff changeset
   148
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 66364
diff changeset
   149
ML_file \<open>Tools/lambda_lifting.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 66364
diff changeset
   150
ML_file \<open>Tools/monomorph.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 66364
diff changeset
   151
ML_file \<open>Tools/ATP/atp_problem_generate.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 66364
diff changeset
   152
ML_file \<open>Tools/ATP/atp_proof_reconstruct.ML\<close>
39951
ff60a6e4edfe factor out "ATP" from "Sledgehammer" (cf. "SAT" vs. "Refute", etc.) -- the theories now reflect the directory structure
blanchet
parents:
diff changeset
   153
ff60a6e4edfe factor out "ATP" from "Sledgehammer" (cf. "SAT" vs. "Refute", etc.) -- the theories now reflect the directory structure
blanchet
parents:
diff changeset
   154
end