| author | nipkow | 
| Fri, 15 Jul 2022 08:46:04 +0200 | |
| changeset 75667 | 33177228aa69 | 
| parent 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  | 
|
| 69605 | 9  | 
ML_file \<open>~~/src/Tools/Argo/argo_expr.ML\<close>  | 
10  | 
ML_file \<open>~~/src/Tools/Argo/argo_term.ML\<close>  | 
|
11  | 
ML_file \<open>~~/src/Tools/Argo/argo_lit.ML\<close>  | 
|
12  | 
ML_file \<open>~~/src/Tools/Argo/argo_proof.ML\<close>  | 
|
13  | 
ML_file \<open>~~/src/Tools/Argo/argo_rewr.ML\<close>  | 
|
14  | 
ML_file \<open>~~/src/Tools/Argo/argo_cls.ML\<close>  | 
|
15  | 
ML_file \<open>~~/src/Tools/Argo/argo_common.ML\<close>  | 
|
16  | 
ML_file \<open>~~/src/Tools/Argo/argo_cc.ML\<close>  | 
|
17  | 
ML_file \<open>~~/src/Tools/Argo/argo_simplex.ML\<close>  | 
|
18  | 
ML_file \<open>~~/src/Tools/Argo/argo_thy.ML\<close>  | 
|
19  | 
ML_file \<open>~~/src/Tools/Argo/argo_heap.ML\<close>  | 
|
20  | 
ML_file \<open>~~/src/Tools/Argo/argo_cdcl.ML\<close>  | 
|
21  | 
ML_file \<open>~~/src/Tools/Argo/argo_core.ML\<close>  | 
|
22  | 
ML_file \<open>~~/src/Tools/Argo/argo_clausify.ML\<close>  | 
|
23  | 
ML_file \<open>~~/src/Tools/Argo/argo_solver.ML\<close>  | 
|
| 
63960
 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 
boehmes 
parents:  
diff
changeset
 | 
24  | 
|
| 69605 | 25  | 
ML_file \<open>Tools/Argo/argo_tactic.ML\<close>  | 
| 
63960
 
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  |