author | nipkow |
Tue, 28 Sep 1999 16:37:04 +0200 | |
changeset 7627 | 6b0709a2f6c7 |
parent 7624 | 9024e9d370c7 |
child 7629 | 68e155f81f88 |
permissions | -rw-r--r-- |
2448 | 1 |
# |
2 |
# $Id$ |
|
3 |
# |
|
4 |
# IsaMakefile for HOL |
|
5 |
# |
|
6 |
||
4518 | 7 |
## targets |
2448 | 8 |
|
4518 | 9 |
default: HOL |
7125 | 10 |
images: HOL HOL-Real TLA |
11 |
test: HOL-Subst HOL-Induct HOL-IMP HOL-Hoare HOL-Lex \ |
|
7627 | 12 |
HOL-Auth HOL-UNITY HOL-Modelcheck HOL-Lambda HOL-W0 HOL-MiniML HOL-BCV \ |
7307 | 13 |
HOL-IOA HOL-AxClasses-Group HOL-AxClasses-Lattice \ |
7393 | 14 |
HOL-AxClasses-Tutorial HOL-Quot HOL-ex HOL-Isar_examples HOL-Real-ex \ |
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
7513
diff
changeset
|
15 |
HOL-Real-HahnBanach TLA-Inc TLA-Buffer TLA-Memory |
6445 | 16 |
|
4518 | 17 |
all: images test |
18 |
||
19 |
||
20 |
## global settings |
|
21 |
||
22 |
SRC = $(ISABELLE_HOME)/src |
|
3118 | 23 |
OUT = $(ISABELLE_OUTPUT) |
4447 | 24 |
LOG = $(OUT)/log |
2448 | 25 |
|
4518 | 26 |
|
27 |
## HOL |
|
2448 | 28 |
|
4518 | 29 |
HOL: Pure $(OUT)/HOL |
30 |
||
31 |
Pure: |
|
32 |
@cd $(SRC)/Pure; $(ISATOOL) make Pure |
|
3232
19a2b853ba7b
Removal of ex/LexProd; TFL files; new treatment of Prover files
paulson
parents:
3222
diff
changeset
|
33 |
|
6905 | 34 |
$(OUT)/HOL: $(OUT)/Pure $(SRC)/Provers/Arith/abel_cancel.ML \ |
7357 | 35 |
$(SRC)/Provers/Arith/cancel_sums.ML \ |
36 |
$(SRC)/Provers/Arith/assoc_fold.ML \ |
|
37 |
$(SRC)/Provers/Arith/combine_coeff.ML \ |
|
6905 | 38 |
$(SRC)/Provers/Arith/fast_lin_arith.ML $(SRC)/Provers/blast.ML \ |
4707 | 39 |
$(SRC)/Provers/clasimp.ML $(SRC)/Provers/classical.ML \ |
40 |
$(SRC)/Provers/hypsubst.ML $(SRC)/Provers/simplifier.ML \ |
|
5699 | 41 |
$(SRC)/Provers/split_paired_all.ML $(SRC)/Provers/splitter.ML \ |
42 |
$(SRC)/Pure/section_utils.ML $(SRC)/TFL/dcterm.sml \ |
|
6496
a185927883e5
Addition of Auth/KerberosIV; renaming of rules.new.sml to rules.sml
paulson
parents:
6474
diff
changeset
|
43 |
$(SRC)/TFL/post.sml $(SRC)/TFL/rules.sml $(SRC)/TFL/rules.sig \ |
5699 | 44 |
$(SRC)/TFL/sys.sml $(SRC)/TFL/tfl.sig $(SRC)/TFL/tfl.sml \ |
45 |
$(SRC)/TFL/thms.sig $(SRC)/TFL/thms.sml $(SRC)/TFL/thry.sig \ |
|
46 |
$(SRC)/TFL/thry.sml $(SRC)/TFL/usyntax.sig $(SRC)/TFL/usyntax.sml \ |
|
47 |
$(SRC)/TFL/utils.sig $(SRC)/TFL/utils.sml Arith.ML Arith.thy \ |
|
6779
2912aff958bd
Calculation.thy: Setup transitivity rules for calculational proofs.
wenzelm
parents:
6764
diff
changeset
|
48 |
Calculation.thy Datatype.thy Divides.ML Divides.thy Finite.ML \ |
2912aff958bd
Calculation.thy: Setup transitivity rules for calculational proofs.
wenzelm
parents:
6764
diff
changeset
|
49 |
Finite.thy Fun.ML Fun.thy Gfp.ML Gfp.thy HOL.ML HOL.thy \ |
7357 | 50 |
HOL_lemmas.ML Inductive.thy Integ/Bin.ML Integ/Bin.thy \ |
51 |
Integ/Equiv.ML Integ/Equiv.thy Integ/IntDef.ML Integ/IntDef.thy \ |
|
52 |
Integ/Int.ML Integ/Int.thy Integ/IntDiv.ML Integ/IntDiv.thy \ |
|
53 |
Integ/NatBin.ML Integ/NatBin.thy Integ/simproc.ML Lfp.ML Lfp.thy \ |
|
54 |
List.ML List.thy Main.thy Map.ML Map.thy Nat.ML Nat.thy NatDef.ML \ |
|
55 |
NatDef.thy Numeral.thy Option.ML Option.thy Ord.ML Ord.thy Power.ML \ |
|
56 |
Power.thy Prod.ML Prod.thy ROOT.ML Recdef.thy Record.thy RelPow.ML \ |
|
57 |
RelPow.thy Relation.ML Relation.thy Set.ML Set.thy Sexp.ML Sexp.thy \ |
|
58 |
String.thy SVC_Oracle.ML SVC_Oracle.thy Sum.ML Sum.thy \ |
|
59 |
Tools/datatype_aux.ML Tools/datatype_abs_proofs.ML \ |
|
60 |
Tools/datatype_package.ML Tools/datatype_prop.ML \ |
|
61 |
Tools/datatype_rep_proofs.ML Tools/induct_method.ML \ |
|
62 |
Tools/inductive_package.ML Tools/numeral_syntax.ML \ |
|
63 |
Tools/primrec_package.ML Tools/recdef_package.ML \ |
|
64 |
Tools/record_package.ML Tools/svc_funcs.ML Tools/typedef_package.ML \ |
|
65 |
Trancl.ML Trancl.thy Univ.ML Univ.thy Vimage.ML Vimage.thy WF.ML \ |
|
66 |
WF.thy WF_Rel.ML WF_Rel.thy blastdata.ML cladata.ML equalities.ML \ |
|
67 |
equalities.thy hologic.ML mono.ML mono.thy simpdata.ML subset.ML \ |
|
68 |
subset.thy thy_syntax.ML |
|
2826 | 69 |
@$(ISATOOL) usedir -b $(OUT)/Pure HOL |
2448 | 70 |
|
4518 | 71 |
|
7125 | 72 |
## HOL-Real |
73 |
||
74 |
HOL-Real: HOL $(OUT)/HOL-Real |
|
75 |
||
7334 | 76 |
$(OUT)/HOL-Real: $(OUT)/HOL Real/Lubs.ML Real/Lubs.thy Real/PNat.ML \ |
77 |
Real/PNat.thy Real/PRat.ML Real/PRat.thy Real/PReal.ML \ |
|
78 |
Real/PReal.thy Real/RComplete.ML Real/RComplete.thy Real/Real.thy \ |
|
79 |
Real/RealDef.ML Real/RealDef.thy Real/RealOrd.ML Real/RealOrd.thy \ |
|
80 |
Real/simproc.ML Real/RealAbs.ML Real/RealAbs.thy Real/ROOT.ML \ |
|
81 |
Real/RealBin.ML Real/RealBin.thy Real/RealInt.ML Real/RealInt.thy \ |
|
82 |
Real/RealPow.ML Real/RealPow.thy Real/Hyperreal/Filter.ML \ |
|
83 |
Real/Hyperreal/Filter.thy Real/Hyperreal/fuf.ML \ |
|
7218
bfa767b4dc51
new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
7186
diff
changeset
|
84 |
Real/Hyperreal/HyperDef.ML Real/Hyperreal/HyperDef.thy \ |
7125 | 85 |
Real/Hyperreal/Zorn.ML Real/Hyperreal/Zorn.thy |
86 |
@cd Real; $(ISATOOL) usedir -b $(OUT)/HOL HOL-Real |
|
87 |
||
7395 | 88 |
|
7393 | 89 |
## HOL-Real-ex |
7392 | 90 |
|
7393 | 91 |
HOL-Real-ex: HOL-Real $(LOG)/HOL-Real-ex.gz |
7392 | 92 |
|
7577 | 93 |
$(LOG)/HOL-Real-ex.gz: $(OUT)/HOL-Real Real/ex/ROOT.ML \ |
94 |
Real/ex/BinEx.thy |
|
7393 | 95 |
@cd Real; $(ISATOOL) usedir $(OUT)/HOL-Real ex |
7125 | 96 |
|
7395 | 97 |
|
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
7513
diff
changeset
|
98 |
## HOL-Real-HahnBanach |
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
7513
diff
changeset
|
99 |
|
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
7513
diff
changeset
|
100 |
HOL-Real-HahnBanach: HOL-Real $(LOG)/HOL-Real-HahnBanach.gz |
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
7513
diff
changeset
|
101 |
|
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
7513
diff
changeset
|
102 |
$(LOG)/HOL-Real-HahnBanach.gz: $(OUT)/HOL-Real Real/HahnBanach/Aux.thy \ |
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
7513
diff
changeset
|
103 |
Real/HahnBanach/Bounds.thy Real/HahnBanach/FunctionNorm.thy \ |
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
7513
diff
changeset
|
104 |
Real/HahnBanach/FunctionOrder.thy Real/HahnBanach/HahnBanach.thy \ |
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
7513
diff
changeset
|
105 |
Real/HahnBanach/HahnBanach_h0_lemmas.thy \ |
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
7513
diff
changeset
|
106 |
Real/HahnBanach/HahnBanach_lemmas.thy \ |
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
7513
diff
changeset
|
107 |
Real/HahnBanach/LinearSpace.thy Real/HahnBanach/Linearform.thy \ |
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
7513
diff
changeset
|
108 |
Real/HahnBanach/NormedSpace.thy Real/HahnBanach/ROOT.ML \ |
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
7513
diff
changeset
|
109 |
Real/HahnBanach/Subspace.thy Real/HahnBanach/Zorn_Lemma.thy |
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
7513
diff
changeset
|
110 |
@cd Real; $(ISATOOL) usedir $(OUT)/HOL-Real HahnBanach |
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
7513
diff
changeset
|
111 |
|
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
7513
diff
changeset
|
112 |
|
4518 | 113 |
## HOL-Subst |
114 |
||
115 |
HOL-Subst: HOL $(LOG)/HOL-Subst.gz |
|
116 |
||
117 |
$(LOG)/HOL-Subst.gz: $(OUT)/HOL Subst/AList.ML Subst/AList.thy \ |
|
118 |
Subst/ROOT.ML Subst/Subst.ML Subst/Subst.thy Subst/UTerm.ML \ |
|
119 |
Subst/UTerm.thy Subst/Unifier.ML Subst/Unifier.thy Subst/Unify.ML \ |
|
120 |
Subst/Unify.thy |
|
121 |
@$(ISATOOL) usedir $(OUT)/HOL Subst |
|
2448 | 122 |
|
123 |
||
4518 | 124 |
## HOL-Induct |
2473 | 125 |
|
4518 | 126 |
HOL-Induct: HOL $(LOG)/HOL-Induct.gz |
3125 | 127 |
|
4518 | 128 |
$(LOG)/HOL-Induct.gz: $(OUT)/HOL Induct/Acc.ML Induct/Acc.thy \ |
129 |
Induct/Com.ML Induct/Com.thy Induct/Comb.ML Induct/Comb.thy \ |
|
5616 | 130 |
Induct/Exp.ML Induct/Exp.thy Induct/LFilter.ML Induct/LFilter.thy \ |
4518 | 131 |
Induct/LList.ML Induct/LList.thy Induct/Mutil.ML Induct/Mutil.thy \ |
5627 | 132 |
Induct/Multiset0.ML Induct/Multiset0.thy \ |
133 |
Induct/Multiset.ML Induct/Multiset.thy \ |
|
4518 | 134 |
Induct/Perm.ML Induct/Perm.thy Induct/PropLog.ML Induct/PropLog.thy \ |
5740
3a466866f7b9
Directory Induct: Added new theory ABexp, removed obsolete
berghofe
parents:
5712
diff
changeset
|
135 |
Induct/ROOT.ML Induct/SList.ML Induct/SList.thy Induct/ABexp.ML \ |
3a466866f7b9
Directory Induct: Added new theory ABexp, removed obsolete
berghofe
parents:
5712
diff
changeset
|
136 |
Induct/ABexp.thy Induct/Term.ML Induct/Term.thy |
3125 | 137 |
@$(ISATOOL) usedir $(OUT)/HOL Induct |
138 |
||
139 |
||
4518 | 140 |
## HOL-IMP |
141 |
||
142 |
HOL-IMP: HOL $(LOG)/HOL-IMP.gz |
|
2448 | 143 |
|
5225
092e77b6f7c6
Removed HOL/IMP/Com.ML because it contained only an "open" declaration
paulson
parents:
5199
diff
changeset
|
144 |
$(LOG)/HOL-IMP.gz: $(OUT)/HOL IMP/Com.thy IMP/Denotation.ML \ |
4518 | 145 |
IMP/Denotation.thy IMP/Expr.ML IMP/Expr.thy IMP/Hoare.ML IMP/Hoare.thy \ |
146 |
IMP/Natural.ML IMP/Natural.thy IMP/ROOT.ML IMP/Transition.ML \ |
|
147 |
IMP/Transition.thy IMP/VC.ML IMP/VC.thy |
|
2826 | 148 |
@$(ISATOOL) usedir $(OUT)/HOL IMP |
2448 | 149 |
|
150 |
||
4518 | 151 |
## HOL-Hoare |
152 |
||
153 |
HOL-Hoare: HOL $(LOG)/HOL-Hoare.gz |
|
2448 | 154 |
|
4518 | 155 |
$(LOG)/HOL-Hoare.gz: $(OUT)/HOL Hoare/Arith2.ML Hoare/Arith2.thy \ |
156 |
Hoare/Examples.ML Hoare/Examples.thy Hoare/Hoare.ML Hoare/Hoare.thy \ |
|
157 |
Hoare/ROOT.ML |
|
2826 | 158 |
@$(ISATOOL) usedir $(OUT)/HOL Hoare |
2448 | 159 |
|
160 |
||
4518 | 161 |
## HOL-Lex |
162 |
||
163 |
HOL-Lex: HOL $(LOG)/HOL-Lex.gz |
|
2448 | 164 |
|
4830 | 165 |
$(LOG)/HOL-Lex.gz: $(OUT)/HOL \ |
4518 | 166 |
Lex/AutoChopper.thy Lex/AutoChopper.ML Lex/AutoChopper1.thy \ |
4830 | 167 |
Lex/AutoMaxChop.thy Lex/AutoMaxChop.ML Lex/AutoProj.thy Lex/AutoProj.ML \ |
168 |
Lex/Automata.thy Lex/Automata.ML Lex/Chopper.thy Lex/DA.thy Lex/DA.ML \ |
|
169 |
Lex/MaxChop.thy Lex/MaxChop.ML Lex/MaxPrefix.thy Lex/MaxPrefix.ML \ |
|
5322 | 170 |
Lex/NA.thy Lex/NA.ML Lex/NAe.thy Lex/NAe.ML Lex/RegExp2NAe.thy \ |
171 |
Lex/RegExp2NAe.ML Lex/RegExp2NA.thy Lex/RegExp2NA.ML \ |
|
4830 | 172 |
Lex/Prefix.thy Lex/Prefix.ML Lex/ROOT.ML \ |
173 |
Lex/RegExp.thy Lex/RegSet.thy Lex/RegSet.ML \ |
|
174 |
Lex/RegSet_of_nat_DA.thy Lex/RegSet_of_nat_DA.ML |
|
4518 | 175 |
@$(ISATOOL) usedir $(OUT)/HOL Lex |
2448 | 176 |
|
4518 | 177 |
|
178 |
## HOL-Auth |
|
3819 | 179 |
|
4518 | 180 |
HOL-Auth: HOL $(LOG)/HOL-Auth.gz |
3819 | 181 |
|
4518 | 182 |
$(LOG)/HOL-Auth.gz: $(OUT)/HOL Auth/Event.ML Auth/Event.thy \ |
183 |
Auth/Message.ML Auth/Message.thy Auth/NS_Public.ML Auth/NS_Public.thy \ |
|
184 |
Auth/NS_Public_Bad.ML Auth/NS_Public_Bad.thy Auth/NS_Shared.ML \ |
|
185 |
Auth/NS_Shared.thy Auth/OtwayRees.ML Auth/OtwayRees.thy \ |
|
186 |
Auth/OtwayRees_AN.ML Auth/OtwayRees_AN.thy Auth/OtwayRees_Bad.ML \ |
|
187 |
Auth/OtwayRees_Bad.thy Auth/Public.ML Auth/Public.thy Auth/ROOT.ML \ |
|
188 |
Auth/Recur.ML Auth/Recur.thy Auth/Shared.ML Auth/Shared.thy \ |
|
189 |
Auth/TLS.ML Auth/TLS.thy Auth/WooLam.ML Auth/WooLam.thy \ |
|
6496
a185927883e5
Addition of Auth/KerberosIV; renaming of rules.new.sml to rules.sml
paulson
parents:
6474
diff
changeset
|
190 |
Auth/Kerberos_BAN.ML Auth/Kerberos_BAN.thy \ |
a185927883e5
Addition of Auth/KerberosIV; renaming of rules.new.sml to rules.sml
paulson
parents:
6474
diff
changeset
|
191 |
Auth/KerberosIV.ML Auth/KerberosIV.thy \ |
6401 | 192 |
Auth/Yahalom.ML Auth/Yahalom.thy Auth/Yahalom2.ML Auth/Yahalom2.thy \ |
193 |
Auth/Yahalom_Bad.ML Auth/Yahalom_Bad.thy |
|
2826 | 194 |
@$(ISATOOL) usedir $(OUT)/HOL Auth |
2448 | 195 |
|
196 |
||
4777 | 197 |
## HOL-UNITY |
198 |
||
199 |
HOL-UNITY: HOL $(LOG)/HOL-UNITY.gz |
|
200 |
||
201 |
$(LOG)/HOL-UNITY.gz: $(OUT)/HOL UNITY/ROOT.ML\ |
|
6817 | 202 |
UNITY/Alloc.ML UNITY/Alloc.thy\ |
4777 | 203 |
UNITY/Channel.ML UNITY/Channel.thy UNITY/Common.ML UNITY/Common.thy\ |
5636 | 204 |
UNITY/Client.ML UNITY/Client.thy UNITY/Comp.ML UNITY/Comp.thy\ |
7400
fbd5582761e6
new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
7395
diff
changeset
|
205 |
UNITY/Guar.ML UNITY/Guar.thy\ |
4777 | 206 |
UNITY/Deadlock.ML UNITY/Deadlock.thy UNITY/FP.ML UNITY/FP.thy\ |
5358 | 207 |
UNITY/Union.ML UNITY/Union.thy UNITY/Handshake.ML UNITY/Handshake.thy\ |
7513 | 208 |
UNITY/TimerArray.ML UNITY/TimerArray.thy\ |
7624 | 209 |
UNITY/Extend.ML UNITY/Extend.thy UNITY/Project.ML UNITY/Project.thy\ |
6730 | 210 |
UNITY/Follows.ML UNITY/Follows.thy\ |
6817 | 211 |
UNITY/GenPrefix.thy UNITY/GenPrefix.ML \ |
7186 | 212 |
UNITY/LessThan.ML UNITY/LessThan.thy \ |
213 |
UNITY/Lift_prog.ML UNITY/Lift_prog.thy UNITY/ListOrder.thy\ |
|
6730 | 214 |
UNITY/Mutex.ML UNITY/Mutex.thy\ |
4777 | 215 |
UNITY/Network.ML UNITY/Network.thy UNITY/Reach.ML UNITY/Reach.thy\ |
216 |
UNITY/SubstAx.ML UNITY/SubstAx.thy UNITY/Token.ML UNITY/Token.thy\ |
|
6535 | 217 |
UNITY/UNITY.ML UNITY/UNITY.thy\ |
5430 | 218 |
UNITY/WFair.ML UNITY/WFair.thy UNITY/Lift.ML UNITY/Lift.thy\ |
5900 | 219 |
UNITY/PPROD.ML UNITY/PPROD.thy UNITY/NSP_Bad.ML UNITY/NSP_Bad.thy |
4777 | 220 |
@$(ISATOOL) usedir $(OUT)/HOL UNITY |
221 |
||
222 |
||
4518 | 223 |
## HOL-Modelcheck |
224 |
||
225 |
HOL-Modelcheck: HOL $(LOG)/HOL-Modelcheck.gz |
|
3218 | 226 |
|
4518 | 227 |
$(LOG)/HOL-Modelcheck.gz: $(OUT)/HOL Modelcheck/CTL.thy \ |
6474 | 228 |
Modelcheck/EindhovenExample.ML Modelcheck/EindhovenExample.thy \ |
229 |
Modelcheck/EindhovenSyn.ML Modelcheck/EindhovenSyn.thy \ |
|
230 |
Modelcheck/MuCalculus.ML Modelcheck/MuCalculus.thy \ |
|
231 |
Modelcheck/MuckeExample1.ML Modelcheck/MuckeExample1.thy \ |
|
6472 | 232 |
Modelcheck/MuckeExample2.ML Modelcheck/MuckeExample2.thy \ |
6474 | 233 |
Modelcheck/MuckeSyn.ML Modelcheck/MuckeSyn.thy Modelcheck/ROOT.ML \ |
234 |
Modelcheck/mucke_oracle.ML |
|
3218 | 235 |
@$(ISATOOL) usedir $(OUT)/HOL Modelcheck |
236 |
||
237 |
||
4518 | 238 |
## HOL-Lambda |
2448 | 239 |
|
4518 | 240 |
HOL-Lambda: HOL $(LOG)/HOL-Lambda.gz |
2448 | 241 |
|
4518 | 242 |
$(LOG)/HOL-Lambda.gz: $(OUT)/HOL Lambda/Commutation.ML \ |
5272 | 243 |
Lambda/Commutation.thy Lambda/Eta.ML Lambda/Eta.thy Lambda/InductTermi.ML \ |
244 |
Lambda/InductTermi.thy Lambda/Lambda.ML Lambda/Lambda.thy \ |
|
245 |
Lambda/ListApplication.ML Lambda/ListApplication.thy Lambda/ListBeta.ML \ |
|
246 |
Lambda/ListBeta.thy Lambda/ListOrder.ML Lambda/ListOrder.thy \ |
|
247 |
Lambda/ParRed.ML Lambda/ParRed.thy Lambda/ROOT.ML |
|
2826 | 248 |
@$(ISATOOL) usedir $(OUT)/HOL Lambda |
2448 | 249 |
|
250 |
||
4518 | 251 |
## HOL-W0 |
2527 | 252 |
|
4518 | 253 |
HOL-W0: HOL $(LOG)/HOL-W0.gz |
2527 | 254 |
|
4518 | 255 |
$(LOG)/HOL-W0.gz: $(OUT)/HOL W0/I.ML W0/I.thy W0/Maybe.ML W0/Maybe.thy \ |
256 |
W0/MiniML.ML W0/MiniML.thy W0/ROOT.ML W0/Type.ML W0/Type.thy W0/W.ML \ |
|
257 |
W0/W.thy |
|
2826 | 258 |
@$(ISATOOL) usedir $(OUT)/HOL W0 |
2527 | 259 |
|
260 |
||
4518 | 261 |
## HOL-MiniML |
2527 | 262 |
|
4518 | 263 |
HOL-MiniML: HOL $(LOG)/HOL-MiniML.gz |
2448 | 264 |
|
4518 | 265 |
$(LOG)/HOL-MiniML.gz: $(OUT)/HOL MiniML/Generalize.ML \ |
266 |
MiniML/Generalize.thy MiniML/Instance.ML MiniML/Instance.thy \ |
|
267 |
MiniML/Maybe.ML MiniML/Maybe.thy MiniML/MiniML.ML MiniML/MiniML.thy \ |
|
268 |
MiniML/ROOT.ML MiniML/Type.ML MiniML/Type.thy MiniML/W.ML MiniML/W.thy |
|
2826 | 269 |
@$(ISATOOL) usedir $(OUT)/HOL MiniML |
2448 | 270 |
|
7627 | 271 |
## HOL-BCV |
272 |
||
273 |
HOL-BCV:HOL $(LOG)/HOL-BCV.gz |
|
274 |
||
275 |
$(LOG)/HOL-BCV.gz: $(OUT)/HOL BCV/DFAandWTI.ML \ |
|
276 |
BCV/DFAandWTI.thy BCV/DFAimpl.ML BCV/DFAimpl.thy \ |
|
277 |
BCV/Fixpoint.ML BCV/Fixpoint.thy BCV/Machine.ML BCV/Machine.thy \ |
|
278 |
BCV/Orders.ML BCV/Orders.thy BCV/Orders0.ML BCV/Orders0.thy \ |
|
279 |
BCV/Plus.ML BCV/Plus.thy BCV/ROOT.ML BCV/SemiLattice.ML BCV/SemiLattice.thy \ |
|
280 |
BCV/Types0.ML BCV/Types0.thy BCV/Types.ML BCV/Types.thy |
|
281 |
@$(ISATOOL) usedir $(OUT)/HOL BCV |
|
2448 | 282 |
|
4518 | 283 |
## HOL-IOA |
284 |
||
285 |
HOL-IOA: HOL $(LOG)/HOL-IOA.gz |
|
2448 | 286 |
|
4518 | 287 |
$(LOG)/HOL-IOA.gz: $(OUT)/HOL IOA/Asig.ML IOA/Asig.thy IOA/IOA.ML \ |
288 |
IOA/IOA.thy IOA/ROOT.ML IOA/Solve.ML IOA/Solve.thy |
|
289 |
@$(ISATOOL) usedir $(OUT)/HOL IOA |
|
290 |
||
291 |
||
292 |
## HOL-AxClasses-Group |
|
2448 | 293 |
|
7307 | 294 |
HOL-AxClasses-Group: HOL $(LOG)/HOL-AxClasses-Group.gz |
4518 | 295 |
|
296 |
$(LOG)/HOL-AxClasses-Group.gz: $(OUT)/HOL AxClasses/Group/Group.ML \ |
|
297 |
AxClasses/Group/Group.thy AxClasses/Group/GroupDefs.ML \ |
|
298 |
AxClasses/Group/GroupDefs.thy AxClasses/Group/GroupInsts.thy \ |
|
299 |
AxClasses/Group/Monoid.thy AxClasses/Group/MonoidGroupInsts.thy \ |
|
300 |
AxClasses/Group/ROOT.ML AxClasses/Group/Sigs.thy |
|
301 |
@$(ISATOOL) usedir -s AxClasses-Group $(OUT)/HOL AxClasses/Group |
|
2448 | 302 |
|
303 |
||
4518 | 304 |
## HOL-AxClasses-Lattice |
2545 | 305 |
|
7307 | 306 |
HOL-AxClasses-Lattice: HOL $(LOG)/HOL-AxClasses-Lattice.gz |
2545 | 307 |
|
4518 | 308 |
$(LOG)/HOL-AxClasses-Lattice.gz: $(OUT)/HOL AxClasses/Lattice/CLattice.ML \ |
309 |
AxClasses/Lattice/CLattice.thy AxClasses/Lattice/LatInsts.ML \ |
|
310 |
AxClasses/Lattice/LatInsts.thy AxClasses/Lattice/LatMorph.ML \ |
|
311 |
AxClasses/Lattice/LatMorph.thy AxClasses/Lattice/LatPreInsts.ML \ |
|
312 |
AxClasses/Lattice/LatPreInsts.thy AxClasses/Lattice/Lattice.ML \ |
|
313 |
AxClasses/Lattice/Lattice.thy AxClasses/Lattice/OrdDefs.ML \ |
|
314 |
AxClasses/Lattice/OrdDefs.thy AxClasses/Lattice/OrdInsts.thy \ |
|
315 |
AxClasses/Lattice/Order.ML AxClasses/Lattice/Order.thy \ |
|
5712 | 316 |
AxClasses/Lattice/ROOT.ML |
2827 | 317 |
@$(ISATOOL) usedir -s AxClasses-Lattice $(OUT)/HOL AxClasses/Lattice |
4447 | 318 |
|
4518 | 319 |
|
320 |
## HOL-AxClasses-Tutorial |
|
321 |
||
7307 | 322 |
HOL-AxClasses-Tutorial: HOL $(LOG)/HOL-AxClasses-Tutorial.gz |
4518 | 323 |
|
324 |
$(LOG)/HOL-AxClasses-Tutorial.gz: $(OUT)/HOL \ |
|
325 |
AxClasses/Tutorial/BoolGroupInsts.thy AxClasses/Tutorial/Group.ML \ |
|
326 |
AxClasses/Tutorial/Group.thy AxClasses/Tutorial/Monoid.thy \ |
|
327 |
AxClasses/Tutorial/MonoidGroupInsts.thy \ |
|
328 |
AxClasses/Tutorial/ProdGroupInsts.thy AxClasses/Tutorial/Product.thy \ |
|
329 |
AxClasses/Tutorial/ProductInsts.thy AxClasses/Tutorial/ROOT.ML \ |
|
330 |
AxClasses/Tutorial/Semigroup.thy AxClasses/Tutorial/Semigroups.thy \ |
|
331 |
AxClasses/Tutorial/Sigs.thy AxClasses/Tutorial/Xor.ML \ |
|
332 |
AxClasses/Tutorial/Xor.thy |
|
2827 | 333 |
@$(ISATOOL) usedir -s AxClasses-Tutorial $(OUT)/HOL AxClasses/Tutorial |
2545 | 334 |
|
335 |
||
4518 | 336 |
## HOL-Quot |
337 |
||
338 |
HOL-Quot: HOL $(LOG)/HOL-Quot.gz |
|
2900 | 339 |
|
4518 | 340 |
$(LOG)/HOL-Quot.gz: $(OUT)/HOL Quot/FRACT.ML Quot/FRACT.thy \ |
341 |
Quot/HQUOT.ML Quot/HQUOT.thy Quot/NPAIR.ML Quot/NPAIR.thy Quot/PER.ML \ |
|
342 |
Quot/PER.thy Quot/PER0.ML Quot/PER0.thy Quot/ROOT.ML |
|
2900 | 343 |
@$(ISATOOL) usedir $(OUT)/HOL Quot |
344 |
||
345 |
||
4518 | 346 |
## HOL-ex |
2448 | 347 |
|
4518 | 348 |
HOL-ex: HOL $(LOG)/HOL-ex.gz |
2448 | 349 |
|
4518 | 350 |
$(LOG)/HOL-ex.gz: $(OUT)/HOL ex/BT.ML ex/BT.thy ex/Fib.ML ex/Fib.thy \ |
351 |
ex/InSort.ML ex/InSort.thy ex/MT.ML ex/MT.thy ex/NatSum.ML \ |
|
5199 | 352 |
ex/NatSum.thy ex/Primes.ML ex/Primes.thy ex/Primrec.ML \ |
353 |
ex/Primrec.thy ex/Puzzle.ML ex/Puzzle.thy ex/Qsort.ML ex/Qsort.thy \ |
|
354 |
ex/ROOT.ML ex/Recdefs.ML ex/Recdefs.thy ex/cla.ML ex/meson.ML \ |
|
355 |
ex/mesontest.ML ex/set.ML ex/Group.ML ex/Group.thy ex/IntRing.ML \ |
|
356 |
ex/IntRing.thy ex/IntRingDefs.ML ex/IntRingDefs.thy ex/Lagrange.ML \ |
|
357 |
ex/Lagrange.thy ex/Ring.ML ex/Ring.thy ex/StringEx.ML \ |
|
7085 | 358 |
ex/StringEx.thy ex/Tarski.ML ex/Tarski.thy \ |
7180 | 359 |
ex/BinEx.ML ex/BinEx.thy ex/svc_test.thy ex/svc_test.ML ex/MonoidGroup.thy \ |
5368 | 360 |
ex/PiSets.thy ex/PiSets.ML ex/LocaleGroup.thy ex/LocaleGroup.ML \ |
6737 | 361 |
ex/Antiquote.thy ex/Antiquote.ML ex/Points.thy |
2826 | 362 |
@$(ISATOOL) usedir $(OUT)/HOL ex |
2448 | 363 |
|
364 |
||
6445 | 365 |
## HOL-Isar_examples |
366 |
||
367 |
HOL-Isar_examples: HOL $(LOG)/HOL-Isar_examples.gz |
|
368 |
||
369 |
$(LOG)/HOL-Isar_examples.gz: $(OUT)/HOL Isar_examples/BasicLogic.thy \ |
|
6516 | 370 |
Isar_examples/Cantor.ML Isar_examples/Cantor.thy \ |
6764 | 371 |
Isar_examples/ExprCompiler.thy Isar_examples/Group.thy \ |
7435 | 372 |
Isar_examples/KnasterTarski.thy Isar_examples/MultisetOrder.thy \ |
7443 | 373 |
Isar_examples/MutilatedCheckerboard.thy Isar_examples/Peirce.thy \ |
374 |
Isar_examples/Summation.thy Isar_examples/ROOT.ML |
|
6445 | 375 |
@$(ISATOOL) usedir $(OUT)/HOL Isar_examples |
376 |
||
377 |
||
4518 | 378 |
## TLA |
379 |
||
380 |
TLA: HOL $(OUT)/TLA |
|
381 |
||
6254 | 382 |
$(OUT)/TLA: $(OUT)/HOL TLA/Action.ML TLA/Action.thy TLA/Init.ML \ |
383 |
TLA/Init.thy TLA/IntLemmas.ML TLA/Intensional.ML TLA/Intensional.thy \ |
|
384 |
TLA/ROOT.ML TLA/Stfun.ML TLA/Stfun.thy TLA/TLA.ML TLA/TLA.thy |
|
4518 | 385 |
@cd TLA; $(ISATOOL) usedir -b $(OUT)/HOL TLA |
386 |
||
387 |
||
388 |
## TLA-Inc |
|
389 |
||
390 |
TLA-Inc: TLA $(LOG)/TLA-Inc.gz |
|
391 |
||
392 |
$(LOG)/TLA-Inc.gz: $(OUT)/TLA TLA/Inc/Inc.thy TLA/Inc/Inc.ML \ |
|
393 |
TLA/Inc/Pcount.thy |
|
394 |
@cd TLA; $(ISATOOL) usedir $(OUT)/TLA Inc |
|
395 |
||
396 |
||
397 |
## TLA-Buffer |
|
398 |
||
399 |
TLA-Buffer: TLA $(LOG)/TLA-Buffer.gz |
|
2448 | 400 |
|
4518 | 401 |
$(LOG)/TLA-Buffer.gz: $(OUT)/TLA TLA/Buffer/Buffer.thy \ |
402 |
TLA/Buffer/Buffer.ML TLA/Buffer/DBuffer.thy TLA/Buffer/DBuffer.ML |
|
403 |
@cd TLA; $(ISATOOL) usedir $(OUT)/TLA Buffer |
|
404 |
||
405 |
||
406 |
## TLA-Memory |
|
407 |
||
408 |
TLA-Memory: TLA $(LOG)/TLA-Memory.gz |
|
4447 | 409 |
|
4518 | 410 |
$(LOG)/TLA-Memory.gz: $(OUT)/TLA TLA/Memory/MIParameters.thy \ |
411 |
TLA/Memory/MIlive.ML TLA/Memory/MIsafe.ML TLA/Memory/MemClerk.ML \ |
|
412 |
TLA/Memory/MemClerk.thy TLA/Memory/MemClerkParameters.ML \ |
|
413 |
TLA/Memory/MemClerkParameters.thy TLA/Memory/Memory.ML \ |
|
414 |
TLA/Memory/Memory.thy TLA/Memory/MemoryImplementation.ML \ |
|
415 |
TLA/Memory/MemoryImplementation.thy TLA/Memory/MemoryParameters.ML \ |
|
416 |
TLA/Memory/MemoryParameters.thy TLA/Memory/ProcedureInterface.ML \ |
|
417 |
TLA/Memory/ProcedureInterface.thy TLA/Memory/RPC.ML TLA/Memory/RPC.thy \ |
|
418 |
TLA/Memory/RPCMemoryParams.thy TLA/Memory/RPCParameters.ML \ |
|
419 |
TLA/Memory/RPCParameters.thy |
|
420 |
@cd TLA; $(ISATOOL) usedir $(OUT)/TLA Memory |
|
421 |
||
422 |
||
423 |
## clean |
|
4447 | 424 |
|
425 |
clean: |
|
7125 | 426 |
@rm -f $(OUT)/HOL $(OUT)/HOL-Real $(LOG)/HOL.gz $(LOG)/HOL-Subst.gz \ |
4518 | 427 |
$(LOG)/HOL-Induct.gz $(LOG)/HOL-IMP.gz $(LOG)/HOL-Hoare.gz \ |
7125 | 428 |
$(LOG)/HOL-Lex.gz $(LOG)/HOL-Auth.gz $(LOG)/HOL-UNITY.gz \ |
7307 | 429 |
$(LOG)/HOL-Modelcheck.gz $(LOG)/HOL-Lambda.gz \ |
7627 | 430 |
$(LOG)/HOL-W0.gz $(LOG)/HOL-MiniML.gz $(LOG)/HOL-BCV.gz $(LOG)/HOL-IOA.gz \ |
7307 | 431 |
$(LOG)/HOL-AxClasses-Group.gz \ |
432 |
$(LOG)/HOL-AxClasses-Lattice.gz \ |
|
4518 | 433 |
$(LOG)/HOL-AxClasses-Tutorial.gz $(LOG)/HOL-Quot.gz \ |
6445 | 434 |
$(LOG)/HOL-ex.gz $(LOG)/HOL-Isar_examples.gz $(OUT)/TLA \ |
435 |
$(LOG)/TLA.gz $(LOG)/TLA-Inc.gz $(LOG)/TLA-Buffer.gz \ |
|
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
7513
diff
changeset
|
436 |
$(LOG)/TLA-Memory.gz $(LOG)/HOL-Real-ex.gz \ |
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
7513
diff
changeset
|
437 |
$(LOG)/HOL-Real-HahnBanach.gz |