author | haftmann |
Sun, 08 Oct 2017 22:28:22 +0200 | |
changeset 66814 | a24cde9588bb |
parent 63960 | 3daf02070be5 |
child 69605 | a96320074298 |
permissions | -rw-r--r-- |
63960
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
1 |
(* Title: HOL/Argo.thy |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
2 |
Author: Sascha Boehme |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
3 |
*) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
4 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
5 |
theory Argo |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
6 |
imports HOL |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
7 |
begin |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
8 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
9 |
ML_file "~~/src/Tools/Argo/argo_expr.ML" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
10 |
ML_file "~~/src/Tools/Argo/argo_term.ML" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
11 |
ML_file "~~/src/Tools/Argo/argo_lit.ML" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
12 |
ML_file "~~/src/Tools/Argo/argo_proof.ML" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
13 |
ML_file "~~/src/Tools/Argo/argo_rewr.ML" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
14 |
ML_file "~~/src/Tools/Argo/argo_cls.ML" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
15 |
ML_file "~~/src/Tools/Argo/argo_common.ML" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
16 |
ML_file "~~/src/Tools/Argo/argo_cc.ML" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
17 |
ML_file "~~/src/Tools/Argo/argo_simplex.ML" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
18 |
ML_file "~~/src/Tools/Argo/argo_thy.ML" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
19 |
ML_file "~~/src/Tools/Argo/argo_heap.ML" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
20 |
ML_file "~~/src/Tools/Argo/argo_cdcl.ML" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
21 |
ML_file "~~/src/Tools/Argo/argo_core.ML" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
22 |
ML_file "~~/src/Tools/Argo/argo_clausify.ML" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
23 |
ML_file "~~/src/Tools/Argo/argo_solver.ML" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
24 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
25 |
ML_file "Tools/Argo/argo_tactic.ML" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
26 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
27 |
end |