├── 100 └── arithmetic.ml ├── ProofTrace └── .gitignore ├── miz3 ├── make.ml ├── Samples │ ├── bug3.ml │ ├── NEEDS │ ├── bug2.ml │ ├── sample.ml │ └── wishes.ml ├── ERRORS ├── exrc └── test.ml ├── load_camlp4.ml ├── .gitattributes ├── Boyer_Moore └── testset │ ├── res1.pdf │ └── res2.pdf ├── load_camlp5.ml ├── Formal_ineqs ├── docs │ └── FormalVerifier.pdf ├── tests │ └── data │ │ └── gen_nat_data.py └── README.md ├── Proofrecording ├── diffs │ ├── proofobjects_trt.ml │ ├── proofobjects_dummy.ml │ └── proofobjects_init.ml └── tools │ ├── startcore.ml │ ├── src │ └── Makefile │ └── init.ml ├── EC └── README ├── RichterHilbertAxiomGeometry └── miz3 │ └── make.ml ├── Rqe └── matinsert_thms.ml ├── Help ├── then_.hlp ├── thenc_.hlp ├── orelse_.hlp ├── orelsec_.hlp ├── thenl_.hlp ├── f_f_.hlp ├── then_tcl_.hlp ├── orelse_tcl_.hlp ├── length.hlp ├── NO_CONV.hlp ├── I.hlp ├── rev.hlp ├── K.hlp ├── type_of.hlp ├── o.hlp ├── W.hlp ├── hd.hlp ├── C.hlp ├── last.hlp ├── omit.hlp ├── dpty.hlp ├── tl.hlp ├── name.hlp ├── exactly.hlp ├── mk_select.hlp ├── dest_conj.hlp ├── butlast.hlp ├── is_conj.hlp ├── is_disj.hlp ├── is_neg.hlp ├── replicate.hlp ├── zip.hlp ├── unzip.hlp ├── is_cond.hlp ├── is_list.hlp ├── is_imp.hlp ├── is_select.hlp ├── PRINT_GOAL_TAC.hlp ├── dest_neg.hlp ├── by.hlp ├── body.hlp ├── is_forall.hlp ├── is_cons.hlp ├── is_exists.hlp ├── mem.hlp ├── rhs.hlp ├── bndvar.hlp ├── mk_cond.hlp ├── mk_eq.hlp ├── constants.hlp ├── is_numeral.hlp ├── lhs.hlp ├── F_F.hlp ├── mk_pair.hlp ├── DISJ1_TAC.hlp ├── uncurry.hlp ├── DISJ2_TAC.hlp ├── use_file.hlp ├── dest_eq.hlp ├── is_uexists.hlp ├── isnum.hlp ├── hol_version.hlp ├── print_thm.hlp ├── dest_disj.hlp ├── map.hlp ├── metisverb.hlp ├── dest_list.hlp ├── issep.hlp ├── rotate.hlp ├── type_abbrevs.hlp ├── bool_ty.hlp ├── mk_neg.hlp ├── dest_thm.hlp ├── flat.hlp ├── mk_imp.hlp ├── report.hlp ├── copverb.hlp ├── dest_select.hlp ├── mk_disj.hlp ├── rand.hlp ├── dest_cond.hlp ├── dest_imp.hlp ├── dest_uexists.hlp ├── isspace.hlp ├── num_0.hlp ├── num_1.hlp ├── num_2.hlp ├── num_10.hlp ├── EQT_INTRO.hlp ├── mk_conj.hlp ├── tyvars.hlp ├── dest_forall.hlp ├── mk_goalstate.hlp ├── print_num.hlp ├── dest_var.hlp ├── is_pair.hlp ├── el.hlp ├── is_var.hlp ├── ALL_CONV.hlp ├── aty.hlp ├── bty.hlp ├── can.hlp ├── dest_exists.hlp ├── isalpha.hlp ├── isbra.hlp ├── pp_print_thm.hlp ├── print_goal.hlp ├── temp_path.hlp ├── ANTS_TAC.hlp ├── filter.hlp ├── is_prefix.hlp ├── .upto.hlp ├── DENUMERAL.hlp ├── isalnum.hlp ├── mk_exists.hlp ├── mk_forall.hlp ├── map2.hlp ├── mk_uexists.hlp ├── pow2.hlp ├── remove.hlp ├── the_inductive_types.hlp ├── dest_pair.hlp ├── is_abs.hlp ├── list_mk_abs.hlp ├── pow10.hlp ├── variables.hlp ├── CONJUNCT1.hlp ├── CONJUNCT2.hlp ├── find.hlp ├── CONJ.hlp ├── DISJ2.hlp ├── infixes.hlp ├── is_comb.hlp ├── strip_forall.hlp ├── concl.hlp ├── strip_exists.hlp ├── chop_list.hlp ├── distinctness_store.hlp ├── get_const_type.hlp ├── hyp.hlp ├── injectivity_store.hlp ├── print_goalstack.hlp ├── DISJ1.hlp ├── print_qterm.hlp ├── is_reserved_word.hlp ├── print_fpf.hlp ├── EQF_ELIM.hlp ├── EQT_ELIM.hlp ├── find_term.hlp ├── print_qtype.hlp ├── end_itlist.hlp ├── parse_as_prefix.hlp ├── current_goalstack.hlp ├── dest_abs.hlp ├── dest_cons.hlp ├── mk_var.hlp ├── unions.hlp ├── uniq.hlp ├── delete_parser.hlp ├── is_let.hlp ├── rev_itlist.hlp ├── curry.hlp ├── explode.hlp ├── installed_parsers.hlp ├── pp_print_qterm.hlp ├── rator.hlp ├── CONJ_TAC.hlp ├── pp_print_term.hlp ├── SYM_CONV.hlp ├── dest_vartype.hlp ├── empty_ss.hlp ├── flush_goalstack.hlp ├── mk_abs.hlp ├── null_meta.hlp ├── parse_type.hlp ├── NO_THEN.hlp ├── REFL.hlp ├── REFL_TAC.hlp ├── string_of_type.hlp ├── unparse_as_prefix.hlp ├── warn.hlp ├── BETAS_CONV.hlp ├── dest_const.hlp ├── mk_cons.hlp ├── pp_print_qtype.hlp ├── pp_print_type.hlp ├── top_realgoal.hlp ├── mk_string.hlp ├── partition.hlp ├── EQF_INTRO.hlp ├── UNDISCH.hlp ├── equals_thm.hlp ├── is_vartype.hlp ├── mk_finty.hlp ├── allpairs.hlp ├── it.hlp ├── setify.hlp ├── UNDISCH_TAC.hlp ├── assoc.hlp ├── dest_string.hlp ├── mk_fun_ty.hlp ├── strip_abs.hlp ├── subset.hlp ├── try_user_parser.hlp ├── tryfind.hlp ├── is_type.hlp ├── p.hlp ├── print_term.hlp ├── reserve_words.hlp ├── the_interface.hlp ├── loaded_files.hlp ├── print_type.hlp ├── itlist.hlp ├── list_mk_gabs.hlp ├── mk_iff.hlp ├── CONTRAPOS_CONV.hlp ├── install_parser.hlp ├── pp_print_num.hlp ├── ASSUME.hlp ├── CONJ_PAIR.hlp ├── refine.hlp ├── null_inst.hlp ├── rev_assoc.hlp ├── equals_goal.hlp ├── is_eq.hlp ├── REDEPTH_SQCONV.hlp ├── TRY_CONV.hlp ├── implode.hlp ├── insert.hlp ├── lcm_num.hlp ├── set_eq.hlp ├── AP_TERM_TAC.hlp ├── remove_interface.hlp ├── CHEAT_TAC.hlp ├── LAND_CONV.hlp ├── TOP_DEPTH_SQCONV.hlp ├── USE_THEN.hlp ├── frees.hlp ├── index.hlp ├── inst_goal.hlp ├── basic_ss.hlp ├── dest_finty.hlp ├── list_mk_forall.hlp ├── parses_as_binder.hlp ├── reduce_interface.hlp ├── thm_frees.hlp ├── type_unify.hlp ├── DEPTH_SQCONV.hlp ├── is_intconst.hlp ├── is_setenum.hlp ├── repeat.hlp ├── ASM.hlp ├── ORELSE_TCL.hlp ├── ABS.hlp ├── binders.hlp ├── exists.hlp ├── forall.hlp ├── gcd.hlp ├── preterm_of_term.hlp ├── quotexpander.hlp ├── decreasing.hlp ├── name_of.hlp ├── pretype_of_type.hlp ├── AP_THM_TAC.hlp ├── GEN_SIMPLIFY_CONV.hlp ├── ONCE_DEPTH_SQCONV.hlp ├── dest_let.hlp ├── help_path.hlp ├── subtract.hlp ├── union.hlp ├── EQ_TAC.hlp ├── SIMP_RULE.hlp ├── list_mk_conj.hlp ├── mk_char.hlp ├── REPLICATE_TAC.hlp ├── extend_basic_rewrites.hlp ├── load_path.hlp ├── mk_comb.hlp ├── parse_preterm.hlp ├── parse_pretype.hlp ├── pp_print_fpf.hlp ├── prefixes.hlp ├── COMB_CONV.hlp ├── NOT_ELIM.hlp ├── TAC_PROOF.hlp ├── THEN_TCL.hlp ├── dest_binary.hlp ├── dest_type.hlp ├── list_mk_disj.hlp ├── nsplit.hlp ├── top_goal.hlp ├── dest_iff.hlp ├── is_realintconst.hlp ├── strip_comb.hlp ├── new_constant.hlp ├── unparse_as_infix.hlp └── CHANGED_TAC.hlp ├── META ├── LP_arith ├── make.ml ├── README └── Makefile ├── Tutorial ├── HOL_basics.ml └── Wellfounded_induction.ml ├── .gitignore ├── Minisat ├── make.ml └── CREDITS ├── load_camlp5_topfind.ml ├── Cadical ├── make.ml └── test.ml ├── GL └── README.md ├── Functionspaces └── README ├── Geometric_Algebra ├── README └── make.ml ├── Divstep └── Makefile ├── QBF ├── make.ml └── mygraph.ml ├── IEEE └── make.ml ├── Complex └── complex_real.ml ├── Mizarlight └── Makefile ├── pa_j └── README ├── bignum_num.ml └── Unity └── make.ml /ProofTrace/.gitignore: -------------------------------------------------------------------------------- 1 | *.names 2 | *.proofs 3 | *.theorems 4 | -------------------------------------------------------------------------------- /miz3/make.ml: -------------------------------------------------------------------------------- 1 | #load "unix.cma";; 2 | loadt "miz3/miz3.ml";; 3 | -------------------------------------------------------------------------------- /load_camlp4.ml: -------------------------------------------------------------------------------- 1 | Topdirs.dir_load Format.std_formatter "camlp4o.cma";; 2 | -------------------------------------------------------------------------------- /miz3/Samples/bug3.ml: -------------------------------------------------------------------------------- 1 | let FOO = thm `; 2 | thus T; 3 | `;; 4 | -------------------------------------------------------------------------------- /.gitattributes: -------------------------------------------------------------------------------- 1 | *.ml linguist-language=OCaml 2 | *.hl linguist-language=OCaml -------------------------------------------------------------------------------- /Boyer_Moore/testset/res1.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/fblanqui/hol-light/master/Boyer_Moore/testset/res1.pdf -------------------------------------------------------------------------------- /Boyer_Moore/testset/res2.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/fblanqui/hol-light/master/Boyer_Moore/testset/res2.pdf -------------------------------------------------------------------------------- /load_camlp5.ml: -------------------------------------------------------------------------------- 1 | Topdirs.dir_directory "+camlp5";; 2 | Topdirs.dir_load Format.std_formatter "camlp5o.cma";; 3 | -------------------------------------------------------------------------------- /miz3/Samples/NEEDS: -------------------------------------------------------------------------------- 1 | needs "Examples/transc.ml";; 2 | needs "Examples/sos.ml";; 3 | needs "Examples/prime.ml";; 4 | -------------------------------------------------------------------------------- /Formal_ineqs/docs/FormalVerifier.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/fblanqui/hol-light/master/Formal_ineqs/docs/FormalVerifier.pdf -------------------------------------------------------------------------------- /Proofrecording/diffs/proofobjects_trt.ml: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/fblanqui/hol-light/master/Proofrecording/diffs/proofobjects_trt.ml -------------------------------------------------------------------------------- /Proofrecording/diffs/proofobjects_dummy.ml: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/fblanqui/hol-light/master/Proofrecording/diffs/proofobjects_dummy.ml -------------------------------------------------------------------------------- /miz3/Samples/bug2.ml: -------------------------------------------------------------------------------- 1 | let FOO = thm `; 2 | let P be num->bool; 3 | assume !x. P x \/ ~P x; 4 | thus (~ ~ ?x. P x) ==> ?x. P x; 5 | `;; 6 | -------------------------------------------------------------------------------- /EC/README: -------------------------------------------------------------------------------- 1 | ELLIPTIC CURVES OVER ARBITRARY FIELDS 2 | 3 | All code in this directory is distributed under the same license as HOL Light. 4 | -------------------------------------------------------------------------------- /RichterHilbertAxiomGeometry/miz3/make.ml: -------------------------------------------------------------------------------- 1 | #load "unix.cma";; 2 | loadt "miz3/miz3.ml";; 3 | loadt "RichterHilbertAxiomGeometry/miz3/HilbertAxiom.ml";; 4 | -------------------------------------------------------------------------------- /Rqe/matinsert_thms.ml: -------------------------------------------------------------------------------- 1 | 2 | let matinsert_lem0 = prove_by_refinement( 3 | `!S. (!x. P x) ==> (!x. S x ==> P x)`, 4 | (* {{{ Proof *) 5 | [MESON_TAC[]]);; 6 | (* }}} *) 7 | -------------------------------------------------------------------------------- /Help/then_.hlp: -------------------------------------------------------------------------------- 1 | \DOC then_ 2 | 3 | \TYPE {then_ : tactic -> tactic -> tactic} 4 | 5 | \SYNOPSIS 6 | Non-infix version of {THEN}. 7 | 8 | \SEEALSO 9 | THEN. 10 | 11 | \ENDDOC 12 | -------------------------------------------------------------------------------- /Help/thenc_.hlp: -------------------------------------------------------------------------------- 1 | \DOC thenc_ 2 | 3 | \TYPE {thenc_ : conv -> conv -> conv} 4 | 5 | \SYNOPSIS 6 | Non-infix version of {THENC}. 7 | 8 | \SEEALSO 9 | THENC. 10 | 11 | \ENDDOC 12 | -------------------------------------------------------------------------------- /META: -------------------------------------------------------------------------------- 1 | version = "3.0.0+" 2 | description = "The HOL Light interactive theorem prover" 3 | requires = "camlp5 zarith" 4 | archive(byte) = "hol_lib.cma" 5 | archive(native) = "hol_lib.cmxa" 6 | -------------------------------------------------------------------------------- /Help/orelse_.hlp: -------------------------------------------------------------------------------- 1 | \DOC orelse_ 2 | 3 | \TYPE {orelse_ : tactic -> tactic -> tactic} 4 | 5 | \SYNOPSIS 6 | Non-infix version of {ORELSE}. 7 | 8 | \SEEALSO 9 | ORELSE. 10 | 11 | \ENDDOC 12 | -------------------------------------------------------------------------------- /Help/orelsec_.hlp: -------------------------------------------------------------------------------- 1 | \DOC orelsec_ 2 | 3 | \TYPE {orelsec_ : conv -> conv -> conv} 4 | 5 | \SYNOPSIS 6 | Non-infix version of {ORELSEC}. 7 | 8 | \SEEALSO 9 | ORELSEC. 10 | 11 | \ENDDOC 12 | -------------------------------------------------------------------------------- /Help/thenl_.hlp: -------------------------------------------------------------------------------- 1 | \DOC thenl_ 2 | 3 | \TYPE {thenl_ : tactic -> tactic list -> tactic} 4 | 5 | \SYNOPSIS 6 | Non-infix version of {THENL}. 7 | 8 | \SEEALSO 9 | THENL. 10 | 11 | \ENDDOC 12 | -------------------------------------------------------------------------------- /Help/f_f_.hlp: -------------------------------------------------------------------------------- 1 | \DOC f_f_ 2 | 3 | \TYPE {f_f_ : ('a -> 'b) -> ('c -> 'd) -> 'a * 'c -> 'b * 'd} 4 | 5 | \SYNOPSIS 6 | Non-infix version of {F_F}. 7 | 8 | \SEEALSO 9 | F_F. 10 | 11 | \ENDDOC 12 | -------------------------------------------------------------------------------- /LP_arith/make.ml: -------------------------------------------------------------------------------- 1 | loadt "LP_arith/lp_arith.ml";; 2 | loadt "LP_arith/lp_tests.ml";; 3 | 4 | time LP_ARITH rec_seq;; 5 | time LP_ARITH test_std;; 6 | 7 | time REAL_ARITH rec_seq;; 8 | time REAL_ARITH test_std;; 9 | -------------------------------------------------------------------------------- /Help/then_tcl_.hlp: -------------------------------------------------------------------------------- 1 | \DOC then_tcl_ 2 | 3 | \TYPE {then_tcl_ : thm_tactical -> thm_tactical -> thm_tactical} 4 | 5 | \SYNOPSIS 6 | Non-infix version of {THEN_TCL}. 7 | 8 | \SEEALSO 9 | THEN_TCL. 10 | 11 | \ENDDOC 12 | -------------------------------------------------------------------------------- /Help/orelse_tcl_.hlp: -------------------------------------------------------------------------------- 1 | \DOC orelse_tcl_ 2 | 3 | \TYPE {orelse_tcl_ : thm_tactical -> thm_tactical -> thm_tactical} 4 | 5 | \SYNOPSIS 6 | Non-infix version of {ORELSE_TCL}. 7 | 8 | \SEEALSO 9 | ORELSE_TCL. 10 | 11 | \ENDDOC 12 | -------------------------------------------------------------------------------- /Help/length.hlp: -------------------------------------------------------------------------------- 1 | \DOC length 2 | 3 | \TYPE {length : 'a list -> int} 4 | 5 | \SYNOPSIS 6 | Computes the length of a list: {length [x1;...;xn]} returns {n}. 7 | 8 | \KEYWORDS 9 | list. 10 | 11 | \FAILURE 12 | Never fails. 13 | 14 | \ENDDOC 15 | -------------------------------------------------------------------------------- /Tutorial/HOL_basics.ml: -------------------------------------------------------------------------------- 1 | ARITH_RULE 2 | `(a * x + b * y + a * y) EXP 3 + (b * x) EXP 3 + 3 | (a * x + b * y + b * x) EXP 3 + (a * y) EXP 3 = 4 | (a * x + a * y + b * x) EXP 3 + (b * y) EXP 3 + 5 | (a * y + b * y + b * x) EXP 3 + (a * x) EXP 3`;; 6 | -------------------------------------------------------------------------------- /Help/NO_CONV.hlp: -------------------------------------------------------------------------------- 1 | \DOC NO_CONV 2 | 3 | \TYPE {NO_CONV : conv} 4 | 5 | \SYNOPSIS 6 | Conversion that always fails. 7 | 8 | \KEYWORDS 9 | conversion. 10 | 11 | \FAILURE 12 | {NO_CONV} always fails. 13 | 14 | \SEEALSO 15 | ALL_CONV. 16 | 17 | \ENDDOC 18 | -------------------------------------------------------------------------------- /Proofrecording/tools/startcore.ml: -------------------------------------------------------------------------------- 1 | 2 | set_jrh_lexer;; (* Uppercase idents *) 3 | 4 | Gc.set { (Gc.get()) with Gc.stack_limit = 16777216 };; (* Up the stack size *) 5 | 6 | include Num;; 7 | 8 | Sys.catch_break true;; 9 | -------------------------------------------------------------------------------- /Help/I.hlp: -------------------------------------------------------------------------------- 1 | \DOC I 2 | 3 | \TYPE {I : 'a -> 'a} 4 | 5 | \SYNOPSIS 6 | Performs identity operation: {I x} = {x}. 7 | 8 | \KEYWORDS 9 | combinator, identity. 10 | 11 | \FAILURE 12 | Never fails. 13 | 14 | \SEEALSO 15 | C, K, F_F, o, W. 16 | 17 | \ENDDOC 18 | -------------------------------------------------------------------------------- /Help/rev.hlp: -------------------------------------------------------------------------------- 1 | \DOC rev 2 | 3 | \TYPE {rev : 'a list -> 'a list} 4 | 5 | \SYNOPSIS 6 | Reverses a list. 7 | 8 | \KEYWORDS 9 | list. 10 | 11 | \DESCRIBE 12 | {rev [x1;...;xn]} returns {[xn;...;x1]}. 13 | 14 | \FAILURE 15 | Never fails. 16 | 17 | \ENDDOC 18 | -------------------------------------------------------------------------------- /.gitignore: -------------------------------------------------------------------------------- 1 | .DS_Store 2 | .vscode 3 | _opam 4 | *.a 5 | *.cma 6 | *.cmi 7 | *.cmo 8 | *.o 9 | *.cmx 10 | *.cmxa 11 | pa_j.ml 12 | update_database.ml 13 | ocaml-hol 14 | hol.sh 15 | hol_lib_inlined.ml 16 | unit_tests_inlined.ml 17 | unit_tests.byte 18 | unit_tests.native -------------------------------------------------------------------------------- /Help/K.hlp: -------------------------------------------------------------------------------- 1 | \DOC K 2 | 3 | \TYPE {K : 'a -> 'b -> 'a} 4 | 5 | \SYNOPSIS 6 | Forms a constant function: {(K x) y} = {x}. 7 | 8 | \KEYWORDS 9 | combinator, constant. 10 | 11 | \FAILURE 12 | Never fails. 13 | 14 | \SEEALSO 15 | C, F_F, I, o, W. 16 | 17 | \ENDDOC 18 | -------------------------------------------------------------------------------- /Help/type_of.hlp: -------------------------------------------------------------------------------- 1 | \DOC type_of 2 | 3 | \TYPE {type_of : term -> hol_type} 4 | 5 | \SYNOPSIS 6 | Returns the type of a term. 7 | 8 | \FAILURE 9 | Never fails. 10 | 11 | \EXAMPLE 12 | { 13 | # type_of `T`;; 14 | val it : hol_type = `:bool` 15 | } 16 | \ENDDOC 17 | -------------------------------------------------------------------------------- /miz3/ERRORS: -------------------------------------------------------------------------------- 1 | 1: inference error 2 | 2: inference time-out 3 | 3: skeleton error 4 | 4: unknown label 5 | 5: underspecified types hol 6 | 6: unbound free variables hol 7 | 7: syntax error justification or ocaml 8 | 8: syntax or type error hol 9 | 9: syntax error mizar 10 | -------------------------------------------------------------------------------- /Help/o.hlp: -------------------------------------------------------------------------------- 1 | \DOC o 2 | 3 | \TYPE {o : ('a -> 'b) -> ('c -> 'a) -> 'c -> 'b} 4 | 5 | \SYNOPSIS 6 | Composes two functions: {(f o g) x} = {f (g x)}. 7 | 8 | \KEYWORDS 9 | combinator. 10 | 11 | \FAILURE 12 | Never fails. 13 | 14 | \SEEALSO 15 | C, F_F, I, K, W. 16 | 17 | \ENDDOC 18 | -------------------------------------------------------------------------------- /Help/W.hlp: -------------------------------------------------------------------------------- 1 | \DOC W 2 | 3 | \TYPE {W : ('a -> 'a -> 'b) -> 'a -> 'b} 4 | 5 | \SYNOPSIS 6 | Duplicates function argument : {W f x} = {f x x}. 7 | 8 | \KEYWORDS 9 | combinator, duplicate. 10 | 11 | \FAILURE 12 | Never fails. 13 | 14 | \SEEALSO 15 | C, F_F, I, K, o. 16 | 17 | \ENDDOC 18 | -------------------------------------------------------------------------------- /Help/hd.hlp: -------------------------------------------------------------------------------- 1 | \DOC hd 2 | 3 | \TYPE {hd : 'a list -> 'a} 4 | 5 | \SYNOPSIS 6 | Computes the first element (the head) of a list. 7 | 8 | \DESCRIBE 9 | {hd [x1;...;xn]} returns {x1}. 10 | 11 | \FAILURE 12 | Fails with {hd} if the list is empty. 13 | 14 | \SEEALSO 15 | tl, el. 16 | 17 | \ENDDOC 18 | -------------------------------------------------------------------------------- /Help/C.hlp: -------------------------------------------------------------------------------- 1 | \DOC C 2 | 3 | \TYPE {C : ('a -> 'b -> 'c) -> 'b -> 'a -> 'c} 4 | 5 | \SYNOPSIS 6 | Permutes first two arguments to curried function: {C f x y} = {f y x}. 7 | 8 | \KEYWORDS 9 | combinator, permute. 10 | 11 | \FAILURE 12 | Never fails. 13 | 14 | \SEEALSO 15 | F_F, I, K, W. 16 | 17 | \ENDDOC 18 | -------------------------------------------------------------------------------- /Help/last.hlp: -------------------------------------------------------------------------------- 1 | \DOC last 2 | 3 | \TYPE {last : 'a list -> 'a} 4 | 5 | \SYNOPSIS 6 | Computes the last element of a list. 7 | 8 | \DESCRIBE 9 | {last [x1;...;xn]} returns {xn}. 10 | 11 | \FAILURE 12 | Fails with {last} if the list is empty. 13 | 14 | \SEEALSO 15 | butlast, hd, tl, el. 16 | 17 | \ENDDOC 18 | -------------------------------------------------------------------------------- /LP_arith/README: -------------------------------------------------------------------------------- 1 | HOL interface to cddlib to provide a faster linear prover. 2 | 3 | http://www.ifor.math.ethz.ch/~fukuda/cdd_home/ 4 | 5 | Once cddlib is installed, just do "make" in this directory, 6 | and then load "LP_arith/lp_arith.ml" into HOL Light. 7 | 8 | (c) Lars Schewe (schewe@mathematik.tu-darmstadt.de), 2007 9 | -------------------------------------------------------------------------------- /miz3/exrc: -------------------------------------------------------------------------------- 1 | " put "source .../miz3/exrc" in your .exrc 2 | " in which ".../miz3" is the path of your miz3 directory 3 | :map  :w! :!miz3 % :e 4 | :map  A {/. !}miz3f / $ D^ 5 | :map  {/. !}miz3f /# l 6 | :map  :g/^::/d 7 | :map  /# 8 | :map  ?# 9 | :map  ^f`s(parse_term "/` s") 10 | -------------------------------------------------------------------------------- /Help/omit.hlp: -------------------------------------------------------------------------------- 1 | \DOC omit 2 | 3 | \TYPE {omit : term -> term} 4 | 5 | \SYNOPSIS 6 | Omit anything satisfying the given {search} query. 7 | 8 | \DESCRIBE 9 | The function {omit} is intended for use solely with the {search} function. 10 | 11 | \FAILURE 12 | Never fails. 13 | 14 | \SEEALSO 15 | search. 16 | 17 | \ENDDOC 18 | -------------------------------------------------------------------------------- /Help/dpty.hlp: -------------------------------------------------------------------------------- 1 | \DOC dpty 2 | 3 | \TYPE {dpty : pretype} 4 | 5 | \SYNOPSIS 6 | Dummy pretype. 7 | 8 | \DESCRIBE 9 | This is a dummy pretype, intended as a placeholder until the correct one is 10 | discovered. 11 | 12 | \FAILURE 13 | Not applicable. 14 | 15 | \SEEALSO 16 | pretype_of_type, type_of_pretype. 17 | 18 | \ENDDOC 19 | -------------------------------------------------------------------------------- /Help/tl.hlp: -------------------------------------------------------------------------------- 1 | \DOC tl 2 | 3 | \TYPE {tl : 'a list -> 'a list} 4 | 5 | \SYNOPSIS 6 | Computes the tail of a list (the original list less the first element). 7 | 8 | \DESCRIBE 9 | {tl [x1;...;xn]} returns {[x2;...;xn]}. 10 | 11 | \FAILURE 12 | Fails with {tl} if the list is empty. 13 | 14 | \SEEALSO 15 | hd, el. 16 | 17 | \ENDDOC 18 | -------------------------------------------------------------------------------- /Help/name.hlp: -------------------------------------------------------------------------------- 1 | \DOC name 2 | 3 | \TYPE {name : string -> term} 4 | 5 | \SYNOPSIS 6 | Query to {search} for a theorem whose name contains a string. 7 | 8 | \DESCRIBE 9 | The function {name} is intended for use solely with the {search} function. 10 | 11 | \FAILURE 12 | Never fails. 13 | 14 | \SEEALSO 15 | search. 16 | 17 | \ENDDOC 18 | -------------------------------------------------------------------------------- /Help/exactly.hlp: -------------------------------------------------------------------------------- 1 | \DOC exactly 2 | 3 | \TYPE {exactly : term -> term} 4 | 5 | \SYNOPSIS 6 | Query to {search} for a term alpha-equivalent to pattern. 7 | 8 | \DESCRIBE 9 | The function {exactly} is intended for use solely with the {search} function. 10 | 11 | \FAILURE 12 | Never fails. 13 | 14 | \SEEALSO 15 | search. 16 | 17 | \ENDDOC 18 | -------------------------------------------------------------------------------- /Help/mk_select.hlp: -------------------------------------------------------------------------------- 1 | \DOC mk_select 2 | 3 | \TYPE {mk_select : term * term -> term} 4 | 5 | \SYNOPSIS 6 | Constructs a choice binding. 7 | 8 | \DESCRIBE 9 | The call {mk_select(`v`,`t`)} returns the choice term {`@v. t`}. 10 | 11 | \FAILURE 12 | Fails if {v} is not a variable. 13 | 14 | \SEEALSO 15 | is_slect, mk_select. 16 | 17 | \ENDDOC 18 | -------------------------------------------------------------------------------- /Minisat/make.ml: -------------------------------------------------------------------------------- 1 | #load "str.cma";; 2 | loads "Minisat/sat_common_tools.ml";; 3 | loads "Minisat/dimacs_tools.ml";; 4 | loads "Minisat/sat_solvers.ml";; 5 | loads "Minisat/sat_script.ml";; 6 | loads "Minisat/sat_tools.ml";; 7 | loads "Minisat/minisat_parse.ml";; 8 | loads "Minisat/minisat_resolve.ml";; 9 | loads "Minisat/minisat_prove.ml";; 10 | -------------------------------------------------------------------------------- /Help/dest_conj.hlp: -------------------------------------------------------------------------------- 1 | \DOC dest_conj 2 | 3 | \TYPE {dest_conj : term -> term * term} 4 | 5 | \SYNOPSIS 6 | Term destructor for conjunctions. 7 | 8 | \DESCRIBE 9 | {dest_conj(`t1 /\ t2`)} returns {(`t1`,`t2`)}. 10 | 11 | \FAILURE 12 | Fails with {dest_conj} if term is not a conjunction. 13 | 14 | \SEEALSO 15 | is_conj, mk_conj. 16 | 17 | \ENDDOC 18 | -------------------------------------------------------------------------------- /Help/butlast.hlp: -------------------------------------------------------------------------------- 1 | \DOC butlast 2 | 3 | \TYPE {butlast : 'a list -> 'a list} 4 | 5 | \SYNOPSIS 6 | Computes the sub-list of a list consisting of all but the last element. 7 | 8 | \DESCRIBE 9 | {butlast [x1;...;xn]} returns {[x1;...;x(n-1)]}. 10 | 11 | \FAILURE 12 | Fails if the list is empty. 13 | 14 | \SEEALSO 15 | last, hd, tl, el. 16 | 17 | \ENDDOC 18 | -------------------------------------------------------------------------------- /Help/is_conj.hlp: -------------------------------------------------------------------------------- 1 | \DOC is_conj 2 | 3 | \TYPE {is_conj : term -> bool} 4 | 5 | \SYNOPSIS 6 | Tests a term to see if it is a conjunction. 7 | 8 | \DESCRIBE 9 | {is_conj `t1 /\ t2`} returns {true}. If the term is not a conjunction the 10 | result is {false}. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \SEEALSO 16 | dest_conj, mk_conj. 17 | 18 | \ENDDOC 19 | -------------------------------------------------------------------------------- /Help/is_disj.hlp: -------------------------------------------------------------------------------- 1 | \DOC is_disj 2 | 3 | \TYPE {is_disj : term -> bool} 4 | 5 | \SYNOPSIS 6 | Tests a term to see if it is a disjunction. 7 | 8 | \DESCRIBE 9 | {is_disj `t1 \/ t2`} returns {true}. If the term is not a disjunction the 10 | result is {false}. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \SEEALSO 16 | dest_disj, mk_disj. 17 | 18 | \ENDDOC 19 | -------------------------------------------------------------------------------- /Help/is_neg.hlp: -------------------------------------------------------------------------------- 1 | \DOC is_neg 2 | 3 | \TYPE {is_neg : term -> bool} 4 | 5 | \SYNOPSIS 6 | Tests a term to see if it is a logical negation. 7 | 8 | \DESCRIBE 9 | {is_neg `~t`} returns {true}. If the term is not a logical negation the result 10 | is {false}. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \SEEALSO 16 | dest_neg, mk_neg. 17 | 18 | \ENDDOC 19 | -------------------------------------------------------------------------------- /Help/replicate.hlp: -------------------------------------------------------------------------------- 1 | \DOC replicate 2 | 3 | \TYPE {replicate : 'a -> int -> 'a list} 4 | 5 | \SYNOPSIS 6 | Makes a list consisting of a value replicated a specified number of times. 7 | 8 | \DESCRIBE 9 | {replicate x n} returns {[x;...;x]}, a list of length {n}. 10 | 11 | \FAILURE 12 | Fails if number of replications is less than zero. 13 | 14 | \ENDDOC 15 | -------------------------------------------------------------------------------- /Help/zip.hlp: -------------------------------------------------------------------------------- 1 | \DOC zip 2 | 3 | \TYPE {zip : 'a list -> 'b list -> ('a * 'b) list} 4 | 5 | \SYNOPSIS 6 | Converts a pair of lists into a list of pairs. 7 | 8 | \DESCRIBE 9 | {zip [x1;...;xn] [y1;...;yn]} returns {[(x1,y1);...;(xn,yn)]}. 10 | 11 | \FAILURE 12 | Fails if the two lists are of different lengths. 13 | 14 | \SEEALSO 15 | unzip. 16 | 17 | \ENDDOC 18 | -------------------------------------------------------------------------------- /load_camlp5_topfind.ml: -------------------------------------------------------------------------------- 1 | Topdirs.dir_use Format.std_formatter "topfind";; 2 | 3 | (* Don't load compiler-libs.common because it conflicts with toplevel's Env. *) 4 | (* See also: https://github.com/ocaml/ocaml/issues/12271 *) 5 | Topfind.don't_load ["compiler-libs.common"];; 6 | 7 | Topfind.load_deeply ["camlp5"];; 8 | Topdirs.dir_load Format.std_formatter "camlp5o.cma";; 9 | -------------------------------------------------------------------------------- /Help/unzip.hlp: -------------------------------------------------------------------------------- 1 | \DOC unzip 2 | 3 | \TYPE {unzip : ('a * 'b) list -> 'a list * 'b list} 4 | 5 | \SYNOPSIS 6 | Converts a list of pairs into a pair of lists. 7 | 8 | \KEYWORDS 9 | list. 10 | 11 | \DESCRIBE 12 | {unzip [(x1,y1);...;(xn,yn)]} returns {([x1;...;xn],[y1;...;yn])}. 13 | 14 | \FAILURE 15 | Never fails. 16 | 17 | \SEEALSO 18 | zip. 19 | 20 | \ENDDOC 21 | -------------------------------------------------------------------------------- /Help/is_cond.hlp: -------------------------------------------------------------------------------- 1 | \DOC is_cond 2 | 3 | \TYPE {is_cond : term -> bool} 4 | 5 | \SYNOPSIS 6 | Tests a term to see if it is a conditional. 7 | 8 | \DESCRIBE 9 | {is_cond `if t then t1 else t2`} returns {true}. If the term is not a 10 | conditional the result is {false}. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \SEEALSO 16 | mk_cond, dest_cond. 17 | 18 | \ENDDOC 19 | -------------------------------------------------------------------------------- /Help/is_list.hlp: -------------------------------------------------------------------------------- 1 | \DOC is_list 2 | 3 | \TYPE {is_list : term -> bool} 4 | 5 | \SYNOPSIS 6 | Tests a term to see if it is a list. 7 | 8 | \DESCRIBE 9 | {is_list} returns {true} of a term representing a list. Otherwise it returns 10 | {false}. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \SEEALSO 16 | dest_cons, dest_list, is_cons, mk_cons, mk_list. 17 | 18 | \ENDDOC 19 | -------------------------------------------------------------------------------- /Help/is_imp.hlp: -------------------------------------------------------------------------------- 1 | \DOC is_imp 2 | 3 | \TYPE {is_imp : term -> bool} 4 | 5 | \SYNOPSIS 6 | Tests if a term is an application of implication. 7 | 8 | \DESCRIBE 9 | The call {is_imp t} returns {true} if {t} is of the form {p ==> q} for some {p} 10 | and {q}, and returns {false} otherwise. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \SEEALSO 16 | dest_imp. 17 | 18 | \ENDDOC 19 | -------------------------------------------------------------------------------- /Help/is_select.hlp: -------------------------------------------------------------------------------- 1 | \DOC is_select 2 | 3 | \TYPE {is_select : term -> bool} 4 | 5 | \SYNOPSIS 6 | Tests a term to see if it is a choice binding. 7 | 8 | \DESCRIBE 9 | {is_select `@var. t`} returns {true}. If the term is not an epsilon-term the 10 | result is {false}. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \SEEALSO 16 | mk_select, dest_select. 17 | 18 | \ENDDOC 19 | -------------------------------------------------------------------------------- /Cadical/make.ml: -------------------------------------------------------------------------------- 1 | (* ========================================================================= *) 2 | (* Proof-checking interface to the Cadical SAT solver. *) 3 | (* ========================================================================= *) 4 | 5 | loadt "Cadical/dimacs.ml";; 6 | loadt "Cadical/cnf.ml";; 7 | loadt "Cadical/ldrup.ml";; 8 | loadt "Cadical/cadical.ml";; 9 | -------------------------------------------------------------------------------- /Help/PRINT_GOAL_TAC.hlp: -------------------------------------------------------------------------------- 1 | \DOC PRINT_GOAL_TAC 2 | 3 | \TYPE {PRINT_GOAL_TAC : tactic} 4 | 5 | \SYNOPSIS 6 | Prints the goal to standard output. 7 | 8 | \DESCRIBE 9 | Given any goal {A ?- p}, the tactic {PRINT_GOAL_TAC} prints the goal 10 | to standard output. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \SEEALSO 16 | print_goal, print_goalstack, print_term 17 | 18 | \ENDDOC 19 | -------------------------------------------------------------------------------- /Help/dest_neg.hlp: -------------------------------------------------------------------------------- 1 | \DOC dest_neg 2 | 3 | \TYPE {dest_neg : term -> term} 4 | 5 | \SYNOPSIS 6 | Breaks apart a negation, returning its body. 7 | 8 | \DESCRIBE 9 | {dest_neg} is a term destructor for negations: 10 | {dest_neg `~t`} returns {`t`}. 11 | 12 | \FAILURE 13 | Fails with {dest_neg} if term is not a negation. 14 | 15 | \SEEALSO 16 | is_neg, mk_neg. 17 | 18 | \ENDDOC 19 | -------------------------------------------------------------------------------- /Minisat/CREDITS: -------------------------------------------------------------------------------- 1 | (* 2 | * Original HOL4 code for satisfiability testing by Mike Gordon (2004) 3 | * DIMACS cnf parsing code improved by Ken Friis Larsen (2004?) 4 | * MiniSat proof reconstruction and HOL Light port by Hasan Amjad (2005-6) 5 | * Implicit definitional CNF and minor improvements by John Harrison (2006-7) 6 | * zChaff to MiniSat proof translator by Hasan Amjad (2007) 7 | *) 8 | -------------------------------------------------------------------------------- /GL/README.md: -------------------------------------------------------------------------------- 1 | # Modal completeness for the provability logic GL in HOL Light 2 | 3 | (c) Copyright, Marco Maggesi, Cosimo Perini Brogi 2020-2022. 4 | 5 | This repository contains a formal proof of the modal completeness 6 | for the provability logic GL in HOL Light. 7 | 8 | The top-level file is `make.ml`. 9 | 10 | The main theorem is `COMPLETENESS_THEOREM_GEN` in file `completeness.ml`. -------------------------------------------------------------------------------- /Functionspaces/README: -------------------------------------------------------------------------------- 1 | Library of complex function vector spaces 2 | 3 | (c) Copyright, Mohamed Yousri Mahmoud, Vincent Aravantinos, 2012-2013 4 | Hardware Verification Group, 5 | Concordia University 6 | 7 | Contact: , 8 | 9 | Distributed with HOL Light under same license terms 10 | -------------------------------------------------------------------------------- /Help/by.hlp: -------------------------------------------------------------------------------- 1 | \DOC by 2 | 3 | \TYPE {by : tactic -> refinement} 4 | 5 | \SYNOPSIS 6 | Converts a tactic to a refinement. 7 | 8 | \DESCRIBE 9 | The call {by tac} for a tactic {tac} gives a refinement of the current list of 10 | subgoals that applies {tac} to the first subgoal. 11 | 12 | \COMMENTS 13 | Only of interest to users who want to handle `refinements' explicitly. 14 | 15 | \ENDDOC 16 | -------------------------------------------------------------------------------- /Help/body.hlp: -------------------------------------------------------------------------------- 1 | \DOC body 2 | 3 | \TYPE {body : term -> term} 4 | 5 | \SYNOPSIS 6 | Returns the body of an abstraction. 7 | 8 | \DESCRIBE 9 | {body `\var. t`} returns {`t`}. 10 | 11 | \FAILURE 12 | Fails unless the term is an abstraction. 13 | 14 | \EXAMPLE 15 | { 16 | # body `\x. x + 1`;; 17 | val it : term = `x + 1` 18 | } 19 | 20 | \SEEALSO 21 | bndvar, dest_abs. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Help/is_forall.hlp: -------------------------------------------------------------------------------- 1 | \DOC is_forall 2 | 3 | \TYPE {is_forall : term -> bool} 4 | 5 | \SYNOPSIS 6 | Tests a term to see if it is a universal quantification. 7 | 8 | \DESCRIBE 9 | {is_forall `!var. t`} returns {true}. If the term is not a universal 10 | quantification the result is {false}. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \SEEALSO 16 | dest_forall, mk_forall. 17 | 18 | \ENDDOC 19 | -------------------------------------------------------------------------------- /Help/is_cons.hlp: -------------------------------------------------------------------------------- 1 | \DOC is_cons 2 | 3 | \TYPE {is_cons : term -> bool} 4 | 5 | \SYNOPSIS 6 | Tests a term to see if it is an application of {CONS}. 7 | 8 | \DESCRIBE 9 | {is_cons} returns {true} of a term representing a non-empty list. Otherwise it 10 | returns {false}. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \SEEALSO 16 | dest_cons, dest_list, is_list, mk_cons, mk_list. 17 | 18 | \ENDDOC 19 | -------------------------------------------------------------------------------- /Help/is_exists.hlp: -------------------------------------------------------------------------------- 1 | \DOC is_exists 2 | 3 | \TYPE {is_exists : term -> bool} 4 | 5 | \SYNOPSIS 6 | Tests a term to see if it as an existential quantification. 7 | 8 | \DESCRIBE 9 | {is_exists `?var. t`} returns {true}. If the term is not an existential 10 | quantification the result is {false}. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \SEEALSO 16 | dest_exists, mk_exists. 17 | 18 | \ENDDOC 19 | -------------------------------------------------------------------------------- /Help/mem.hlp: -------------------------------------------------------------------------------- 1 | \DOC mem 2 | 3 | \TYPE {mem : 'a -> 'a list -> bool} 4 | 5 | \SYNOPSIS 6 | Tests whether a list contains a certain member. 7 | 8 | \DESCRIBE 9 | {mem x [x1;...;xn]} returns {true} if some {xi} in the list is equal to {x}. 10 | Otherwise it returns {false}. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \SEEALSO 16 | find, tryfind, exists, forall, assoc, rev_assoc. 17 | 18 | \ENDDOC 19 | -------------------------------------------------------------------------------- /Help/rhs.hlp: -------------------------------------------------------------------------------- 1 | \DOC rhs 2 | 3 | \TYPE {rhs : term -> term} 4 | 5 | \SYNOPSIS 6 | Returns the right-hand side of an equation. 7 | 8 | \DESCRIBE 9 | {rhs `t1 = t2`} returns {`t2`}. 10 | 11 | \FAILURE 12 | Fails with {rhs} if term is not an equality. 13 | 14 | \EXAMPLE 15 | { 16 | # rhs `2 + 2 = 4`;; 17 | val it : term = `4` 18 | } 19 | 20 | \SEEALSO 21 | dest_eq, lhs, rand. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Geometric_Algebra/README: -------------------------------------------------------------------------------- 1 | Formulation of geometric algebra G(P,Q,R) with the multivector structure 2 | (P,Q,R)multivector, which can formulate positive definite, negative definite 3 | and zero quadratic forms. 4 | 5 | (c) Copyright, Capital Normal University, China, 2018. 6 | Authors: Liming Li, Zhiping Shi, Yong Guan, Guohui Wang, Sha Ma. 7 | 8 | Distributed with HOL Light under the same license terms 9 | -------------------------------------------------------------------------------- /Help/bndvar.hlp: -------------------------------------------------------------------------------- 1 | \DOC bndvar 2 | 3 | \TYPE {bndvar : term -> term} 4 | 5 | \SYNOPSIS 6 | Returns the bound variable of an abstraction. 7 | 8 | \DESCRIBE 9 | {bndvar `\var. t`} returns {`var`}. 10 | 11 | \FAILURE 12 | Fails unless the term is an abstraction. 13 | 14 | \EXAMPLE 15 | { 16 | # bndvar `\x. x + 1`;; 17 | val it : term = `x` 18 | } 19 | 20 | \SEEALSO 21 | body, dest_abs. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Help/mk_cond.hlp: -------------------------------------------------------------------------------- 1 | \DOC mk_cond 2 | 3 | \TYPE {mk_cond : term * term * term -> term} 4 | 5 | \SYNOPSIS 6 | Constructs a conditional term. 7 | 8 | \DESCRIBE 9 | {mk_cond(`t`,`t1`,`t2`)} returns {`if t then t1 else t2`}. 10 | 11 | \FAILURE 12 | Fails with {mk_cond} if {t} is not of type {`:bool`} or if {t1} and {t2} are 13 | of different types. 14 | 15 | \SEEALSO 16 | dest_cond, is_cond. 17 | 18 | \ENDDOC 19 | -------------------------------------------------------------------------------- /Help/mk_eq.hlp: -------------------------------------------------------------------------------- 1 | \DOC mk_eq 2 | 3 | \TYPE {mk_eq : term * term -> term} 4 | 5 | \SYNOPSIS 6 | Constructs an equation. 7 | 8 | \DESCRIBE 9 | {mk_eq(`t1`,`t2`)} returns {`t1 = t2`}. 10 | 11 | \FAILURE 12 | Fails with {mk_eq} if {t1} and {t2} have different types. 13 | 14 | \EXAMPLE 15 | { 16 | # mk_eq(`1`,`2`);; 17 | val it : term = `1 = 2` 18 | } 19 | 20 | \SEEALSO 21 | dest_eq, is_eq. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Help/constants.hlp: -------------------------------------------------------------------------------- 1 | \DOC constants 2 | 3 | \TYPE {constants : unit -> (string * hol_type) list} 4 | 5 | \SYNOPSIS 6 | Returns a list of the constants currently defined. 7 | 8 | \DESCRIBE 9 | The call 10 | { 11 | constants();; 12 | } 13 | returns a list of all the constants that have been defined so far. 14 | 15 | \FAILURE 16 | Never fails. 17 | 18 | \SEEALSO 19 | axioms, binders, infixes. 20 | 21 | \ENDDOC 22 | -------------------------------------------------------------------------------- /Help/is_numeral.hlp: -------------------------------------------------------------------------------- 1 | \DOC is_numeral 2 | 3 | \TYPE {is_numeral : term -> bool} 4 | 5 | \SYNOPSIS 6 | Tests if a term is a natural number numeral. 7 | 8 | \DESCRIBE 9 | When applied to a term, {is_numeral} returns {true} if and only if the term is 10 | a canonical natural number numeral ({0}, {1}, {2} etc.) 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \SEEALSO 16 | dest_numeral, is_numeral. 17 | 18 | \ENDDOC 19 | -------------------------------------------------------------------------------- /Help/lhs.hlp: -------------------------------------------------------------------------------- 1 | \DOC lhs 2 | 3 | \TYPE {lhs : term -> term} 4 | 5 | \SYNOPSIS 6 | Returns the left-hand side of an equation. 7 | 8 | \DESCRIBE 9 | {lhs `t1 = t2`} returns {`t1`}. 10 | 11 | \FAILURE 12 | Fails with {lhs} if the term is not an equation. 13 | 14 | \EXAMPLE 15 | { 16 | # lhs `2 + 2 = 4`;; 17 | val it : term = `2 + 2` 18 | } 19 | 20 | \SEEALSO 21 | dest_eq, lhand, rand, rhs. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Help/F_F.hlp: -------------------------------------------------------------------------------- 1 | \DOC F_F 2 | 3 | \TYPE {(F_F) : ('a -> 'b) -> ('c -> 'd) -> 'a * 'c -> 'b * 'd} 4 | 5 | \SYNOPSIS 6 | 7 | Infix operator. Applies two functions to a pair: 8 | {(f F_F g) (x,y)} = {(f x, g y)}. 9 | 10 | \KEYWORDS 11 | 12 | \LIBRARY 13 | 14 | \DESCRIBE 15 | 16 | \FAILURE 17 | 18 | Never fails. 19 | 20 | \EXAMPLE 21 | 22 | \USES 23 | 24 | \COMMENTS 25 | 26 | \SEEALSO 27 | f_f_ 28 | 29 | \ENDDOC 30 | -------------------------------------------------------------------------------- /Help/mk_pair.hlp: -------------------------------------------------------------------------------- 1 | \DOC mk_pair 2 | 3 | \TYPE {mk_pair : term * term -> term} 4 | 5 | \SYNOPSIS 6 | Constructs object-level pair from a pair of terms. 7 | 8 | \DESCRIBE 9 | {mk_pair(`t1`,`t2`)} returns {`(t1,t2)`}. 10 | 11 | \FAILURE 12 | Never fails. 13 | 14 | \EXAMPLE 15 | { 16 | # mk_pair(`x:real`,`T`);; 17 | val it : term = `x,T` 18 | } 19 | 20 | \SEEALSO 21 | dest_pair, is_pair, mk_cons. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Help/DISJ1_TAC.hlp: -------------------------------------------------------------------------------- 1 | \DOC DISJ1_TAC 2 | 3 | \TYPE {DISJ1_TAC : tactic} 4 | 5 | \SYNOPSIS 6 | Selects the left disjunct of a disjunctive goal. 7 | 8 | \KEYWORDS 9 | tactic, disjunction. 10 | 11 | \DESCRIBE 12 | { 13 | A ?- t1 \/ t2 14 | =============== DISJ1_TAC 15 | A ?- t1 16 | } 17 | 18 | \FAILURE 19 | Fails if the goal is not a disjunction. 20 | 21 | \SEEALSO 22 | DISJ1, DISJ2, DISJ2_TAC. 23 | 24 | \ENDDOC 25 | -------------------------------------------------------------------------------- /Help/uncurry.hlp: -------------------------------------------------------------------------------- 1 | \DOC uncurry 2 | 3 | \TYPE {uncurry : ('a -> 'b -> 'c) -> 'a * 'b -> 'c} 4 | 5 | \SYNOPSIS 6 | Converts a function taking two arguments into a function taking a single 7 | paired argument. 8 | 9 | \DESCRIBE 10 | The application {uncurry f} returns {fun (x,y) -> f x y}, so that 11 | { 12 | uncurry f (x,y) = f x y 13 | } 14 | \FAILURE 15 | Never fails. 16 | 17 | \SEEALSO 18 | curry. 19 | 20 | \ENDDOC 21 | -------------------------------------------------------------------------------- /Divstep/Makefile: -------------------------------------------------------------------------------- 1 | # The hull_light.ml file can be freshly generated from the Sage script, 2 | # using the first (default) Makefile target, while "clobber" removes 3 | # it so that it needs to be regenerated. 4 | 5 | .PHONY: clean clobber 6 | 7 | hull_light.ml: hull-light-20230416.sage; sage hull-light-20230416.sage >hull_light.ml 8 | 9 | clean:; rm -f *.sage.py 10 | 11 | clobber:; rm -f hull-light-20230416.sage.py hull_light.ml 12 | -------------------------------------------------------------------------------- /Help/DISJ2_TAC.hlp: -------------------------------------------------------------------------------- 1 | \DOC DISJ2_TAC 2 | 3 | \TYPE {DISJ2_TAC : tactic} 4 | 5 | \SYNOPSIS 6 | Selects the right disjunct of a disjunctive goal. 7 | 8 | \KEYWORDS 9 | tactic, disjunction. 10 | 11 | \DESCRIBE 12 | { 13 | A ?- t1 \/ t2 14 | =============== DISJ2_TAC 15 | A ?- t2 16 | } 17 | 18 | \FAILURE 19 | Fails if the goal is not a disjunction. 20 | 21 | \SEEALSO 22 | DISJ1, DISJ1_TAC, DISJ2. 23 | 24 | \ENDDOC 25 | -------------------------------------------------------------------------------- /Help/use_file.hlp: -------------------------------------------------------------------------------- 1 | \DOC use_file 2 | 3 | \TYPE {use_file : string -> unit} 4 | 5 | \SYNOPSIS 6 | Load a file, much like OCaml's {#use} directive. 7 | 8 | \DESCRIBE 9 | Essentially the same as OCaml's {#use} directive, but a regular OCaml function 10 | and therefore easier to exploit programmatically. 11 | 12 | \FAILURE 13 | Only fails if the included file causes failure. 14 | 15 | \SEEALSO 16 | loads, loadt. 17 | 18 | \ENDDOC 19 | -------------------------------------------------------------------------------- /QBF/make.ml: -------------------------------------------------------------------------------- 1 | (* ========================================================================= *) 2 | (* Ondrej Kuncar's HOL Light QBF code. *) 3 | (* ========================================================================= *) 4 | 5 | #use "topfind";; 6 | #require "ocamlgraph";; 7 | 8 | loads "Minisat/make.ml";; 9 | loads "QBF/mygraph.ml";; 10 | loads "QBF/qbfr.ml";; 11 | loads "QBF/qbf.ml";; 12 | -------------------------------------------------------------------------------- /Help/dest_eq.hlp: -------------------------------------------------------------------------------- 1 | \DOC dest_eq 2 | 3 | \TYPE {dest_eq : term -> term * term} 4 | 5 | \SYNOPSIS 6 | Term destructor for equality. 7 | 8 | \DESCRIBE 9 | {dest_eq(`t1 = t2`)} returns {(`t1`,`t2`)}. 10 | 11 | \FAILURE 12 | Fails with {dest_eq} if term is not an equality. 13 | 14 | \EXAMPLE 15 | { 16 | # dest_eq `2 + 2 = 4`;; 17 | val it : term * term = (`2 + 2`, `4`) 18 | } 19 | 20 | \SEEALSO 21 | is_eq, mk_eq. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Help/is_uexists.hlp: -------------------------------------------------------------------------------- 1 | \DOC is_uexists 2 | 3 | \TYPE {is_uexists : term -> bool} 4 | 5 | \SYNOPSIS 6 | Tests if a term is of the form `there exists a unique ...' 7 | 8 | \DESCRIBE 9 | If {t} has the form {?!x. p[x]} (there exists a unique {x} such that {p[x]} 10 | then {is_uexists t} returns {true}, otherwise {false}. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \SEEALSO 16 | dest_uexists, is_exists, is_forall. 17 | 18 | \ENDDOC 19 | -------------------------------------------------------------------------------- /Help/isnum.hlp: -------------------------------------------------------------------------------- 1 | \DOC isnum 2 | 3 | \TYPE {isnum : string -> bool} 4 | 5 | \SYNOPSIS 6 | Tests if a one-character string is a decimal digit. 7 | 8 | \DESCRIBE 9 | The call {isnum s} tests whether the first character of string {s} (normally 10 | it is the only character) is a decimal digit. 11 | 12 | \FAILURE 13 | Fails if the string is empty. 14 | 15 | \SEEALSO 16 | isalnum, isalpha, isbra, issep, isspace, issymb. 17 | 18 | \ENDDOC 19 | -------------------------------------------------------------------------------- /Help/hol_version.hlp: -------------------------------------------------------------------------------- 1 | \DOC hol_version 2 | 3 | \TYPE {hol_version : string} 4 | 5 | \SYNOPSIS 6 | A string indicating the version of HOL Light. 7 | 8 | \DESCRIBE 9 | This string is a numeric version number for HOL Light. 10 | 11 | \FAILURE 12 | Not applicable. 13 | 14 | \EXAMPLE 15 | On my laptop, the value is: 16 | { 17 | # hol_version;; 18 | val it : string = "2.10" 19 | } 20 | 21 | \SEEALSO 22 | startup_banner. 23 | 24 | \ENDDOC 25 | -------------------------------------------------------------------------------- /Help/print_thm.hlp: -------------------------------------------------------------------------------- 1 | \DOC print_thm 2 | 3 | \TYPE {print_thm : thm -> unit} 4 | 5 | \SYNOPSIS 6 | Prints a HOL theorem to the standard output. 7 | 8 | \DESCRIBE 9 | The call {print_thm th} prints the usual textual representation of the 10 | theorem {th} to the standard output. 11 | 12 | \COMMENTS 13 | This is invoked automatically at the toplevel when theorems are printed. 14 | 15 | \SEEALSO 16 | print_type, print_term. 17 | 18 | \ENDDOC 19 | -------------------------------------------------------------------------------- /Help/dest_disj.hlp: -------------------------------------------------------------------------------- 1 | \DOC dest_disj 2 | 3 | \TYPE {dest_disj : term -> term * term} 4 | 5 | \SYNOPSIS 6 | Breaks apart a disjunction into the two disjuncts. 7 | 8 | \DESCRIBE 9 | {dest_disj} is a term destructor for disjunctions: 10 | { 11 | dest_disj `t1 \/ t2` 12 | } 13 | \noindent returns {(`t1`,`t2`)}. 14 | 15 | \FAILURE 16 | Fails with {dest_disj} if term is not a disjunction. 17 | 18 | \SEEALSO 19 | is_disj, mk_disj. 20 | 21 | \ENDDOC 22 | -------------------------------------------------------------------------------- /Help/map.hlp: -------------------------------------------------------------------------------- 1 | \DOC map 2 | 3 | \TYPE {map : ('a -> 'b) -> 'a list -> 'b list} 4 | 5 | \SYNOPSIS 6 | Applies a function to every element of a list. 7 | 8 | \DESCRIBE 9 | {map f [x1;...;xn]} returns {[(f x1);...;(f xn)]}. 10 | 11 | \FAILURE 12 | Never fails. 13 | 14 | \EXAMPLE 15 | { 16 | # map (fun x -> x * 2) [];; 17 | val it : int list = [] 18 | # map (fun x -> x * 2) [1;2;3];; 19 | val it : int list = [2; 4; 6] 20 | } 21 | 22 | \ENDDOC 23 | -------------------------------------------------------------------------------- /Help/metisverb.hlp: -------------------------------------------------------------------------------- 1 | \DOC metisverb 2 | 3 | \TYPE {metisverb : bool ref} 4 | 5 | \SYNOPSIS 6 | Make {METIS}'s output more verbose and detailed. 7 | 8 | \DESCRIBE 9 | When this reference variable is set to {true}, it makes any applications of 10 | {METIS}, {METIS_TAC} and related rules and tactics provide more verbose output 11 | about their working. 12 | 13 | \FAILURE 14 | Not applicable. 15 | 16 | \SEEALSO 17 | meson_chatty.ml 18 | 19 | \ENDDOC 20 | -------------------------------------------------------------------------------- /Help/dest_list.hlp: -------------------------------------------------------------------------------- 1 | \DOC dest_list 2 | 3 | \TYPE {dest_list : term -> term list} 4 | 5 | \SYNOPSIS 6 | Iteratively breaks apart a list term. 7 | 8 | \DESCRIBE 9 | {dest_list} is a term destructor for lists: 10 | {dest_list(`[t1;...;tn]:(ty)list`)} returns {[`t1`;...;`tn`]}. 11 | 12 | \FAILURE 13 | Fails with {dest_list} if the term is not a list. 14 | 15 | \SEEALSO 16 | dest_cons, dest_setenum, is_cons, is_list, mk_cons, mk_list. 17 | 18 | \ENDDOC 19 | -------------------------------------------------------------------------------- /Help/issep.hlp: -------------------------------------------------------------------------------- 1 | \DOC issep 2 | 3 | \TYPE {issep : string -> bool} 4 | 5 | \SYNOPSIS 6 | Tests if a one-character string is a separator. 7 | 8 | \DESCRIBE 9 | The call {issep s} tests whether the first character of string {s} (normally 10 | it is the only character) is one of the separators `{,}' or `{;}'. 11 | 12 | \FAILURE 13 | Fails if the string is empty. 14 | 15 | \SEEALSO 16 | isalnum, isalpha, isbra, isnum, isspace, issymb. 17 | 18 | \ENDDOC 19 | -------------------------------------------------------------------------------- /Help/rotate.hlp: -------------------------------------------------------------------------------- 1 | \DOC rotate 2 | 3 | \TYPE {rotate : int -> refinement} 4 | 5 | \SYNOPSIS 6 | Rotate a goalstate. 7 | 8 | \DESCRIBE 9 | The function {rotate n gl} rotates a list {gl} of subgoals by {n} places. The 10 | function {r} is the special case where this modification is applied to the 11 | imperative variable of unproven subgoals. 12 | 13 | \FAILURE 14 | Fails only if the list of goals is empty. 15 | 16 | \SEEALSO 17 | r. 18 | 19 | \ENDDOC 20 | -------------------------------------------------------------------------------- /Help/type_abbrevs.hlp: -------------------------------------------------------------------------------- 1 | \DOC type_abbrevs 2 | 3 | \TYPE {type_abbrevs : unit -> (string * hol_type) list} 4 | 5 | \SYNOPSIS 6 | Lists all current type abbreviations. 7 | 8 | \DESCRIBE 9 | The call {type_abbrevs()} returns a list of all current type abbreviations, 10 | which are applied when parsing types but have no logical significance. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \SEEALSO 16 | new_type_abbrev, remove_type_abbrev. 17 | 18 | \ENDDOC 19 | -------------------------------------------------------------------------------- /Help/bool_ty.hlp: -------------------------------------------------------------------------------- 1 | \DOC bool_ty 2 | 3 | \TYPE {bool_ty : hol_type} 4 | 5 | \SYNOPSIS 6 | The type {`:bool`}. 7 | 8 | \DESCRIBE 9 | This name is bound to the HOL type {:bool}. 10 | 11 | \FAILURE 12 | Not applicable. 13 | 14 | \USES 15 | Exploiting the very common type {:bool} inside derived rules without the 16 | inefficiency or inconvenience of calling a quotation parser or explicit 17 | constructor. 18 | 19 | \SEEALSO 20 | aty, bty. 21 | 22 | \ENDDOC 23 | -------------------------------------------------------------------------------- /Help/mk_neg.hlp: -------------------------------------------------------------------------------- 1 | \DOC mk_neg 2 | 3 | \TYPE {mk_neg : term -> term} 4 | 5 | \SYNOPSIS 6 | Constructs a logical negation. 7 | 8 | \DESCRIBE 9 | {mk_neg `t`} returns {`~t`}. 10 | 11 | \FAILURE 12 | Fails with {mk_neg} unless {t} is of type {bool}. 13 | 14 | \EXAMPLE 15 | { 16 | # mk_neg `p /\ q`;; 17 | val it : term = `~(p /\ q)` 18 | 19 | # mk_neg `~p`;; 20 | val it : term = `~ ~p` 21 | } 22 | 23 | \SEEALSO 24 | dest_neg, is_neg. 25 | 26 | \ENDDOC 27 | -------------------------------------------------------------------------------- /Help/dest_thm.hlp: -------------------------------------------------------------------------------- 1 | \DOC dest_thm 2 | 3 | \TYPE {dest_thm : thm -> term list * term} 4 | 5 | \SYNOPSIS 6 | Breaks a theorem into assumption list and conclusion. 7 | 8 | \DESCRIBE 9 | {dest_thm (t1,...,tn |- t)} returns {([`t1`;...;`tn`],`t`)}. 10 | 11 | \FAILURE 12 | Never fails. 13 | 14 | \EXAMPLE 15 | { 16 | # dest_thm (ASSUME `1 = 0`);; 17 | val it : term list * term = ([`1 = 0`], `1 = 0`) 18 | } 19 | 20 | \SEEALSO 21 | concl, hyp. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Help/flat.hlp: -------------------------------------------------------------------------------- 1 | \DOC flat 2 | 3 | \TYPE {flat : 'a list list -> 'a list} 4 | 5 | \SYNOPSIS 6 | Flattens a list of lists into one long list. 7 | 8 | \KEYWORDS 9 | list. 10 | 11 | \DESCRIBE 12 | {flat [l1;...;ln]} returns {(l1 @ ... @ ln)} where each li is a list and {@} 13 | is list concatenation. 14 | 15 | \FAILURE 16 | Never fails. 17 | 18 | \EXAMPLE 19 | { 20 | # flat [[1;2];[3;4;5];[6]];; 21 | val it : int list = [1; 2; 3; 4; 5; 6] 22 | } 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Help/mk_imp.hlp: -------------------------------------------------------------------------------- 1 | \DOC mk_imp 2 | 3 | \TYPE {mk_imp : term * term -> term} 4 | 5 | \SYNOPSIS 6 | Constructs an implication. 7 | 8 | \DESCRIBE 9 | {mk_imp(`t1`,`t2`)} returns {`t1 ==> t2`}. 10 | 11 | \FAILURE 12 | Fails with {mk_imp} if either {t1} or {t2} are not of type {`:bool`}. 13 | 14 | \EXAMPLE 15 | { 16 | # mk_imp(`p /\ q`,`r:bool`);; 17 | val it : term = `p /\ q ==> r` 18 | } 19 | 20 | \SEEALSO 21 | dest_imp, is_imp, list_mk_imp. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Help/report.hlp: -------------------------------------------------------------------------------- 1 | \DOC report 2 | 3 | \TYPE {report : string -> unit} 4 | 5 | \SYNOPSIS 6 | Prints a string and a following line break. 7 | 8 | \DESCRIBE 9 | The call {report s} prints the string {s} to the terminal and then a following 10 | newline. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \EXAMPLE 16 | { 17 | # report "Proof completed OK";; 18 | Proof completed OK 19 | val it : unit = () 20 | } 21 | 22 | \SEEALSO 23 | remark, warn. 24 | 25 | \ENDDOC 26 | -------------------------------------------------------------------------------- /Help/copverb.hlp: -------------------------------------------------------------------------------- 1 | \DOC metisverb 2 | 3 | \TYPE {metisverb : bool ref} 4 | 5 | \SYNOPSIS 6 | Make {METIS}'s output more verbose and detailed. 7 | 8 | \DESCRIBE 9 | When this reference variable is set to {true}, it makes any applications of 10 | {METIS}, {METIS_TAC} and related rules and tactics provide more verbose output 11 | about their working. 12 | 13 | \FAILURE 14 | Not applicable. 15 | 16 | \SEEALSO 17 | copverb, meson_chatty, METIS, METIS_TAC. 18 | 19 | \ENDDOC 20 | -------------------------------------------------------------------------------- /Help/dest_select.hlp: -------------------------------------------------------------------------------- 1 | \DOC dest_select 2 | 3 | \TYPE {dest_select : term -> term * term} 4 | 5 | \SYNOPSIS 6 | Breaks apart a choice term into selected variable and body. 7 | 8 | \DESCRIBE 9 | {dest_select} is a term destructor for choice terms: 10 | { 11 | dest_select `@var. t` 12 | } 13 | \noindent returns {(`var`,`t`)}. 14 | 15 | \FAILURE 16 | Fails with {dest_select} if term is not an epsilon-term. 17 | 18 | \SEEALSO 19 | mk_select, is_select. 20 | 21 | \ENDDOC 22 | -------------------------------------------------------------------------------- /Help/mk_disj.hlp: -------------------------------------------------------------------------------- 1 | \DOC mk_disj 2 | 3 | \TYPE {mk_disj : term * term -> term} 4 | 5 | \SYNOPSIS 6 | Constructs a disjunction. 7 | 8 | \DESCRIBE 9 | {mk_disj(`t1`,`t2`)} returns {`t1 \/ t2`}. 10 | 11 | \FAILURE 12 | Fails with {mk_disj} if either {t1} or {t2} are not of type {`:bool`}. 13 | 14 | \EXAMPLE 15 | { 16 | # mk_disj(`x = 1`,`y <= 2`);; 17 | val it : term = `x = 1 \/ y <= 2` 18 | } 19 | 20 | \SEEALSO 21 | dest_disj, is_disj, list_mk_disj. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Help/rand.hlp: -------------------------------------------------------------------------------- 1 | \DOC rand 2 | 3 | \TYPE {rand : term -> term} 4 | 5 | \SYNOPSIS 6 | Returns the operand from a combination (function application). 7 | 8 | \DESCRIBE 9 | {rand `t1 t2`} returns {`t2`}. 10 | 11 | \FAILURE 12 | Fails with {rand} if term is not a combination. 13 | 14 | \EXAMPLE 15 | { 16 | # rand `SUC 0`;; 17 | val it : term = `0` 18 | # rand `x + y`;; 19 | val it : term = `y` 20 | } 21 | 22 | \SEEALSO 23 | rator, lhand, dest_comb. 24 | 25 | \ENDDOC 26 | -------------------------------------------------------------------------------- /Help/dest_cond.hlp: -------------------------------------------------------------------------------- 1 | \DOC dest_cond 2 | 3 | \TYPE {dest_cond : term -> term * (term * term)} 4 | 5 | \SYNOPSIS 6 | Breaks apart a conditional into the three terms involved. 7 | 8 | \DESCRIBE 9 | {dest_cond} is a term destructor for conditionals: 10 | { 11 | dest_cond `if t then t1 else t2` 12 | } 13 | \noindent returns {(`t`,`t1`,`t2`)}. 14 | 15 | \FAILURE 16 | Fails with {dest_cond} if term is not a conditional. 17 | 18 | \SEEALSO 19 | mk_cond, is_cond. 20 | 21 | \ENDDOC 22 | -------------------------------------------------------------------------------- /Help/dest_imp.hlp: -------------------------------------------------------------------------------- 1 | \DOC dest_imp 2 | 3 | \TYPE {dest_imp : term -> term * term} 4 | 5 | \SYNOPSIS 6 | Breaks apart an implication into antecedent and consequent. 7 | 8 | \DESCRIBE 9 | {dest_imp} is a term destructor for implications. Thus 10 | { 11 | dest_imp `t1 ==> t2` 12 | } 13 | \noindent returns 14 | { 15 | (`t1`,`t2`) 16 | } 17 | \FAILURE 18 | Fails with {dest_imp} if term is not an implication. 19 | 20 | \SEEALSO 21 | is_imp, mk_imp, strip_imp. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Help/dest_uexists.hlp: -------------------------------------------------------------------------------- 1 | \DOC dest_uexists 2 | 3 | \TYPE {dest_uexists : term -> term * term} 4 | 5 | \SYNOPSIS 6 | Breaks apart a unique existence term. 7 | 8 | \DESCRIBE 9 | If {t} has the form {?!x. p[x]} (there exists a unique [x} such that {p[x]} 10 | then {dest_uexists t} returns the pair {x,p[x]}; otherwise it fails. 11 | 12 | \FAILURE 13 | Fails if the term is not a `unique existence' term. 14 | 15 | \SEEALSO 16 | dest_exists, dest_forall, is_uexists. 17 | 18 | \ENDDOC 19 | -------------------------------------------------------------------------------- /Help/isspace.hlp: -------------------------------------------------------------------------------- 1 | \DOC isspace 2 | 3 | \TYPE {isspace : string -> bool} 4 | 5 | \SYNOPSIS 6 | Tests if a one-character string is some kind of space. 7 | 8 | \DESCRIBE 9 | The call {isspace s} tests whether the first character of string {s} (normally 10 | it is the only character) is a `space' of some kind, including tab and newline. 11 | 12 | \FAILURE 13 | Fails if the string is empty. 14 | 15 | \SEEALSO 16 | isalnum, isalpha, isbra, isnum, issep, issymb. 17 | 18 | \ENDDOC 19 | -------------------------------------------------------------------------------- /Help/num_0.hlp: -------------------------------------------------------------------------------- 1 | \DOC num_0 2 | 3 | \TYPE {num_0 : num} 4 | 5 | \SYNOPSIS 6 | Constant zero in unlimited-size integers. 7 | 8 | \DESCRIBE 9 | The constant {num_0} is bound to the integer constant 0 in the 10 | unlimited-precision numbers provided by the OCaml {Num} library. 11 | 12 | \FAILURE 13 | Not applicable. 14 | 15 | \USES 16 | Exactly the same as {Int 0}, but may save recreation of a cons cell each time. 17 | 18 | \SEEALSO 19 | num_1, num_2, num_10. 20 | 21 | \ENDDOC 22 | -------------------------------------------------------------------------------- /Help/num_1.hlp: -------------------------------------------------------------------------------- 1 | \DOC num_1 2 | 3 | \TYPE {num_1 : num} 4 | 5 | \SYNOPSIS 6 | Constant one in unlimited-size integers. 7 | 8 | \DESCRIBE 9 | The constant {num_1} is bound to the integer constant 1 in the 10 | unlimited-precision numbers provided by the OCaml {Num} library. 11 | 12 | \FAILURE 13 | Not applicable. 14 | 15 | \USES 16 | Exactly the same as {Int 1}, but may save recreation of a cons cell each time. 17 | 18 | \SEEALSO 19 | num_0, num_2, num_10. 20 | 21 | \ENDDOC 22 | -------------------------------------------------------------------------------- /Help/num_2.hlp: -------------------------------------------------------------------------------- 1 | \DOC num_2 2 | 3 | \TYPE {num_2 : num} 4 | 5 | \SYNOPSIS 6 | Constant two in unlimited-size integers. 7 | 8 | \DESCRIBE 9 | The constant {num_2} is bound to the integer constant 2 in the 10 | unlimited-precision numbers provided by the OCaml {Num} library. 11 | 12 | \FAILURE 13 | Not applicable. 14 | 15 | \USES 16 | Exactly the same as {Int 2}, but may save recreation of a cons cell each time. 17 | 18 | \SEEALSO 19 | num_0, num_1, num_10. 20 | 21 | \ENDDOC 22 | -------------------------------------------------------------------------------- /IEEE/make.ml: -------------------------------------------------------------------------------- 1 | (* -------------------------------------------------------------------------- *) 2 | (* Load this file to get all theorems and formalizations. *) 3 | (* -------------------------------------------------------------------------- *) 4 | 5 | loadt "IEEE/common.hl";; 6 | loadt "IEEE/fixed.hl";; 7 | loadt "IEEE/fixed_thms.hl";; 8 | loadt "IEEE/float.hl";; 9 | loadt "IEEE/float_thms.hl";; 10 | loadt "IEEE/ieee.hl";; 11 | loadt "IEEE/ieee_thms.hl";; 12 | -------------------------------------------------------------------------------- /Proofrecording/tools/src/Makefile: -------------------------------------------------------------------------------- 1 | install: nametheorems.jar detecteq.jar; mv nametheorems.jar detecteq.jar .. 2 | 3 | nametheorems.jar: NamedTheorem.java; \ 4 | javac NamedTheorem.java; \ 5 | jar cvfe nametheorems.jar NamedTheorem NamedTheorem.class; \ 6 | rm -f NamedTheorem.class 7 | 8 | detecteq.jar: DetectEq.java; \ 9 | javac DetectEq.java; \ 10 | jar cvfe detecteq.jar DetectEq DetectEq.class; \ 11 | rm -f DetectEq*.class 12 | 13 | clean:; rm -f *.class *.jar 14 | -------------------------------------------------------------------------------- /Help/num_10.hlp: -------------------------------------------------------------------------------- 1 | \DOC num_10 2 | 3 | \TYPE {num_10 : num} 4 | 5 | \SYNOPSIS 6 | Constant ten in unlimited-size integers. 7 | 8 | \DESCRIBE 9 | The constant {num_10} is bound to the integer constant 10 in the 10 | unlimited-precision numbers provided by the OCaml {Num} library. 11 | 12 | \FAILURE 13 | Not applicable. 14 | 15 | \USES 16 | Exactly the same as {Int 10}, but may save recreation of a cons cell each time. 17 | 18 | \SEEALSO 19 | num_0, num_1, num_2. 20 | 21 | \ENDDOC 22 | -------------------------------------------------------------------------------- /Help/EQT_INTRO.hlp: -------------------------------------------------------------------------------- 1 | \DOC EQT_INTRO 2 | 3 | \TYPE {EQT_INTRO : thm -> thm} 4 | 5 | \SYNOPSIS 6 | Introduces equality with {T}. 7 | 8 | \KEYWORDS 9 | rule, truth. 10 | 11 | \DESCRIBE 12 | { 13 | A |- tm 14 | --------------- EQF_INTRO 15 | A |- tm <=> T 16 | } 17 | 18 | \FAILURE 19 | Never fails. 20 | 21 | \EXAMPLE 22 | { 23 | # EQT_INTRO (REFL `2`);; 24 | val it : thm = |- 2 = 2 <=> T 25 | } 26 | 27 | \SEEALSO 28 | EQF_ELIM, EQF_INTRO, EQT_ELIM. 29 | 30 | \ENDDOC 31 | -------------------------------------------------------------------------------- /Help/mk_conj.hlp: -------------------------------------------------------------------------------- 1 | \DOC mk_conj 2 | 3 | \TYPE {mk_conj : term * term -> term} 4 | 5 | \SYNOPSIS 6 | Constructs a conjunction. 7 | 8 | \DESCRIBE 9 | {mk_conj(`t1`,`t2`)} returns {`t1 /\ t2`}. 10 | 11 | \FAILURE 12 | Fails with {mk_conj} if either {t1} or {t2} are not of type {`:bool`}. 13 | 14 | \EXAMPLE 15 | { 16 | # mk_conj(`1 + 1 = 2`,`2 + 2 = 4`);; 17 | val it : term = `1 + 1 = 2 /\ 2 + 2 = 4` 18 | } 19 | 20 | \SEEALSO 21 | dest_conj, is_conj, list_mk_conj. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Help/tyvars.hlp: -------------------------------------------------------------------------------- 1 | \DOC tyvars 2 | 3 | \TYPE {tyvars : hol_type -> hol_type list} 4 | 5 | \SYNOPSIS 6 | Returns a list of the type variables in a type. 7 | 8 | \DESCRIBE 9 | When applied to a type, {tyvars} returns a list (possibly empty) of the type 10 | variables that it involves. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \EXAMPLE 16 | { 17 | # tyvars `:(A->bool)->A`;; 18 | val it : hol_type list = [`:A`] 19 | } 20 | 21 | \SEEALSO 22 | type_vars_in_term. 23 | 24 | \ENDDOC 25 | -------------------------------------------------------------------------------- /QBF/mygraph.ml: -------------------------------------------------------------------------------- 1 | unset_jrh_lexer;; 2 | 3 | module Intvertex = struct 4 | type t = int 5 | let compare : t -> t -> int = compare 6 | let hash = Hashtbl.hash 7 | let equal = (=) 8 | let default = 0 9 | end;; 10 | 11 | module Gr = Graph.Imperative.Digraph.ConcreteBidirectional(Intvertex);; 12 | 13 | module Topo = Graph.Topological.Make(Gr);; 14 | 15 | let make_vertex var_index = Gr.V.create var_index;; 16 | let dest_vertex var_index = Gr.V.label var_index;; 17 | 18 | set_jrh_lexer;; 19 | -------------------------------------------------------------------------------- /Help/dest_forall.hlp: -------------------------------------------------------------------------------- 1 | \DOC dest_forall 2 | 3 | \TYPE {dest_forall : term -> term * term} 4 | 5 | \SYNOPSIS 6 | Breaks apart a universally quantified term into quantified variable and body. 7 | 8 | \DESCRIBE 9 | {dest_forall} is a term destructor for universal quantification: 10 | {dest_forall `!var. t`} returns {(`var`,`t`)}. 11 | 12 | \FAILURE 13 | Fails with {dest_forall} if term is not a universal quantification. 14 | 15 | \SEEALSO 16 | is_forall, mk_forall, strip_forall. 17 | 18 | \ENDDOC 19 | -------------------------------------------------------------------------------- /Help/mk_goalstate.hlp: -------------------------------------------------------------------------------- 1 | \DOC mk_goalstate 2 | 3 | \TYPE {mk_goalstate : goal -> goalstate} 4 | 5 | \SYNOPSIS 6 | Converts a goal into a 1-element goalstate. 7 | 8 | \DESCRIBE 9 | Given a goal {g}, the call {mk_goalstate g} converts it into a goalstate with 10 | that goal as its only member. (A goalstate consists of a list of subgoals as 11 | well as justification and metavariable information.) 12 | 13 | \FAILURE 14 | Never fails. 15 | 16 | \SEEALSO 17 | g, set_goal, TAC_PROOF. 18 | 19 | \ENDDOC 20 | -------------------------------------------------------------------------------- /Help/print_num.hlp: -------------------------------------------------------------------------------- 1 | \DOC print_num 2 | 3 | \TYPE {print_num : num -> unit} 4 | 5 | \SYNOPSIS 6 | Print an arbitrary-precision number to the terminal. 7 | 8 | \DESCRIBE 9 | This function prints an arbitrary-precision (type {num}) number to the 10 | terminal. It is automatically invoked on anything of type {num} at the toplevel 11 | anyway, but it may sometimes be useful to issue it under user control. 12 | 13 | \FAILURE 14 | Never fails. 15 | 16 | \SEEALSO 17 | pp_print_num. 18 | 19 | \ENDDOC 20 | -------------------------------------------------------------------------------- /miz3/test.ml: -------------------------------------------------------------------------------- 1 | loadt "miz3/Samples/samples.ml";; 2 | loadt "miz3/Samples/sample.ml";; 3 | loadt "miz3/Samples/talk.ml";; 4 | loadt "miz3/Samples/drinker.ml";; 5 | loadt "miz3/Samples/irrat2.ml";; 6 | loadt "miz3/Samples/lagrange.ml";; 7 | loadt "miz3/Samples/lagrange1.ml";; 8 | loadt "miz3/Samples/icms.ml";; 9 | loadt "miz3/Samples/other_mizs.ml";; 10 | loadt "miz3/Samples/robbins.ml";; 11 | loadt "miz3/Samples/forster.ml";; 12 | loadt "miz3/Samples/luxury.ml";; 13 | loadt "miz3/Samples/tobias.ml";; 14 | -------------------------------------------------------------------------------- /Help/dest_var.hlp: -------------------------------------------------------------------------------- 1 | \DOC dest_var 2 | 3 | \TYPE {dest_var : term -> string * hol_type} 4 | 5 | \SYNOPSIS 6 | Breaks apart a variable into name and type. 7 | 8 | \DESCRIBE 9 | {dest_var `var:ty`} returns {("var",`:ty`)}. 10 | 11 | \FAILURE 12 | Fails with {dest_var} if term is not a variable. 13 | 14 | \EXAMPLE 15 | { 16 | # dest_var `x:num`;; 17 | val it : string * hol_type = ("x", `:num`) 18 | } 19 | 20 | \SEEALSO 21 | mk_var, is_var, dest_const, dest_comb, dest_abs, name_of. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Help/is_pair.hlp: -------------------------------------------------------------------------------- 1 | \DOC is_pair 2 | 3 | \TYPE {is_pair : term -> bool} 4 | 5 | \SYNOPSIS 6 | Tests a term to see if it is a pair. 7 | 8 | \DESCRIBE 9 | {is_pair `(t1,t2)`} returns {true}. If the term is not a pair the result is 10 | {false}. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \EXAMPLE 16 | { 17 | # is_pair `1,2,3`;; 18 | val it : bool = true 19 | 20 | # is_pair `[1;2;3]`;; 21 | val it : bool = false 22 | } 23 | 24 | \SEEALSO 25 | dest_pair, is_cons, mk_pair. 26 | 27 | \ENDDOC 28 | -------------------------------------------------------------------------------- /Help/el.hlp: -------------------------------------------------------------------------------- 1 | \DOC el 2 | 3 | \TYPE {el : int -> 'a list -> 'a} 4 | 5 | \SYNOPSIS 6 | Extracts a specified element from a list. 7 | 8 | \DESCRIBE 9 | {el i [x0;x1;...;xn]} returns {xi}. Note that the elements are numbered 10 | starting from {0}, not {1}. 11 | 12 | \FAILURE 13 | Fails with {el} if the integer argument is negative or greater than the 14 | length of the list. 15 | 16 | \EXAMPLE 17 | { 18 | # el 3 [1;2;7;1];; 19 | val it : int = 1 20 | } 21 | 22 | \SEEALSO 23 | hd, tl. 24 | 25 | \ENDDOC 26 | -------------------------------------------------------------------------------- /Help/is_var.hlp: -------------------------------------------------------------------------------- 1 | \DOC is_var 2 | 3 | \TYPE {is_var : term -> bool} 4 | 5 | \SYNOPSIS 6 | Tests a term to see if it is a variable. 7 | 8 | \DESCRIBE 9 | {is_var `var:ty`} returns {true}. If the term is not a variable the result 10 | is {false}. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \EXAMPLE 16 | { 17 | # is_var `x:bool`;; 18 | val it : bool = true 19 | # is_var `T`;; 20 | val it : bool = false 21 | } 22 | 23 | \SEEALSO 24 | mk_var, dest_var, is_const, is_comb, is_abs. 25 | 26 | \ENDDOC 27 | -------------------------------------------------------------------------------- /Help/ALL_CONV.hlp: -------------------------------------------------------------------------------- 1 | \DOC ALL_CONV 2 | 3 | \TYPE {ALL_CONV : conv} 4 | 5 | \SYNOPSIS 6 | Conversion that always succeeds and leaves a term unchanged. 7 | 8 | \KEYWORDS 9 | conversion, identity. 10 | 11 | \DESCRIBE 12 | When applied to a term {`t`}, the conversion {ALL_CONV} returns the 13 | theorem {|- t = t}. It is just {REFL} explicitly regarded as a conversion. 14 | 15 | \FAILURE 16 | Never fails. 17 | 18 | \USES 19 | Identity element for {THENC}. 20 | 21 | \SEEALSO 22 | NO_CONV, REFL. 23 | 24 | \ENDDOC 25 | -------------------------------------------------------------------------------- /Help/aty.hlp: -------------------------------------------------------------------------------- 1 | \DOC aty 2 | 3 | \TYPE {aty : hol_type} 4 | 5 | \SYNOPSIS 6 | The type variable {`:A`}. 7 | 8 | \DESCRIBE 9 | This name is bound to the HOL type {:A}. 10 | 11 | \FAILURE 12 | Not applicable. 13 | 14 | \USES 15 | Exploiting the very common type variable {:A} inside derived rules (e.g. an 16 | instantiation list for {inst} or {type_subst}) without the inefficiency or 17 | inconvenience of calling a quotation parser or explicit constructor. 18 | 19 | \SEEALSO 20 | bty, bool_ty. 21 | 22 | \ENDDOC 23 | -------------------------------------------------------------------------------- /Help/bty.hlp: -------------------------------------------------------------------------------- 1 | \DOC bty 2 | 3 | \TYPE {bty : hol_type} 4 | 5 | \SYNOPSIS 6 | The type variable {`:B`}. 7 | 8 | \DESCRIBE 9 | This name is bound to the HOL type {:B}. 10 | 11 | \FAILURE 12 | Not applicable. 13 | 14 | \USES 15 | Exploiting the very common type variable {:B} inside derived rules (e.g. an 16 | instantiation list for {inst} or {type_subst}) without the inefficiency or 17 | inconvenience of calling a quotation parser or explicit constructor. 18 | 19 | \SEEALSO 20 | aty, bool_ty. 21 | 22 | \ENDDOC 23 | -------------------------------------------------------------------------------- /Help/can.hlp: -------------------------------------------------------------------------------- 1 | \DOC can 2 | 3 | \TYPE {can : ('a -> 'b) -> 'a -> bool} 4 | 5 | \SYNOPSIS 6 | Tests for failure. 7 | 8 | \DESCRIBE 9 | {can f x} evaluates to {true} if the application of {f} to {x} succeeds. 10 | It evaluates to {false} if the application fails with a {Failure _} exception. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \EXAMPLE 16 | { 17 | # can hd [1;2];; 18 | val it : bool = true 19 | # can hd [];; 20 | val it : bool = false 21 | } 22 | 23 | \SEEALSO 24 | check. 25 | 26 | \ENDDOC 27 | -------------------------------------------------------------------------------- /Help/dest_exists.hlp: -------------------------------------------------------------------------------- 1 | \DOC dest_exists 2 | 3 | \TYPE {dest_exists : term -> term * term} 4 | 5 | \SYNOPSIS 6 | Breaks apart an existentially quantified term into quantified variable and 7 | body. 8 | 9 | \DESCRIBE 10 | {dest_exists} is a term destructor for existential quantification: 11 | {dest_exists `?var. t`} returns {(`var`,`t`)}. 12 | 13 | \FAILURE 14 | Fails with {dest_exists} if term is not an existential quantification. 15 | 16 | \SEEALSO 17 | is_exists, mk_exists, strip_exists. 18 | 19 | \ENDDOC 20 | -------------------------------------------------------------------------------- /Help/isalpha.hlp: -------------------------------------------------------------------------------- 1 | \DOC isalpha 2 | 3 | \TYPE {isalpha : string -> bool} 4 | 5 | \SYNOPSIS 6 | Tests if a one-character string is alphabetic. 7 | 8 | \DESCRIBE 9 | The call {isalpha s} tests whether the first character of string {s} (normally 10 | it is the only character) is alphabetic, i.e. an uppercase or lowercase 11 | letter, an underscore or a prime character. 12 | 13 | \FAILURE 14 | Fails if the string is empty. 15 | 16 | \SEEALSO 17 | isalnum, isbra, isnum, issep, isspace, issymb. 18 | 19 | \ENDDOC 20 | -------------------------------------------------------------------------------- /Help/isbra.hlp: -------------------------------------------------------------------------------- 1 | \DOC isbra 2 | 3 | \TYPE {isbra : string -> bool} 4 | 5 | \SYNOPSIS 6 | Tests if a one-character string is some kind of bracket. 7 | 8 | \DESCRIBE 9 | The call {isbra s} tests whether the first character of string {s} (normally 10 | it is the only character) is a bracket, meaning an opening or closing 11 | parenthesis, square bracket or curly brace. 12 | 13 | \FAILURE 14 | Fails if the string is empty. 15 | 16 | \SEEALSO 17 | isalnum, isalpha, isnum, issep, isspace, issymb. 18 | 19 | \ENDDOC 20 | -------------------------------------------------------------------------------- /Help/pp_print_thm.hlp: -------------------------------------------------------------------------------- 1 | \DOC pp_print_thm 2 | 3 | \TYPE {pp_print_thm : formatter -> thm -> unit} 4 | 5 | \SYNOPSIS 6 | Prints a theorem to formatter. 7 | 8 | \DESCRIBE 9 | The call {pp_print_thm fmt th} prints the usual textual representation of the 10 | theorem {th} to the formatter {fmt}. 11 | 12 | \FAILURE 13 | Should never fail unless the formatter does. 14 | 15 | \COMMENTS 16 | The usual case where the formatter is the standard output is {print_thm}. 17 | 18 | \SEEALSO 19 | print_thm. 20 | 21 | \ENDDOC 22 | -------------------------------------------------------------------------------- /Help/print_goal.hlp: -------------------------------------------------------------------------------- 1 | \DOC print_goal 2 | 3 | \TYPE {print_goal : goal -> unit} 4 | 5 | \SYNOPSIS 6 | Print a goal. 7 | 8 | \DESCRIBE 9 | {print_goal g} prints the goal {g} to standard output, with no following 10 | newline. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \COMMENTS 16 | This is invoked automatically when something of type {goal} is produced at 17 | the top level, so manual invocation is not normally needed. 18 | 19 | \SEEALSO 20 | print_goalstack, print_term, PRINT_GOAL_TAC. 21 | 22 | \ENDDOC 23 | -------------------------------------------------------------------------------- /Help/temp_path.hlp: -------------------------------------------------------------------------------- 1 | \DOC temp_path 2 | 3 | \TYPE {temp_path : string ref} 4 | 5 | \SYNOPSIS 6 | Directory in which to create temporary files. 7 | 8 | \DESCRIBE 9 | Some HOL Light derived rules in the libraries (none in the core system) need to 10 | create temporary files. This is the directory in which they do so. 11 | 12 | \FAILURE 13 | Not applicable. 14 | 15 | \EXAMPLE 16 | On my laptop: 17 | { 18 | # !temp_path;; 19 | val it : string = "/tmp" 20 | } 21 | 22 | \SEEALSO 23 | hol_dir. 24 | 25 | \ENDDOC 26 | -------------------------------------------------------------------------------- /Help/ANTS_TAC.hlp: -------------------------------------------------------------------------------- 1 | \DOC ANTS_TAC 2 | 3 | \TYPE {ANTS_TAC : tactic} 4 | 5 | \SYNOPSIS 6 | Split off antecedent of antecedent of goal as a new subgoal. 7 | 8 | \DESCRIBE 9 | { 10 | A ?- (p ==> q) ==> r 11 | ======================= ANTS_TAC 12 | A ?- p A ?- q ==> r 13 | } 14 | 15 | \FAILURE 16 | Fails unless the goal is of the specified form. 17 | 18 | \USES 19 | Convenient for focusing on assumptions of an implicational theorem that one 20 | wants to use. 21 | 22 | \SEEALSO 23 | MP_TAC. 24 | 25 | \ENDDOC 26 | -------------------------------------------------------------------------------- /Help/filter.hlp: -------------------------------------------------------------------------------- 1 | \DOC filter 2 | 3 | \TYPE {filter : ('a -> bool) -> 'a list -> 'a list} 4 | 5 | \SYNOPSIS 6 | Filters a list to the sublist of elements satisfying a predicate. 7 | 8 | \KEYWORDS 9 | list. 10 | 11 | \DESCRIBE 12 | {filter p l} applies {p} to every element of {l}, returning a list of those 13 | that satisfy {p}, in the order they appeared in the original list. 14 | 15 | \FAILURE 16 | Fails if the predicate fails on any element. 17 | 18 | \SEEALSO 19 | mapfilter, partition, remove. 20 | 21 | \ENDDOC 22 | -------------------------------------------------------------------------------- /Help/is_prefix.hlp: -------------------------------------------------------------------------------- 1 | \DOC is_prefix 2 | 3 | \TYPE {is_prefix : string -> bool} 4 | 5 | \SYNOPSIS 6 | Tests if an identifier has prefix status. 7 | 8 | \DESCRIBE 9 | Certain identifiers {c} have prefix status, meaning that combinations of the 10 | form {c f x} will be parsed as {c (f x)} rather than the usual {(c f) x}. The 11 | call {is_prefix "c"} tests if {c} is one of those identifiers. 12 | 13 | \FAILURE 14 | Never fails. 15 | 16 | \SEEALSO 17 | parse_as_prefix, prefixes, unparse_as_prefix. 18 | 19 | \ENDDOC 20 | -------------------------------------------------------------------------------- /Help/.upto.hlp: -------------------------------------------------------------------------------- 1 | \DOC -- 2 | 3 | \TYPE {(--) : int -> int -> int list} 4 | 5 | \SYNOPSIS 6 | Gives a finite list of integers between the given bounds. 7 | 8 | \DESCRIBE 9 | The call {m--n} returns the list of consecutive numbers from {m} to {n}. 10 | 11 | \EXAMPLE 12 | { 13 | # 1--10;; 14 | val it : int list = [1; 2; 3; 4; 5; 6; 7; 8; 9; 10] 15 | # 5--5;; 16 | val it : int list = [5] 17 | # (-1)--1;; 18 | val it : int list = [-1; 0; 1] 19 | # 2--1;; 20 | val it : int list = [] 21 | } 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Help/DENUMERAL.hlp: -------------------------------------------------------------------------------- 1 | \DOC DENUMERAL 2 | 3 | \TYPE {DENUMERAL : thm -> thm} 4 | 5 | \SYNOPSIS 6 | Remove instances of the {NUMERAL} constant from a theorem. 7 | 8 | \DESCRIBE 9 | The call {DENUMERAL th} removes from the conclusion of {th} any instances of 10 | the constant {NUMERAL}, used in the internal representation of numerals. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \USES 16 | Only intended for users manipulating the internal structure of numerals. 17 | 18 | \SEEALSO 19 | NUM_REDUCE_CONV. 20 | 21 | \ENDDOC 22 | -------------------------------------------------------------------------------- /Help/isalnum.hlp: -------------------------------------------------------------------------------- 1 | \DOC isalnum 2 | 3 | \TYPE {isalnum : string -> bool} 4 | 5 | \SYNOPSIS 6 | Tests if a one-character string is alphanumeric. 7 | 8 | \DESCRIBE 9 | The call {isalnum s} tests whether the first character of string {s} (normally 10 | it is the only character) is alphanumeric, i.e. an uppercase or lowercase 11 | letter, a digit, an underscore or a prime character. 12 | 13 | \FAILURE 14 | Fails if the string is empty. 15 | 16 | \SEEALSO 17 | isalpha, isbra, isnum, issep, isspace, issymb. 18 | 19 | \ENDDOC 20 | -------------------------------------------------------------------------------- /Help/mk_exists.hlp: -------------------------------------------------------------------------------- 1 | \DOC mk_exists 2 | 3 | \TYPE {mk_exists : term * term -> term} 4 | 5 | \SYNOPSIS 6 | Term constructor for existential quantification. 7 | 8 | \DESCRIBE 9 | {mk_exists(`v`,`t`)} returns {`?v. t`}. 10 | 11 | \FAILURE 12 | Fails with if first term is not a variable or if {t} is not of type {`:bool`}. 13 | 14 | \EXAMPLE 15 | { 16 | # mk_exists(`x:num`,`x + 1 = 1 + x`);; 17 | val it : term = `?x. x + 1 = 1 + x` 18 | } 19 | 20 | \SEEALSO 21 | dest_exists, is_exists, list_mk_exists. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Help/mk_forall.hlp: -------------------------------------------------------------------------------- 1 | \DOC mk_forall 2 | 3 | \TYPE {mk_forall : term * term -> term} 4 | 5 | \SYNOPSIS 6 | Term constructor for universal quantification. 7 | 8 | \DESCRIBE 9 | {mk_forall(`v`,`t`)} returns {`!v. t`}. 10 | 11 | \FAILURE 12 | Fails with if first term is not a variable or if {t} is not of type {`:bool`}. 13 | 14 | \EXAMPLE 15 | { 16 | # mk_forall(`x:num`,`x + 1 = 1 + x`);; 17 | val it : term = `!x. x + 1 = 1 + x` 18 | } 19 | 20 | \SEEALSO 21 | dest_forall, is_forall, list_mk_forall. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Complex/complex_real.ml: -------------------------------------------------------------------------------- 1 | (* ========================================================================= *) 2 | (* Trivial restriction of complex Groebner bases to reals. *) 3 | (* ========================================================================= *) 4 | 5 | let GROBNER_REAL_ARITH = 6 | let trans_conv = GEN_REWRITE_CONV TOP_SWEEP_CONV 7 | [GSYM CX_INJ; CX_POW; CX_MUL; CX_ADD; CX_NEG; CX_SUB] in 8 | fun tm -> let th = trans_conv tm in 9 | EQ_MP (SYM th) (COMPLEX_ARITH(rand(concl th)));; 10 | -------------------------------------------------------------------------------- /Help/map2.hlp: -------------------------------------------------------------------------------- 1 | \DOC map2 2 | 3 | \TYPE {map2 : ('a -> 'b -> 'c) -> 'a list -> 'b list -> 'c list} 4 | 5 | \SYNOPSIS 6 | Maps a binary function over two lists to create one new list. 7 | 8 | \DESCRIBE 9 | {map2 f ([x1;...;xn],[y1;...;yn])} returns {[f(x1,y1);...;f(xn,yn)]}. 10 | 11 | \FAILURE 12 | Fails with {map2} if the two lists are of different lengths. 13 | 14 | \EXAMPLE 15 | { 16 | # map2 (+) [1;2;3] [30;20;10];; 17 | val it : int list = [31; 22; 13] 18 | } 19 | 20 | \SEEALSO 21 | map, uncurry. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Help/mk_uexists.hlp: -------------------------------------------------------------------------------- 1 | \DOC mk_uexists 2 | 3 | \TYPE {mk_uexists : term * term -> term} 4 | 5 | \SYNOPSIS 6 | Term constructor for unique existence. 7 | 8 | \DESCRIBE 9 | {mk_uexists(`v`,`t`)} returns {`?!v. t`}. 10 | 11 | \FAILURE 12 | Fails with if first term is not a variable or if {t} is not of type {`:bool`}. 13 | 14 | \EXAMPLE 15 | { 16 | # mk_uexists(`n:num`,`prime(n) /\ EVEN(n)`);; 17 | val it : term = `?!n. prime n /\ EVEN n` 18 | } 19 | 20 | \SEEALSO 21 | dest_uexists, is_uexists, mk_exists. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Help/pow2.hlp: -------------------------------------------------------------------------------- 1 | \DOC pow2 2 | 3 | \TYPE {pow2 : int -> num} 4 | 5 | \SYNOPSIS 6 | Returns power of 2 as unlimited-size integer. 7 | 8 | \DESCRIBE 9 | When applied to an integer {n} (type {int}), {pow2} returns $2^n$ as an 10 | unlimited-precision integer (type {num}). The argument may be negative. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \EXAMPLE 16 | { 17 | # pow2(-2);; 18 | val it : num = 1/4 19 | # pow2(64);; 20 | val it : num = 18446744073709551616 21 | } 22 | 23 | \SEEALSO 24 | pow10. 25 | 26 | \ENDDOC 27 | -------------------------------------------------------------------------------- /Help/remove.hlp: -------------------------------------------------------------------------------- 1 | \DOC remove 2 | 3 | \TYPE {remove : ('a -> bool) -> 'a list -> 'a * 'a list} 4 | 5 | \SYNOPSIS 6 | Separates the first element of a list to satisfy a predicate from the rest of 7 | the list. 8 | 9 | \FAILURE 10 | Fails if no element satisfies the predicate. This will always be the case for 11 | an empty list. 12 | 13 | \EXAMPLE 14 | { 15 | # remove (fun x -> x >= 3) [1;2;3;4;5;6];; 16 | val it : int * int list = (3, [1; 2; 4; 5; 6]) 17 | } 18 | 19 | \SEEALSO 20 | partition, filter. 21 | 22 | \ENDDOC 23 | -------------------------------------------------------------------------------- /Help/the_inductive_types.hlp: -------------------------------------------------------------------------------- 1 | \DOC the_inductive_types 2 | 3 | \TYPE {the_inductive_types : (string * (thm * thm)) list ref} 4 | 5 | \SYNOPSIS 6 | List of previously declared inductive types. 7 | 8 | \DESCRIBE 9 | This reference variable contains a list of the inductive types, together with 10 | their induction and recursion theorems as returned by {define_type}. The list 11 | is automatically extended by a call of {define_type}. 12 | 13 | \FAILURE 14 | Not applicable. 15 | 16 | \SEEALSO 17 | define_type. 18 | 19 | \ENDDOC 20 | -------------------------------------------------------------------------------- /Mizarlight/Makefile: -------------------------------------------------------------------------------- 1 | # 2 | # Syntax extensions to support Mizar Light: new infix operators 3 | # 4 | 5 | # Build the camlp4 syntax extension file 6 | 7 | pa_f.cmo: pa_f.ml; if test `ocamlc -version | cut -c1-3` = "3.0" ; \ 8 | then ocamlc -c -pp "camlp4r pa_extend.cmo q_MLast.cmo" -I +camlp4 pa_f.ml; \ 9 | else ocamlfind ocamlc -package camlp5 -c -pp "camlp5r pa_lexer.cmo pa_extend.cmo q_MLast.cmo" -I +camlp5 pa_f.ml; \ 10 | fi 11 | 12 | clean:; rm -f pa_f.cmi pa_f.cmo 13 | -------------------------------------------------------------------------------- /Help/dest_pair.hlp: -------------------------------------------------------------------------------- 1 | \DOC dest_pair 2 | 3 | \TYPE {dest_pair : term -> term * term} 4 | 5 | \SYNOPSIS 6 | Breaks apart a pair into two separate terms. 7 | 8 | \DESCRIBE 9 | {dest_pair} is a term destructor for pairs: {dest_pair `(t1,t2)`} returns 10 | {(`t1`,`t2`)}. 11 | 12 | \FAILURE 13 | Fails with {dest_pair} if term is not a pair. 14 | 15 | \EXAMPLE 16 | { 17 | # dest_pair `(1,2),(3,4),(5,6)`;; 18 | val it : term * term = (`1,2`, `(3,4),5,6`) 19 | } 20 | 21 | \SEEALSO 22 | dest_cons, is_pair, mk_pair. 23 | 24 | \ENDDOC 25 | -------------------------------------------------------------------------------- /Help/is_abs.hlp: -------------------------------------------------------------------------------- 1 | \DOC is_abs 2 | 3 | \TYPE {is_abs : term -> bool} 4 | 5 | \SYNOPSIS 6 | Tests a term to see if it is an abstraction. 7 | 8 | \DESCRIBE 9 | {is_abs `\var. t`} returns {true}. If the term is not an abstraction the 10 | result is {false}. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \EXAMPLE 16 | { 17 | # is_abs `\x. x + 1`;; 18 | val it : bool = true 19 | 20 | # is_abs `!x. x >= 0`;; 21 | val it : bool = false 22 | } 23 | 24 | \SEEALSO 25 | mk_abs, dest_abs, is_var, is_const, is_comb. 26 | 27 | \ENDDOC 28 | -------------------------------------------------------------------------------- /Help/list_mk_abs.hlp: -------------------------------------------------------------------------------- 1 | \DOC list_mk_abs 2 | 3 | \TYPE {list_mk_abs : term list * term -> term} 4 | 5 | \SYNOPSIS 6 | Iteratively constructs abstractions. 7 | 8 | \DESCRIBE 9 | {list_mk_abs([`x1`;...;`xn`],`t`)} returns {`\x1 ... xn. t`}. 10 | 11 | \FAILURE 12 | Fails with {list_mk_abs} if the terms in the list are not variables. 13 | 14 | \EXAMPLE 15 | { 16 | # list_mk_abs([`m:num`; `n:num`],`m + n + 1`);; 17 | val it : term = `\m n. m + n + 1` 18 | } 19 | 20 | \SEEALSO 21 | dest_abs, mk_abs, strip_abs. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Help/pow10.hlp: -------------------------------------------------------------------------------- 1 | \DOC pow10 2 | 3 | \TYPE {pow10 : int -> num} 4 | 5 | \SYNOPSIS 6 | Returns power of 10 as unlimited-size integer. 7 | 8 | \DESCRIBE 9 | When applied to an integer {n} (type {int}), {pow10} returns $10^n$ as an 10 | unlimited-precision integer (type {num}). The argument may be negative. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \EXAMPLE 16 | { 17 | # pow10(-1);; 18 | val it : num = 1/10 19 | # pow10(16);; 20 | val it : num = 10000000000000000 21 | } 22 | 23 | \SEEALSO 24 | pow2. 25 | 26 | \ENDDOC 27 | -------------------------------------------------------------------------------- /Help/variables.hlp: -------------------------------------------------------------------------------- 1 | \DOC variables 2 | 3 | \TYPE {variables : term -> term list} 4 | 5 | \SYNOPSIS 6 | Determines the variables used, free or bound, in a given term. 7 | 8 | \DESCRIBE 9 | Given a term argument, {variables} returns a list of variables that occur free 10 | or bound in that term. 11 | 12 | \EXAMPLE 13 | { 14 | # variables `\a:bool. a`;; 15 | val it : term list = [`a`] 16 | # variables `(a:num) + (b:num)`;; 17 | val it : term list = [`b`; `a`] 18 | } 19 | 20 | \SEEALSO 21 | frees, free_in. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Help/CONJUNCT1.hlp: -------------------------------------------------------------------------------- 1 | \DOC CONJUNCT1 2 | 3 | \TYPE {CONJUNCT1 : thm -> thm} 4 | 5 | \SYNOPSIS 6 | Extracts left conjunct of theorem. 7 | 8 | \KEYWORDS 9 | rule, conjunction. 10 | 11 | \DESCRIBE 12 | { 13 | A |- t1 /\ t2 14 | --------------- CONJUNCT1 15 | A |- t1 16 | } 17 | \FAILURE 18 | Fails unless the input theorem is a conjunction. 19 | 20 | \EXAMPLE 21 | { 22 | # CONJUNCT1(ASSUME `p /\ q`);; 23 | val it : thm = p /\ q |- p 24 | } 25 | 26 | \SEEALSO 27 | CONJ_PAIR, CONJUNCT2, CONJ, CONJUNCTS. 28 | 29 | \ENDDOC 30 | -------------------------------------------------------------------------------- /Help/CONJUNCT2.hlp: -------------------------------------------------------------------------------- 1 | \DOC CONJUNCT2 2 | 3 | \TYPE {CONJUNCT2 : thm -> thm} 4 | 5 | \SYNOPSIS 6 | Extracts right conjunct of theorem. 7 | 8 | \KEYWORDS 9 | rule, conjunction. 10 | 11 | \DESCRIBE 12 | { 13 | A |- t1 /\ t2 14 | --------------- CONJUNCT2 15 | A |- t2 16 | } 17 | \FAILURE 18 | Fails unless the input theorem is a conjunction. 19 | 20 | \EXAMPLE 21 | { 22 | # CONJUNCT2(ASSUME `p /\ q`);; 23 | val it : thm = p /\ q |- q 24 | } 25 | 26 | \SEEALSO 27 | CONJ_PAIR, CONJUNCT1, CONJ, CONJUNCTS. 28 | 29 | \ENDDOC 30 | -------------------------------------------------------------------------------- /Help/find.hlp: -------------------------------------------------------------------------------- 1 | \DOC find 2 | 3 | \TYPE {find : ('a -> bool) -> 'a list -> 'a} 4 | 5 | \SYNOPSIS 6 | Returns the first element of a list which satisfies a predicate. 7 | 8 | \KEYWORDS 9 | list. 10 | 11 | \DESCRIBE 12 | {find p [x1;...;xn]} returns the first {xi} in the list such that {(p xi)} 13 | is {true}. 14 | 15 | \FAILURE 16 | Fails with {find} if no element satisfies the predicate. This will always be 17 | the case if the list is empty. 18 | 19 | \SEEALSO 20 | tryfind, mem, exists, forall, assoc, rev_assoc. 21 | 22 | \ENDDOC 23 | -------------------------------------------------------------------------------- /Help/CONJ.hlp: -------------------------------------------------------------------------------- 1 | \DOC CONJ 2 | 3 | \TYPE {CONJ : thm -> thm -> thm} 4 | 5 | \SYNOPSIS 6 | Introduces a conjunction. 7 | 8 | \KEYWORDS 9 | rule, conjunction. 10 | 11 | \DESCRIBE 12 | { 13 | A1 |- t1 A2 |- t2 14 | ------------------------ CONJ 15 | A1 u A2 |- t1 /\ t2 16 | } 17 | \FAILURE 18 | Never fails. 19 | 20 | \EXAMPLE 21 | { 22 | # CONJ (NUM_REDUCE_CONV `2 + 2`) (ASSUME `p:bool`);; 23 | val it : thm = p |- 2 + 2 = 4 /\ p 24 | } 25 | 26 | \SEEALSO 27 | CONJUNCT1, CONJUNCT2, CONJUNCTS, CONJ_PAIR. 28 | 29 | \ENDDOC 30 | -------------------------------------------------------------------------------- /Help/DISJ2.hlp: -------------------------------------------------------------------------------- 1 | \DOC DISJ2 2 | 3 | \TYPE {DISJ2 : term -> thm -> thm} 4 | 5 | \SYNOPSIS 6 | Introduces a left disjunct into the conclusion of a theorem. 7 | 8 | \KEYWORDS 9 | rule, disjunction. 10 | 11 | \DESCRIBE 12 | { 13 | A |- t2 14 | --------------- DISJ2 `t1` 15 | A |- t1 \/ t2 16 | } 17 | \FAILURE 18 | Fails if the term argument is not boolean. 19 | 20 | \EXAMPLE 21 | { 22 | # DISJ2 `F` TRUTH;; 23 | val it : thm = |- F \/ T 24 | } 25 | 26 | \SEEALSO 27 | DISJ1, DISJ1_TAC, DISJ2_TAC, DISJ_CASES. 28 | 29 | \ENDDOC 30 | -------------------------------------------------------------------------------- /Help/infixes.hlp: -------------------------------------------------------------------------------- 1 | \DOC infixes 2 | 3 | \TYPE {infixes : unit -> (string * (int * string)) list} 4 | 5 | \SYNOPSIS 6 | Lists the infixes currently recognized by the parser. 7 | 8 | \DESCRIBE 9 | The function {infixes} should be applied to the unit {()} and will then return 10 | a list of all the infixes currently recognized by the parser together with 11 | their precedence and associativity (left or right). 12 | 13 | \FAILURE 14 | Never fails. 15 | 16 | \SEEALSO 17 | get_infix_status, parse_as_infix, unparse_as_infix. 18 | 19 | \ENDDOC 20 | -------------------------------------------------------------------------------- /Help/is_comb.hlp: -------------------------------------------------------------------------------- 1 | \DOC is_comb 2 | 3 | \TYPE {is_comb : term -> bool} 4 | 5 | \SYNOPSIS 6 | Tests a term to see if it is a combination (function application). 7 | 8 | \DESCRIBE 9 | {is_comb "t1 t2"} returns {true}. If the term is not a combination the result 10 | is {false}. 11 | 12 | \FAILURE 13 | Never fails 14 | 15 | \EXAMPLE 16 | { 17 | # is_comb `x + 1`;; 18 | val it : bool = true 19 | # is_comb `T`;; 20 | val it : bool = false 21 | } 22 | 23 | \SEEALSO 24 | dest_comb, is_var, is_const, is_abs, mk_comb. 25 | 26 | \ENDDOC 27 | -------------------------------------------------------------------------------- /Help/strip_forall.hlp: -------------------------------------------------------------------------------- 1 | \DOC strip_forall 2 | 3 | \TYPE {strip_forall : term -> term list * term} 4 | 5 | \SYNOPSIS 6 | Iteratively breaks apart universal quantifications. 7 | 8 | \DESCRIBE 9 | {strip_forall `!x1 ... xn. t`} returns {([`x1`;...;`xn`],`t`)}. Note that 10 | { 11 | strip_forall(list_mk_forall([`x1`;...;`xn`],`t`)) 12 | } 13 | \noindent will not return {([`x1`;...;`xn`],`t`)} if {t} is a universal 14 | quantification. 15 | 16 | \FAILURE 17 | Never fails. 18 | 19 | \SEEALSO 20 | dest_forall, list_mk_forall. 21 | 22 | \ENDDOC 23 | -------------------------------------------------------------------------------- /Help/concl.hlp: -------------------------------------------------------------------------------- 1 | \DOC concl 2 | 3 | \TYPE {concl : thm -> term} 4 | 5 | \SYNOPSIS 6 | Returns the conclusion of a theorem. 7 | 8 | \DESCRIBE 9 | When applied to a theorem {A |- t}, the function {concl} returns {t}. 10 | 11 | \FAILURE 12 | Never fails. 13 | 14 | \EXAMPLE 15 | { 16 | # ADD_SYM;; 17 | val it : thm = |- !m n. m + n = n + m 18 | # concl ADD_SYM;; 19 | val it : term = `!m n. m + n = n + m` 20 | 21 | # concl (ASSUME `1 = 0`);; 22 | val it : term = `1 = 0` 23 | } 24 | 25 | \SEEALSO 26 | dest_thm, hyp. 27 | 28 | \ENDDOC 29 | -------------------------------------------------------------------------------- /Help/strip_exists.hlp: -------------------------------------------------------------------------------- 1 | \DOC strip_exists 2 | 3 | \TYPE {strip_exists : term -> term list * term} 4 | 5 | \SYNOPSIS 6 | Iteratively breaks apart existential quantifications. 7 | 8 | \DESCRIBE 9 | {strip_exists `?x1 ... xn. t`} returns {([`x1`;...;`xn`],`t`)}. Note that 10 | { 11 | strip_exists(list_mk_exists([`x1`;...;`xn`],`t`)) 12 | } 13 | \noindent will not return {([`x1`;...;`xn`],`t`)} if {t} is an existential 14 | quantification. 15 | 16 | \FAILURE 17 | Never fails. 18 | 19 | \SEEALSO 20 | dest_exists, list_mk_exists. 21 | 22 | \ENDDOC 23 | -------------------------------------------------------------------------------- /pa_j/README: -------------------------------------------------------------------------------- 1 | These .ml files are copies of the lib/plexer.ml and etc/pa_o.ml files in 2 | the camlp5 project (https://github.com/camlp5/camlp5) with proper edits to 3 | enable lexing specific to HOL Light and `..`-style quotation. 4 | The edited parts are annotated with comments including specific keywords such as 5 | "JRH" or "HOL Light". 6 | A user can disable/enable this lexer using `unset_jrh_lexer;;`/`set_jrh_lexer;;`. 7 | If a file has name 'pa_j_AA_BB.ml', it must be used for OCaml version AA and 8 | camlp5 version BB. '_BB' is omitted in some old files. -------------------------------------------------------------------------------- /Help/chop_list.hlp: -------------------------------------------------------------------------------- 1 | \DOC chop_list 2 | 3 | \TYPE {chop_list : int -> 'a list -> 'a list * 'a list} 4 | 5 | \SYNOPSIS 6 | Chops a list into two parts at a specified point. 7 | 8 | \DESCRIBE 9 | {chop_list i [x1;...;xn]} returns {([x1;...;xi],[x(i+1);...;xn])}. 10 | 11 | \FAILURE 12 | Fails with {chop_list} if {i} is negative or greater than the length of the 13 | list. 14 | 15 | \EXAMPLE 16 | { 17 | # chop_list 3 [1;2;3;4;5];; 18 | val it : int list * int list = ([1; 2; 3], [4; 5]) 19 | } 20 | 21 | \SEEALSO 22 | partition. 23 | 24 | \ENDDOC 25 | -------------------------------------------------------------------------------- /Help/distinctness_store.hlp: -------------------------------------------------------------------------------- 1 | \DOC distinctness_store 2 | 3 | \TYPE {distinctness_store : (string * thm) list ref} 4 | 5 | \SYNOPSIS 6 | Internal theorem list of distinctness theorems. 7 | 8 | \DESCRIBE 9 | This list contains all the distinctness theorems (see {distinct}) for the 10 | recursive types defined so far. It is automatically extended by {define_type} 11 | and used as a cache by {distinct}. 12 | 13 | \FAILURE 14 | Not applicable. 15 | 16 | \SEEALSO 17 | define_type, distinctness, extend_rectype_net, injectivity_store. 18 | 19 | \ENDDOC 20 | -------------------------------------------------------------------------------- /Help/get_const_type.hlp: -------------------------------------------------------------------------------- 1 | \DOC get_const_type 2 | 3 | \TYPE {get_const_type : string -> hol_type} 4 | 5 | \SYNOPSIS 6 | Gets the generic type of a constant from the name of the constant. 7 | 8 | \DESCRIBE 9 | {get_const_type "c"} returns the generic type of {`c`}, if {`c`} is a constant. 10 | 11 | \FAILURE 12 | {get_const_type st} fails if {st} is not the name of a constant. 13 | 14 | \EXAMPLE 15 | { 16 | # get_const_type "COND";; 17 | val it : hol_type = `:bool->A->A->A` 18 | } 19 | 20 | \SEEALSO 21 | dest_const, is_const. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Help/hyp.hlp: -------------------------------------------------------------------------------- 1 | \DOC hyp 2 | 3 | \TYPE {hyp : thm -> term list} 4 | 5 | \SYNOPSIS 6 | Returns the hypotheses of a theorem. 7 | 8 | \DESCRIBE 9 | When applied to a theorem {A |- t}, the function {hyp} returns {A}, the 10 | list of hypotheses of the theorem. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \EXAMPLE 16 | { 17 | # let th = ADD_ASSUM `x = 1` (ASSUME `y = 2`);; 18 | val th : thm = y = 2, x = 1 |- y = 2 19 | 20 | # hyp th;; 21 | val it : term list = [`y = 2`; `x = 1`] 22 | } 23 | 24 | \SEEALSO 25 | dest_thm, concl. 26 | 27 | \ENDDOC 28 | -------------------------------------------------------------------------------- /Help/injectivity_store.hlp: -------------------------------------------------------------------------------- 1 | \DOC injectivity_store 2 | 3 | \TYPE {injectivity_store : (string * thm) list ref} 4 | 5 | \SYNOPSIS 6 | Internal theorem list of injectivity theorems. 7 | 8 | \DESCRIBE 9 | This list contains all the injectivity theorems (see {injectivity}) for the 10 | recursive types defined so far. It is automatically extended by {define_type} 11 | and used as a cache by {injectivity}. 12 | 13 | \FAILURE 14 | Not applicable. 15 | 16 | \SEEALSO 17 | define_type, distinctness_store, extend_rectype_net, injectivity. 18 | 19 | \ENDDOC 20 | -------------------------------------------------------------------------------- /Help/print_goalstack.hlp: -------------------------------------------------------------------------------- 1 | \DOC print_goalstack 2 | 3 | \TYPE {print_goalstack : goalstack -> unit} 4 | 5 | \SYNOPSIS 6 | Print a goalstack. 7 | 8 | \DESCRIBE 9 | {print_goalstack gs} prints the goalstack {gs} to standard output, with no 10 | following newline. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \COMMENTS 16 | This is invoked automatically when something of type {goalstack} is produced at 17 | the top level, so manual invocation is not normally needed. 18 | 19 | \SEEALSO 20 | print_goal, print_term, PRINT_GOAL_TAC. 21 | 22 | \ENDDOC 23 | -------------------------------------------------------------------------------- /miz3/Samples/sample.ml: -------------------------------------------------------------------------------- 1 | horizon := 1;; 2 | 3 | thm `; 4 | let R be num->num->bool; 5 | assume !x. R x x [1]; 6 | assume !x y z. R x y /\ R y z ==> R x z [2]; 7 | thus (!m n. m <= n ==> R m n) <=> (!n. R n (SUC n)) 8 | proof 9 | now [3] // back direction first 10 | assume !n. R n (SUC n); 11 | let m n be num; 12 | !d. R m (m + d) ==> R m (m + SUC d) [4] by 2,ADD_CLAUSES; 13 | R m (m + 0) by 1,ADD_CLAUSES; 14 | !d. R m (m + d) by 4,INDUCT_TAC; 15 | thus m <= n ==> R m n by LE_EXISTS; 16 | end; 17 | !n. n <= SUC n; 18 | qed by 3`;; 19 | 20 | -------------------------------------------------------------------------------- /Help/DISJ1.hlp: -------------------------------------------------------------------------------- 1 | \DOC DISJ1 2 | 3 | \TYPE {DISJ1 : thm -> term -> thm} 4 | 5 | \SYNOPSIS 6 | Introduces a right disjunct into the conclusion of a theorem. 7 | 8 | \KEYWORDS 9 | rule, disjunction. 10 | 11 | \DESCRIBE 12 | { 13 | A |- t1 14 | --------------- DISJ1 (A |- t1) `t2` 15 | A |- t1 \/ t2 16 | } 17 | \FAILURE 18 | Fails unless the term argument is boolean. 19 | 20 | \EXAMPLE 21 | { 22 | # DISJ1 TRUTH `F`;; 23 | val it : thm = |- T \/ F 24 | } 25 | 26 | \SEEALSO 27 | DISJ1_TAC, DISJ2, DISJ2_TAC, DISJ_CASES. 28 | 29 | \ENDDOC 30 | -------------------------------------------------------------------------------- /Help/print_qterm.hlp: -------------------------------------------------------------------------------- 1 | \DOC print_qterm 2 | 3 | \TYPE {print_qterm : term -> unit} 4 | 5 | \SYNOPSIS 6 | Prints a HOL term with surrounding quotes to standard output. 7 | 8 | \DESCRIBE 9 | The call {print_term tm} prints the usual textual representation of the 10 | term {tm} to the standard output, that is {`:tm`}. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \COMMENTS 16 | This is the function that is invoked automatically in the toplevel when 17 | printing terms. 18 | 19 | \SEEALSO 20 | pp_print_qterm, pp_print_term, print_term. 21 | 22 | \ENDDOC 23 | -------------------------------------------------------------------------------- /Help/is_reserved_word.hlp: -------------------------------------------------------------------------------- 1 | \DOC is_reserved_word 2 | 3 | \TYPE {is_reserved_word : string -> bool} 4 | 5 | \SYNOPSIS 6 | Tests if a string is one of the reserved words. 7 | 8 | \DESCRIBE 9 | Certain identifiers in HOL are reserved, e.g. `{if}', `{let}' and `{|}', 10 | meaning that they are special to the parser and cannot be used as ordinary 11 | identifiers. The call {is_reserved_word s} tests if the string {s} is one of 12 | them. 13 | 14 | \FAILURE 15 | Never fails. 16 | 17 | \SEEALSO 18 | reserved_words, reserve_words, unreserve_words. 19 | 20 | \ENDDOC 21 | -------------------------------------------------------------------------------- /Help/print_fpf.hlp: -------------------------------------------------------------------------------- 1 | \DOC print_fpf 2 | 3 | \TYPE {print_fpf : ('a, 'b) func -> unit} 4 | 5 | \SYNOPSIS 6 | Print a finite partial function. 7 | 8 | \DESCRIBE 9 | This prints a finite partial function but only as a trivial string `{}'. 10 | Installed automatically at the top level and probably not useful for most 11 | users. 12 | 13 | \FAILURE 14 | Never fails. 15 | 16 | \SEEALSO 17 | |->, |=>, apply, applyd, choose, combine, defined, dom, foldl, foldr, 18 | graph, is_undefined, mapf, pp_print_fpf, ran, tryapplyd, undefine, undefined. 19 | 20 | \ENDDOC 21 | -------------------------------------------------------------------------------- /Proofrecording/tools/init.ml: -------------------------------------------------------------------------------- 1 | (* ------------------------------------------------------------------------- *) 2 | (* Set up a quotation expander for my `...` quotes. *) 3 | (* ------------------------------------------------------------------------- *) 4 | 5 | let quotexpander s = 6 | if String.sub s 0 1 = ":" then 7 | "parse_type \""^ 8 | (String.escaped (String.sub s 1 (String.length s - 1)))^"\"" 9 | else "parse_term \""^(String.escaped s)^"\"";; 10 | 11 | Quotation.add "tot" (Quotation.ExStr (fun x -> quotexpander));; 12 | -------------------------------------------------------------------------------- /Help/EQF_ELIM.hlp: -------------------------------------------------------------------------------- 1 | \DOC EQF_ELIM 2 | 3 | \TYPE {EQF_ELIM : thm -> thm} 4 | 5 | \SYNOPSIS 6 | Replaces equality with {F} by negation. 7 | 8 | \KEYWORDS 9 | rule, negation, falsity. 10 | 11 | \DESCRIBE 12 | { 13 | A |- tm <=> F 14 | --------------- EQF_ELIM 15 | A |- ~tm 16 | } 17 | 18 | \FAILURE 19 | Fails if the argument theorem is not of the form {A |- tm <=> F}. 20 | 21 | \EXAMPLE 22 | { 23 | # EQF_ELIM(REFL `F`);; 24 | val it : thm = |- ~F 25 | } 26 | 27 | \SEEALSO 28 | EQF_INTRO, EQT_ELIM, EQT_INTRO, NOT_ELIM, NOT_INTRO. 29 | 30 | \ENDDOC 31 | -------------------------------------------------------------------------------- /Help/EQT_ELIM.hlp: -------------------------------------------------------------------------------- 1 | \DOC EQT_ELIM 2 | 3 | \TYPE {EQT_ELIM : thm -> thm} 4 | 5 | \SYNOPSIS 6 | Eliminates equality with {T}. 7 | 8 | \KEYWORDS 9 | rule, truth. 10 | 11 | \DESCRIBE 12 | { 13 | A |- tm <=> T 14 | --------------- EQT_ELIM 15 | A |- tm 16 | } 17 | \FAILURE 18 | Fails if the argument theorem is not of the form {A |- tm <=> T}. 19 | 20 | \EXAMPLE 21 | { 22 | # REFL `T`;; 23 | val it : thm = |- T <=> T 24 | 25 | # EQT_ELIM it;; 26 | val it : thm = |- T 27 | } 28 | 29 | \SEEALSO 30 | EQF_ELIM, EQF_INTRO, EQT_INTRO. 31 | 32 | \ENDDOC 33 | -------------------------------------------------------------------------------- /Help/find_term.hlp: -------------------------------------------------------------------------------- 1 | \DOC find_term 2 | 3 | \TYPE {find_term : (term -> bool) -> term -> term} 4 | 5 | \SYNOPSIS 6 | Searches a term for a subterm that satisfies a given predicate. 7 | 8 | \DESCRIBE 9 | The largest subterm, in a depth-first, left-to-right search 10 | of the given term, that satisfies the predicate is returned. 11 | 12 | \FAILURE 13 | Fails if no subterm of the given term satisfies the predicate. 14 | 15 | \EXAMPLE 16 | { 17 | # find_term is_var `x + y + z`;; 18 | val it : term = `x` 19 | } 20 | 21 | \SEEALSO 22 | find_terms. 23 | 24 | \ENDDOC 25 | -------------------------------------------------------------------------------- /Help/print_qtype.hlp: -------------------------------------------------------------------------------- 1 | \DOC print_qtype 2 | 3 | \TYPE {print_qtype : hol_type -> unit} 4 | 5 | \SYNOPSIS 6 | Prints a type with colon and surrounding quotes to standard output. 7 | 8 | \DESCRIBE 9 | The call {print_type ty} prints the usual textual representation of the 10 | type {ty} to the standard output, that is {`:ty`}. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \COMMENTS 16 | This is the function that is invoked automatically in the toplevel when 17 | printing types. 18 | 19 | \SEEALSO 20 | pp_print_qtype, pp_print_type, print_type. 21 | 22 | \ENDDOC 23 | -------------------------------------------------------------------------------- /Help/end_itlist.hlp: -------------------------------------------------------------------------------- 1 | \DOC end_itlist 2 | 3 | \TYPE {end_itlist : ('a -> 'a -> 'a) -> 'a list -> 'a} 4 | 5 | \SYNOPSIS 6 | List iteration function. Applies a binary function between adjacent elements 7 | of a list. 8 | 9 | \DESCRIBE 10 | {end_itlist f [x1;...;xn]} returns {f x1 ( ... (f x(n-1) xn)...)}. 11 | Returns {x} for a one-element list {[x]}. 12 | 13 | \FAILURE 14 | Fails with {end_itlist} if list is empty. 15 | 16 | \EXAMPLE 17 | { 18 | # end_itlist (+) [1;2;3;4];; 19 | val it : int = 10 20 | } 21 | 22 | \SEEALSO 23 | itlist, rev_itlist. 24 | 25 | \ENDDOC 26 | -------------------------------------------------------------------------------- /Help/parse_as_prefix.hlp: -------------------------------------------------------------------------------- 1 | \DOC parse_as_prefix 2 | 3 | \TYPE {parse_as_prefix : string -> unit} 4 | 5 | \SYNOPSIS 6 | Gives an identifier prefix status. 7 | 8 | \DESCRIBE 9 | Certain identifiers {c} have prefix status, meaning that combinations of the 10 | form {c f x} will be parsed as {c (f x)} rather than the usual {(c f) x}. The 11 | call {parse_as_prefix "c"} adds {c} to the list of such identifiers. 12 | 13 | \FAILURE 14 | Never fails, even if the string already has prefix status. 15 | 16 | \SEEALSO 17 | is_prefix, prefixes, unparse_as_prefix. 18 | 19 | \ENDDOC 20 | -------------------------------------------------------------------------------- /Help/current_goalstack.hlp: -------------------------------------------------------------------------------- 1 | \DOC current_goalstack 2 | 3 | \TYPE {current_goalstack : goalstack ref} 4 | 5 | \SYNOPSIS 6 | Reference variable holding current goalstack. 7 | 8 | \DESCRIBE 9 | The reference variable {current_goalstack} contains the current goalstack. A 10 | goalstack is a type containing a list of goalstates. 11 | 12 | \FAILURE 13 | Not applicable. 14 | 15 | \COMMENTS 16 | Users will probably not often want to examine this variable explicitly, since 17 | various proof commands modify it in various ways. 18 | 19 | \SEEALSO 20 | b, g, e, r. 21 | 22 | \ENDDOC 23 | -------------------------------------------------------------------------------- /Help/dest_abs.hlp: -------------------------------------------------------------------------------- 1 | \DOC dest_abs 2 | 3 | \TYPE {dest_abs : term -> term * term} 4 | 5 | \SYNOPSIS 6 | Breaks apart an abstraction into abstracted variable and body. 7 | 8 | \DESCRIBE 9 | {dest_abs} is a term destructor for abstractions: 10 | {dest_abs `\var. t`} returns {(`var`,`t`)}. 11 | 12 | \FAILURE 13 | Fails with {dest_abs} if term is not an abstraction. 14 | 15 | \EXAMPLE 16 | { 17 | # dest_abs `\x. x + 1`;; 18 | val it : term * term = (`x`, `x + 1`) 19 | } 20 | 21 | \SEEALSO 22 | dest_comb, dest_const, dest_var, is_abs, mk_abs, strip_abs. 23 | 24 | \ENDDOC 25 | -------------------------------------------------------------------------------- /Help/dest_cons.hlp: -------------------------------------------------------------------------------- 1 | \DOC dest_cons 2 | 3 | \TYPE {dest_cons : term -> term * term} 4 | 5 | \SYNOPSIS 6 | Breaks apart a `CONS pair' into head and tail. 7 | 8 | \DESCRIBE 9 | {dest_cons} is a term destructor for `CONS pairs'. When applied to a term 10 | representing a nonempty list {`[t;t1;...;tn]`} (which is equivalent to {`CONS t 11 | [t1;...;tn]`}), it returns the pair of terms {(`t`,`[t1;...;tn]`)}. 12 | 13 | \FAILURE 14 | Fails with {dest_cons} if the term is not a non-empty list. 15 | 16 | \SEEALSO 17 | dest_list, is_cons, is_list, mk_cons, mk_list. 18 | 19 | \ENDDOC 20 | -------------------------------------------------------------------------------- /Help/mk_var.hlp: -------------------------------------------------------------------------------- 1 | \DOC mk_var 2 | 3 | \TYPE {mk_var : string * hol_type -> term} 4 | 5 | \SYNOPSIS 6 | Constructs a variable of given name and type. 7 | 8 | \DESCRIBE 9 | {mk_var("var",`:ty`)} returns the variable {`var:ty`}. 10 | 11 | \FAILURE 12 | Never fails. 13 | 14 | \COMMENTS 15 | {mk_var} can be used to construct variables with names which are not 16 | acceptable to the term parser. In particular, a variable with the name of a 17 | known constant can be constructed using {mk_var}. 18 | 19 | \SEEALSO 20 | dest_var, is_var, mk_const, mk_comb, mk_abs. 21 | 22 | \ENDDOC 23 | -------------------------------------------------------------------------------- /Help/unions.hlp: -------------------------------------------------------------------------------- 1 | \DOC unions 2 | 3 | \TYPE {unions : 'a list list -> 'a list} 4 | 5 | \SYNOPSIS 6 | Performs the union of a set of sets. 7 | 8 | \DESCRIBE 9 | Applied to a list of lists, {union} returns a list of all the elements of them, 10 | in some unspecified order, with no repetitions. It can be considered as the 11 | union of the family of `sets'. 12 | 13 | \FAILURE 14 | Never fails. 15 | 16 | \EXAMPLE 17 | { 18 | # unions [[1;2]; [2;2;2;]; [2;3;4;5]];; 19 | val it : int list = [1; 2; 3; 4; 5] 20 | } 21 | 22 | \SEEALSO 23 | intersect, subtract. 24 | 25 | \ENDDOC 26 | -------------------------------------------------------------------------------- /Help/uniq.hlp: -------------------------------------------------------------------------------- 1 | \DOC uniq 2 | 3 | \TYPE {uniq : 'a list -> 'a list} 4 | 5 | \SYNOPSIS 6 | Eliminate adjacent identical elements from a list. 7 | 8 | \DESCRIBE 9 | When applied to a list, {uniq} gives a new list that results from coalescing 10 | adjacent (only) elements of the list into one. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \EXAMPLE 16 | { 17 | # uniq [1;2;3;1;2;3];; 18 | val it : int list = [1; 2; 3; 1; 2; 3] 19 | 20 | # uniq [1;1;1;2;3;3;3;3;4];; 21 | val it : int list = [1; 2; 3; 4] 22 | } 23 | 24 | \SEEALSO 25 | setify, sort. 26 | 27 | \ENDDOC 28 | -------------------------------------------------------------------------------- /Formal_ineqs/tests/data/gen_nat_data.py: -------------------------------------------------------------------------------- 1 | #!/usr/bin/python 2 | 3 | # -*- coding: utf-8 -*- 4 | """ 5 | Created on Tue May 12 12:23:42 2015 6 | 7 | @author: alexey 8 | """ 9 | 10 | import sys 11 | import random 12 | 13 | if len(sys.argv) != 3: 14 | print "Usage: gen_nat_data length n" 15 | sys.exit(1) 16 | 17 | length = int(sys.argv[1]) 18 | n = int(sys.argv[2]) 19 | 20 | assert (length >= 1 and length <= 1000) 21 | assert (n >= 1 and n <= 10000000) 22 | 23 | a = 10 ** length 24 | b = 10 ** (length + 1) - 1 25 | 26 | for i in range(0, n): 27 | print random.randint(a, b) -------------------------------------------------------------------------------- /Help/delete_parser.hlp: -------------------------------------------------------------------------------- 1 | \DOC delete_parser 2 | 3 | \TYPE {delete_parser : string -> unit} 4 | 5 | \SYNOPSIS 6 | Uninstall a user parser. 7 | 8 | \DESCRIBE 9 | HOL Light allows user parsing functions to be installed, and will try them on 10 | all terms during parsing before the usual parsers. The call {delete_parser "s"} 11 | removes any parsers associated with string {"s"}. 12 | 13 | \FAILURE 14 | Never fails, regardless of whether there are any parsers associated with the 15 | string. 16 | 17 | \SEEALSO 18 | install_parser, installed_parsers, try_user_parser. 19 | 20 | \ENDDOC 21 | -------------------------------------------------------------------------------- /Help/is_let.hlp: -------------------------------------------------------------------------------- 1 | \DOC is_let 2 | 3 | \TYPE {is_let : term -> bool} 4 | 5 | \SYNOPSIS 6 | Tests a term to see if it is a {let}-expression. 7 | 8 | \DESCRIBE 9 | {is_let `let x1 = e1 and ... and xn = en in E`} returns {true}. 10 | If the term is not a {let}-expression of any kind, the result is {false}. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \EXAMPLE 16 | { 17 | # is_let `let x = 1 in x + x`;; 18 | val it : bool = true 19 | 20 | # is_let `let x = 2 and y = 3 in y + x`;; 21 | val it : bool = true 22 | } 23 | 24 | \SEEALSO 25 | mk_let, dest_let. 26 | 27 | \ENDDOC 28 | -------------------------------------------------------------------------------- /Help/rev_itlist.hlp: -------------------------------------------------------------------------------- 1 | \DOC rev_itlist 2 | 3 | \TYPE {rev_itlist : ('a -> 'b -> 'b) -> 'a list -> 'b -> 'b} 4 | 5 | \SYNOPSIS 6 | Applies a binary function between adjacent elements of the reverse of a list. 7 | 8 | \KEYWORDS 9 | list. 10 | 11 | \DESCRIBE 12 | {rev_itlist f [x1;...;xn] y} returns {f xn ( ... (f x2 (f x1 y))...)}. 13 | It returns {y} if the list is empty. 14 | 15 | \FAILURE 16 | Never fails. 17 | 18 | \EXAMPLE 19 | { 20 | # rev_itlist (fun x y -> x * y) [1;2;3;4] 1;; 21 | val it : int = 24 22 | } 23 | 24 | \SEEALSO 25 | itlist, end_itlist. 26 | 27 | \ENDDOC 28 | -------------------------------------------------------------------------------- /Help/curry.hlp: -------------------------------------------------------------------------------- 1 | \DOC curry 2 | 3 | \TYPE {curry : ('a * 'b -> 'c) -> 'a -> 'b -> 'c} 4 | 5 | \SYNOPSIS 6 | Converts a function on a pair to a corresponding curried function. 7 | 8 | \DESCRIBE 9 | The application {curry f} returns {\x y. f(x,y)}, so that 10 | { 11 | curry f x y = f(x,y) 12 | } 13 | \FAILURE 14 | Never fails. 15 | 16 | \EXAMPLE 17 | { 18 | # curry mk_var;; 19 | val it : string -> hol_type -> term = 20 | # it "x";; 21 | val it : hol_type -> term = 22 | # it `:bool`;; 23 | val it : term = `x` 24 | } 25 | \SEEALSO 26 | uncurry. 27 | 28 | \ENDDOC 29 | -------------------------------------------------------------------------------- /Help/explode.hlp: -------------------------------------------------------------------------------- 1 | \DOC explode 2 | 3 | \TYPE {explode : string -> string list} 4 | 5 | \SYNOPSIS 6 | Converts a string into a list of single-character strings. 7 | 8 | \DESCRIBE 9 | {explode s} returns the list of single-character strings that make up {s}, in 10 | the order in which they appear in {s}. If {s} is the empty string, then an 11 | empty list is returned. 12 | 13 | \FAILURE 14 | Never fails. 15 | 16 | \EXAMPLE 17 | { 18 | # explode "example";; 19 | val it : string list = ["e"; "x"; "a"; "m"; "p"; "l"; "e"] 20 | } 21 | 22 | \SEEALSO 23 | implode. 24 | 25 | \ENDDOC 26 | -------------------------------------------------------------------------------- /LP_arith/Makefile: -------------------------------------------------------------------------------- 1 | CDDLIBPATH=/usr/lib/x86_64-linux-gnu 2 | CDDINCLUDEPATH=/usr/include/cdd 3 | 4 | GXX = g++ 5 | CC = $(GXX) 6 | CCLD = $(CC) 7 | 8 | 9 | .SUFFIXES: .c .o 10 | 11 | COMPILE = $(CC) -O2 -DHAVE_LIBGMP=1 -DGMPRATIONAL -I. -I$(CDDLIBPATH) -I$(CDDINCLUDEPATH) 12 | 13 | LINK = $(CCLD) $(AM_CFLAGS) $(CFLAGS) $(AM_LDFLAGS) $(LDFLAGS) -o $@ 14 | 15 | .c.o: 16 | $(COMPILE) -c `test -f '$<' || echo ''`$< 17 | 18 | all: cdd_cert 19 | 20 | cdd_cert: cdd_cert.o 21 | $(LINK) cdd_cert.o $(CDDLIBPATH)/libcddgmp.so -l gmp 22 | 23 | clean: 24 | rm -f *~ *.o cdd_cert cdd_cert.exe 25 | -------------------------------------------------------------------------------- /Help/installed_parsers.hlp: -------------------------------------------------------------------------------- 1 | \DOC installed_parsers 2 | 3 | \TYPE {installed_parsers : unit -> (string * (lexcode list -> preterm * lexcode list)) list} 4 | 5 | \SYNOPSIS 6 | List the user parsers currently installed. 7 | 8 | \DESCRIBE 9 | HOL Light allows user parsing functions to be installed, and will try them on 10 | all terms during parsing before the usual parsers. The call 11 | {installed_parsers()} lists the parsing functions that have been so installed. 12 | 13 | \FAILURE 14 | Never fails. 15 | 16 | \SEEALSO 17 | delete_parser, install_parser, try_user_parser. 18 | 19 | \ENDDOC 20 | -------------------------------------------------------------------------------- /Help/pp_print_qterm.hlp: -------------------------------------------------------------------------------- 1 | \DOC pp_print_qterm 2 | 3 | \TYPE {pp_print_qterm : formatter -> term -> unit} 4 | 5 | \SYNOPSIS 6 | Prints a term with surrounding quotes to formatter. 7 | 8 | \DESCRIBE 9 | The call {pp_print_term fmt tm} prints the usual textual representation of the 10 | term {tm} to the formatter {fmt}, in the form {`tm`}. 11 | 12 | \FAILURE 13 | Should never fail unless the formatter does. 14 | 15 | \COMMENTS 16 | The usual case where the formatter is the standard output is {print_qterm}. 17 | 18 | \SEEALSO 19 | pp_print_term, print_qterm, print_term. 20 | 21 | \ENDDOC 22 | -------------------------------------------------------------------------------- /Help/rator.hlp: -------------------------------------------------------------------------------- 1 | \DOC rator 2 | 3 | \TYPE {rator : term -> term} 4 | 5 | \SYNOPSIS 6 | Returns the operator from a combination (function application). 7 | 8 | \DESCRIBE 9 | {rator(`t1 t2`)} returns {`t1`}. 10 | 11 | \FAILURE 12 | Fails with {rator} if term is not a combination. 13 | 14 | \EXAMPLE 15 | { 16 | # rator `f(x)`;; 17 | Warning: inventing type variables 18 | val it : term = `f` 19 | 20 | # rator `~p`;; 21 | val it : term = `(~)` 22 | 23 | # rator `x + y`;; 24 | val it : term = `(+) x` 25 | } 26 | 27 | \SEEALSO 28 | dest_comb, lhand, lhs, rand. 29 | 30 | \ENDDOC 31 | -------------------------------------------------------------------------------- /Help/CONJ_TAC.hlp: -------------------------------------------------------------------------------- 1 | \DOC CONJ_TAC 2 | 3 | \TYPE {CONJ_TAC : tactic} 4 | 5 | \SYNOPSIS 6 | Reduces a conjunctive goal to two separate subgoals. 7 | 8 | \KEYWORDS 9 | tactic, conjunction. 10 | 11 | \DESCRIBE 12 | When applied to a goal {A ?- t1 /\ t2}, the tactic {CONJ_TAC} reduces it to the 13 | two subgoals corresponding to each conjunct separately. 14 | { 15 | A ?- t1 /\ t2 16 | ====================== CONJ_TAC 17 | A ?- t1 A ?- t2 18 | } 19 | 20 | \FAILURE 21 | Fails unless the conclusion of the goal is a conjunction. 22 | 23 | \SEEALSO 24 | STRIP_TAC. 25 | 26 | \ENDDOC 27 | -------------------------------------------------------------------------------- /Help/pp_print_term.hlp: -------------------------------------------------------------------------------- 1 | \DOC pp_print_term 2 | 3 | \TYPE {pp_print_term : formatter -> term -> unit} 4 | 5 | \SYNOPSIS 6 | Prints a term (without quotes) to formatter. 7 | 8 | \DESCRIBE 9 | The call {pp_print_term fmt tm} prints the usual textual representation of the 10 | term {tm} to the formatter {fmt}. The string is just {tm} not {`tm`}. 11 | 12 | \FAILURE 13 | Should never fail unless the formatter does. 14 | 15 | \COMMENTS 16 | The usual case where the formatter is the standard output is {print_term}. 17 | 18 | \SEEALSO 19 | pp_print_qterm, print_qterm, print_term. 20 | 21 | \ENDDOC 22 | -------------------------------------------------------------------------------- /Help/SYM_CONV.hlp: -------------------------------------------------------------------------------- 1 | \DOC SYM_CONV 2 | 3 | \TYPE {SYM_CONV : term -> thm} 4 | 5 | \SYNOPSIS 6 | Interchanges the left and right-hand sides of an equation. 7 | 8 | \KEYWORDS 9 | conversion, symmetry, equality. 10 | 11 | \DESCRIBE 12 | When applied to an equational term {t1 = t2}, the conversion 13 | {SYM_CONV} returns the theorem: 14 | { 15 | |- t1 = t2 <=> t2 = t1 16 | } 17 | 18 | \FAILURE 19 | Fails if applied to a term that is not an equation. 20 | 21 | \EXAMPLE 22 | { 23 | # SYM_CONV `2 = x`;; 24 | val it : thm = |- 2 = x <=> x = 2 25 | } 26 | 27 | \SEEALSO 28 | SYM. 29 | 30 | \ENDDOC 31 | -------------------------------------------------------------------------------- /Help/dest_vartype.hlp: -------------------------------------------------------------------------------- 1 | \DOC dest_vartype 2 | 3 | \TYPE {dest_vartype : hol_type -> string} 4 | 5 | \SYNOPSIS 6 | Breaks a type variable down to its name. 7 | 8 | \DESCRIBE 9 | {dest_vartype ":A"} returns {"A"} when {A} is a type variable. 10 | 11 | \FAILURE 12 | Fails with {dest_vartype} if the type is not a type variable. 13 | 14 | \EXAMPLE 15 | { 16 | # dest_vartype `:A`;; 17 | val it : string = "A" 18 | 19 | # dest_vartype `:A->B`;; 20 | Exception: Failure "dest_vartype: type constructor not a variable". 21 | } 22 | \SEEALSO 23 | mk_vartype, is_vartype, dest_type. 24 | 25 | \ENDDOC 26 | -------------------------------------------------------------------------------- /Help/empty_ss.hlp: -------------------------------------------------------------------------------- 1 | \DOC empty_ss 2 | 3 | \TYPE {empty_ss : simpset} 4 | 5 | \SYNOPSIS 6 | Simpset consisting of only the default rewrites and conversions. 7 | 8 | \DESCRIBE 9 | In their maximal generality, simplification operations in HOL Light (as invoked 10 | by {SIMP_TAC}) are controlled by a `simpset'. The simpset {empty_ss} has just 11 | the basic rewrites and conversions (see {basic_rewrites} and {basic_convs}), 12 | and no other provers. 13 | 14 | \FAILURE 15 | Not applicable. 16 | 17 | \SEEALSO 18 | basic_convs, basic_rewrites, basic_ss, SIMP_CONV, SIMP_RULE, SIMP_TAC. 19 | 20 | \ENDDOC 21 | -------------------------------------------------------------------------------- /Help/flush_goalstack.hlp: -------------------------------------------------------------------------------- 1 | \DOC flush_goalstack 2 | 3 | \TYPE {flush_goalstack : unit -> unit} 4 | 5 | \SYNOPSIS 6 | Eliminate all but the current goalstate from the current goalstack. 7 | 8 | \DESCRIBE 9 | Normally, the current goalstack has the current goalstate at the head and all 10 | previous intermediate states further back in the list. This function 11 | {flush_goalstack()} keeps just the current goalstate and eliminates all 12 | previous states. 13 | 14 | \FAILURE 15 | Fails if there is no current goalstate, i.e. if the goalstack is empty. 16 | 17 | \SEEALSO 18 | b, g, r. 19 | 20 | \ENDDOC 21 | -------------------------------------------------------------------------------- /Help/mk_abs.hlp: -------------------------------------------------------------------------------- 1 | \DOC mk_abs 2 | 3 | \TYPE {mk_abs : term * term -> term} 4 | 5 | \SYNOPSIS 6 | Constructs an abstraction. 7 | 8 | \DESCRIBE 9 | If {v} is a variable and {t} any term, then {mk_abs(v,t)} produces the 10 | abstraction term {\v. t}. It is not necessary that {v} should occur free in 11 | {t}. 12 | 13 | \FAILURE 14 | Fails if {v} is not a variable. See {mk_gabs} for constructing generalized 15 | abstraction terms. 16 | 17 | \EXAMPLE 18 | { 19 | # mk_abs(`x:num`,`x + 1`);; 20 | val it : term = `\x. x + 1` 21 | } 22 | 23 | \SEEALSO 24 | dest_abs, is_abs, mk_gabs. 25 | 26 | \ENDDOC 27 | -------------------------------------------------------------------------------- /Help/null_meta.hlp: -------------------------------------------------------------------------------- 1 | \DOC null_meta 2 | 3 | \TYPE {null_meta : term list * instantiation} 4 | 5 | \SYNOPSIS 6 | Empty metavariable information. 7 | 8 | \DESCRIBE 9 | This is a pair consisting of an empty list of terms and a null instantiation 10 | (see {null_inst}). It is used inside most tactics to indicate that they do 11 | nothing interesting with metavariables. 12 | 13 | \FAILURE 14 | Not applicable. 15 | 16 | \COMMENTS 17 | This is not intended for general use, but readers writing custom tactics from 18 | scratch may find it convenient. 19 | 20 | \SEEALSO 21 | null_inst. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Help/parse_type.hlp: -------------------------------------------------------------------------------- 1 | \DOC parse_type 2 | 3 | \TYPE {parse_type : string -> hol_type} 4 | 5 | \SYNOPSIS 6 | Parses a string into a HOL type. 7 | 8 | \DESCRIBE 9 | The call {parse_type "s"} parses the string {s} into a HOL type. This is the 10 | function that is invoked automatically when a type is written in quotations 11 | with an initial colon {`:s`}. 12 | 13 | \FAILURE 14 | Fails in the event of a syntax error or unparsed input. 15 | 16 | \EXAMPLE 17 | { 18 | # parse_type "num->bool";; 19 | val it : hol_type = `:num->bool` 20 | } 21 | 22 | \SEEALSO 23 | lex, parse_term. 24 | 25 | \ENDDOC 26 | -------------------------------------------------------------------------------- /Cadical/test.ml: -------------------------------------------------------------------------------- 1 | (* ========================================================================= *) 2 | (* Test the Cadical interface on a set of (relatively easy) tautologies. *) 3 | (* ========================================================================= *) 4 | 5 | needs "Minisat/taut.ml";; 6 | 7 | let TEST_TAUT TAUTCHECKER p = 8 | let th = time TAUTCHECKER p in 9 | if hyp th = [] && concl th = p 10 | then true else failwith "Wrong theorem";; 11 | 12 | map (fun (p,s) -> print_string("Attempting "^s); print_newline(); 13 | s,TEST_TAUT CADICAL_PROVE p) 14 | all_taut;; 15 | -------------------------------------------------------------------------------- /Help/NO_THEN.hlp: -------------------------------------------------------------------------------- 1 | \DOC NO_THEN 2 | 3 | \TYPE {NO_THEN : thm_tactical} 4 | 5 | \SYNOPSIS 6 | Theorem-tactical which always fails. 7 | 8 | \KEYWORDS 9 | theorem-tactic. 10 | 11 | \DESCRIBE 12 | When applied to a theorem-tactic and a theorem, the theorem-tactical 13 | {NO_THEN} always fails with {Failwith "NO_THEN"}. 14 | 15 | \FAILURE 16 | Always fails when applied to a theorem-tactic and a theorem (note that it 17 | never gets as far as being applied to a goal!) 18 | 19 | \USES 20 | Writing compound tactics or tacticals. 21 | 22 | \SEEALSO 23 | ALL_TAC, ALL_THEN, FAIL_TAC, NO_TAC. 24 | 25 | \ENDDOC 26 | -------------------------------------------------------------------------------- /Help/REFL.hlp: -------------------------------------------------------------------------------- 1 | \DOC REFL 2 | 3 | \TYPE {REFL : term -> thm} 4 | 5 | \SYNOPSIS 6 | Returns theorem expressing reflexivity of equality. 7 | 8 | \KEYWORDS 9 | rule, reflexive, equality. 10 | 11 | \DESCRIBE 12 | {REFL} maps any term {`t`} to the corresponding theorem {|- t = t}. 13 | 14 | \FAILURE 15 | Never fails. 16 | 17 | \EXAMPLE 18 | { 19 | # REFL `2`;; 20 | val it : thm = |- 2 = 2 21 | 22 | # REFL `p:bool`;; 23 | val it : thm = |- p <=> p 24 | } 25 | 26 | \COMMENTS 27 | This is one of HOL Light's 10 primitive inference rules. 28 | 29 | \SEEALSO 30 | ALL_CONV, REFL_TAC. 31 | 32 | \ENDDOC 33 | -------------------------------------------------------------------------------- /Help/REFL_TAC.hlp: -------------------------------------------------------------------------------- 1 | \DOC REFL_TAC 2 | 3 | \TYPE {REFL_TAC : tactic} 4 | 5 | \SYNOPSIS 6 | Solves a goal that is an equation between alpha-equivalent terms. 7 | 8 | \KEYWORDS 9 | tactic, reflexive, equality, alpha. 10 | 11 | \DESCRIBE 12 | When applied to a goal {A ?- t = t'}, where {t} and {t'} are alpha-equivalent, 13 | {REFL_TAC} completely solves it. 14 | { 15 | A ?- t = t' 16 | ============= REFL_TAC 17 | 18 | } 19 | 20 | \FAILURE 21 | Fails unless the goal is an equation between alpha-equivalent terms. 22 | 23 | \SEEALSO 24 | ACCEPT_TAC, MATCH_ACCEPT_TAC, REWRITE_TAC. 25 | 26 | \ENDDOC 27 | -------------------------------------------------------------------------------- /Help/string_of_type.hlp: -------------------------------------------------------------------------------- 1 | \DOC string_of_type 2 | 3 | \TYPE {string_of_type : hol_type -> string} 4 | 5 | \SYNOPSIS 6 | Converts a HOL type to a string representation. 7 | 8 | \DESCRIBE 9 | The call {string_of_type ty} produces a textual representation of the type {ty} 10 | as a string, similar to what is printed automatically at the toplevel, though 11 | without the surrounding quotes and colon. 12 | 13 | \FAILURE 14 | Never fails. 15 | 16 | \EXAMPLE 17 | { 18 | # string_of_type bool_ty;; 19 | val it : string = "bool" 20 | } 21 | 22 | \SEEALSO 23 | string_of_term, string_of_thm. 24 | 25 | \ENDDOC 26 | -------------------------------------------------------------------------------- /Help/unparse_as_prefix.hlp: -------------------------------------------------------------------------------- 1 | \DOC unparse_as_prefix 2 | 3 | \TYPE {unparse_as_prefix : string -> unit} 4 | 5 | \SYNOPSIS 6 | Removes prefix status for an identifier. 7 | 8 | \DESCRIBE 9 | Certain identifiers {c} have prefix status, meaning that combinations of the 10 | form {c f x} will be parsed as {c (f x)} rather than the usual {(c f) x}. The 11 | call {unparse_as_prefix "c"} removes {c} from the list of such identifiers. 12 | 13 | \FAILURE 14 | Never fails, regardless of whether {c} originally did have prefix status. 15 | 16 | \SEEALSO 17 | is_prefix, parse_as_prefix, prefixes. 18 | 19 | \ENDDOC 20 | -------------------------------------------------------------------------------- /Help/warn.hlp: -------------------------------------------------------------------------------- 1 | \DOC warn 2 | 3 | \TYPE {warn : bool -> string -> unit} 4 | 5 | \SYNOPSIS 6 | Prints out a warning string 7 | 8 | \DESCRIBE 9 | When applied to a boolean value {b} and a string {s}, the call {warn b s} 10 | prints out ``{Warning: s}'' and a following newline to the terminal if {b} is 11 | true and otherwise does nothing. 12 | 13 | \FAILURE 14 | Never fails. 15 | 16 | \EXAMPLE 17 | { 18 | # let n = 7;; 19 | val n : int = 7 20 | # warn (n <> 0) "Nonzero value";; 21 | Warning: Nonzero value 22 | val it : unit = () 23 | } 24 | 25 | \SEEALSO 26 | remark, report. 27 | 28 | \ENDDOC 29 | -------------------------------------------------------------------------------- /Help/BETAS_CONV.hlp: -------------------------------------------------------------------------------- 1 | \DOC BETAS_CONV 2 | 3 | \TYPE {BETAS_CONV : conv} 4 | 5 | \SYNOPSIS 6 | Beta conversion over multiple arguments. 7 | 8 | \DESCRIBE 9 | Given a term {t} of the form {`(\x1 ... xn. t[x1,...,xn]) s1 ... sn`}, the call 10 | {BETAS_CONV t} returns 11 | { 12 | |- (\x1 ... xn. t[x1,...,xn]) s1 ... sn = t[s1,...,sn] 13 | } 14 | 15 | \FAILURE 16 | Fails if the term is not of the form shown, for some {n}. 17 | 18 | \EXAMPLE 19 | { 20 | # BETAS_CONV `(\x y. x + y) 1 2`;; 21 | val it : thm = |- (\x y. x + y) 1 2 = 1 + 2 22 | } 23 | 24 | \SEEALSO 25 | BETA_CONV, RIGHT_BETAS. 26 | 27 | \ENDDOC 28 | -------------------------------------------------------------------------------- /Help/dest_const.hlp: -------------------------------------------------------------------------------- 1 | \DOC dest_const 2 | 3 | \TYPE {dest_const : term -> string * hol_type} 4 | 5 | \SYNOPSIS 6 | Breaks apart a constant into name and type. 7 | 8 | \DESCRIBE 9 | {dest_const} is a term destructor for constants: 10 | { 11 | dest_const `const:ty` 12 | } 13 | \noindent returns {("const",`:ty`)}. 14 | 15 | \FAILURE 16 | Fails with {dest_const} if term is not a constant. 17 | 18 | \EXAMPLE 19 | { 20 | # dest_const `T`;; 21 | val it : string * hol_type = ("T", `:bool`) 22 | } 23 | 24 | \SEEALSO 25 | dest_abs, dest_comb, dest_var, is_const, mk_const, mk_mconst, name_of. 26 | 27 | \ENDDOC 28 | -------------------------------------------------------------------------------- /Help/mk_cons.hlp: -------------------------------------------------------------------------------- 1 | \DOC mk_cons 2 | 3 | \TYPE {mk_cons : term -> term -> term} 4 | 5 | \SYNOPSIS 6 | Constructs a {CONS} pair. 7 | 8 | \DESCRIBE 9 | {mk_cons `h` `t`} returns {`CONS h t`}. 10 | 11 | \FAILURE 12 | Fails if second term is not of list type or if the first term is not of the 13 | same type as the elements of the list. 14 | 15 | \EXAMPLE 16 | { 17 | # mk_cons `1` `l:num list`;; 18 | val it : term = `CONS 1 l` 19 | 20 | # mk_cons `1` `[2;3;4]`;; 21 | val it : term = `[1; 2; 3; 4]` 22 | } 23 | 24 | \SEEALSO 25 | dest_cons, dest_list, is_cons, is_list, mk_flist, mk_list. 26 | 27 | \ENDDOC 28 | -------------------------------------------------------------------------------- /Help/pp_print_qtype.hlp: -------------------------------------------------------------------------------- 1 | \DOC pp_print_qtype 2 | 3 | \TYPE {pp_print_qtype : formatter -> hol_type -> unit} 4 | 5 | \SYNOPSIS 6 | Prints a type with initial colon and surrounding quotes to formatter. 7 | 8 | \DESCRIBE 9 | The call {pp_print_type fmt ty} prints the usual textual representation of the 10 | type {ty} to the formatter {fmt}, in the form {`:ty`}. 11 | 12 | \FAILURE 13 | Should never fail unless the formatter does. 14 | 15 | \COMMENTS 16 | The usual case where the formatter is the standard output is {print_qtype}. 17 | 18 | \SEEALSO 19 | pp_print_type, print_qtype, print_type. 20 | 21 | \ENDDOC 22 | -------------------------------------------------------------------------------- /Help/pp_print_type.hlp: -------------------------------------------------------------------------------- 1 | \DOC pp_print_type 2 | 3 | \TYPE {pp_print_type : formatter -> hol_type -> unit} 4 | 5 | \SYNOPSIS 6 | Prints a type (without colon or quotes) to formatter. 7 | 8 | \DESCRIBE 9 | The call {pp_print_type fmt ty} prints the usual textual representation of the 10 | type {ty} to the formatter {fmt}. The string is just {ty} not {`:ty`}. 11 | 12 | \FAILURE 13 | Should never fail unless the formatter does. 14 | 15 | \COMMENTS 16 | The usual case where the formatter is the standard output is {print_type}. 17 | 18 | \SEEALSO 19 | pp_print_qtype, print_qtype, print_type. 20 | 21 | \ENDDOC 22 | -------------------------------------------------------------------------------- /Help/top_realgoal.hlp: -------------------------------------------------------------------------------- 1 | \DOC top_realgoal 2 | 3 | \TYPE {top_realgoal : unit -> (string * thm) list * term} 4 | 5 | \SYNOPSIS 6 | Returns the actual internal structure of the current goal. 7 | 8 | \DESCRIBE 9 | Returns the actual internal representation of the current goal, including the 10 | labels and the theorems that are the assumptions. 11 | 12 | \USES 13 | For users interested in the precise internal structure of the goal, e.g. to 14 | debug subtle free variable problems. Normally the simpler structure returned by 15 | {top_goal} is entirely adequate. 16 | 17 | \SEEALSO 18 | top_goal. 19 | 20 | \ENDDOC 21 | -------------------------------------------------------------------------------- /Help/mk_string.hlp: -------------------------------------------------------------------------------- 1 | \DOC mk_string 2 | 3 | \TYPE {mk_string : string -> term} 4 | 5 | \SYNOPSIS 6 | Constructs object-level string from OCaml string. 7 | 8 | \DESCRIBE 9 | {mk_string "..."} produces the HOL term of type {string} (which is an 10 | abbreviation for {char list}) corresponding to the OCaml string {"..."}. 11 | 12 | \FAILURE 13 | Never fails 14 | 15 | \EXAMPLE 16 | { 17 | # mk_string "hello";; 18 | val it : term = `"hello"` 19 | 20 | # type_of it;; 21 | val it : hol_type = `:(char)list` 22 | } 23 | 24 | \SEEALSO 25 | dest_char, dest_list, dest_string, mk_char, mk_list. 26 | 27 | \ENDDOC 28 | -------------------------------------------------------------------------------- /Help/partition.hlp: -------------------------------------------------------------------------------- 1 | \DOC partition 2 | 3 | \TYPE {partition : ('a -> bool) -> 'a list -> 'a list * 'a list} 4 | 5 | \SYNOPSIS 6 | Separates a list into two lists using a predicate. 7 | 8 | \DESCRIBE 9 | {partition p l} returns a pair of lists. The first list contains the elements 10 | which satisfy {p}. The second list contains all the other elements. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \EXAMPLE 16 | { 17 | # partition (fun x -> x mod 2 = 0) (1--10);; 18 | val it : int list * int list = ([2; 4; 6; 8; 10], [1; 3; 5; 7; 9]) 19 | } 20 | 21 | \SEEALSO 22 | chop_list, remove, filter. 23 | 24 | \ENDDOC 25 | -------------------------------------------------------------------------------- /Geometric_Algebra/make.ml: -------------------------------------------------------------------------------- 1 | (* ========================================================================= *) 2 | (* Load geometric algebra theory plus complex/quaternions application *) 3 | (* *) 4 | (* (c) Copyright, Capital Normal University, China, 2018. *) 5 | (* Authors: Liming Li, Zhiping Shi, Yong Guan, Guohui Wang, Sha Ma. *) 6 | (* ========================================================================= *) 7 | 8 | loadt "Geometric_Algebra/geometricalgebra.ml";; 9 | loadt "Geometric_Algebra/quaternions.ml";; 10 | -------------------------------------------------------------------------------- /Help/EQF_INTRO.hlp: -------------------------------------------------------------------------------- 1 | \DOC EQF_INTRO 2 | 3 | \TYPE {EQF_INTRO : thm -> thm} 4 | 5 | \SYNOPSIS 6 | Converts negation to equality with {F}. 7 | 8 | \KEYWORDS 9 | rule, negation, falsity. 10 | 11 | \DESCRIBE 12 | { 13 | A |- ~tm 14 | --------------- EQF_INTRO 15 | A |- tm <=> F 16 | } 17 | 18 | \FAILURE 19 | Fails if the argument theorem is not a negation. 20 | 21 | \EXAMPLE 22 | { 23 | # let th = ASSUME `~p`;; 24 | val th : thm = ~p |- ~p 25 | 26 | # EQF_INTRO th;; 27 | val it : thm = ~p |- p <=> F 28 | } 29 | 30 | \SEEALSO 31 | EQF_ELIM, EQT_ELIM, EQT_INTRO, NOT_ELIM, NOT_INTRO. 32 | 33 | \ENDDOC 34 | -------------------------------------------------------------------------------- /Help/UNDISCH.hlp: -------------------------------------------------------------------------------- 1 | \DOC UNDISCH 2 | 3 | \TYPE {UNDISCH : thm -> thm} 4 | 5 | \SYNOPSIS 6 | Undischarges the antecedent of an implicative theorem. 7 | 8 | \KEYWORDS 9 | rule, undischarge, antecedent. 10 | 11 | \DESCRIBE 12 | { 13 | A |- t1 ==> t2 14 | ---------------- UNDISCH 15 | A, t1 |- t2 16 | } 17 | 18 | \FAILURE 19 | {UNDISCH} will fail on theorems which are not implications. 20 | 21 | \EXAMPLE 22 | { 23 | # UNDISCH(TAUT `p /\ q ==> p`);; 24 | val it : thm = p /\ q |- p 25 | } 26 | 27 | \SEEALSO 28 | DISCH, DISCH_ALL, DISCH_TAC, DISCH_THEN, STRIP_TAC, UNDISCH_ALL, UNDISCH_TAC. 29 | 30 | \ENDDOC 31 | -------------------------------------------------------------------------------- /Help/equals_thm.hlp: -------------------------------------------------------------------------------- 1 | \DOC equals_thm 2 | 3 | \TYPE {equals_thm : thm -> thm -> bool} 4 | 5 | \SYNOPSIS 6 | Equality test on theorems. 7 | 8 | \DESCRIBE 9 | The call {equals_thm th1 th2} returns {true} if and only if both the 10 | conclusions and assumptions of the two theorems {th1} and {th2} are exactly the 11 | same. The same can be achieved by a simple equality test, but it is better 12 | practice to use this function because it will also work in the proof recording 13 | version of HOL Light (see the {Proofrecording} subdirectory). 14 | 15 | \FAILURE 16 | Never fails. 17 | 18 | \SEEALSO 19 | =?. 20 | 21 | \ENDDOC 22 | -------------------------------------------------------------------------------- /Help/is_vartype.hlp: -------------------------------------------------------------------------------- 1 | \DOC is_vartype 2 | 3 | \TYPE {is_vartype : hol_type -> bool} 4 | 5 | \SYNOPSIS 6 | Tests a type to see if it is a type variable. 7 | 8 | \DESCRIBE 9 | Returns {true} if applied to a type variable. For types that are not type 10 | variables it returns {false}. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \EXAMPLE 16 | { 17 | # is_vartype `:A`;; 18 | val it : bool = true 19 | 20 | # is_vartype `:bool`;; 21 | val it : bool = false 22 | 23 | # is_vartype (mk_vartype "bool");; 24 | val it : bool = true 25 | 26 | } 27 | 28 | \SEEALSO 29 | mk_vartype, dest_vartype. 30 | 31 | \ENDDOC 32 | -------------------------------------------------------------------------------- /Help/mk_finty.hlp: -------------------------------------------------------------------------------- 1 | \DOC mk_finty 2 | 3 | \TYPE {mk_finty :num -> hol_type} 4 | 5 | \SYNOPSIS 6 | Converts an integer to a standard finite type. 7 | 8 | \DESCRIBE 9 | Finite types parsed and printed as numerals are provided, and this operation 10 | when applied to a number gives a type of that size. 11 | 12 | \FAILURE 13 | Fails if the number is not a strictly positive integer. 14 | 15 | \EXAMPLE 16 | Here we use a 6-element type: 17 | { 18 | # mk_finty (Int 6);; 19 | val it : hol_type = `:6` 20 | } 21 | 22 | \SEEALSO 23 | dest_finty, DIMINDEX_CONV, DIMINDEX_TAC, HAS_SIZE_DIMINDEX_RULE, mk_type. 24 | 25 | \ENDDOC 26 | -------------------------------------------------------------------------------- /Help/allpairs.hlp: -------------------------------------------------------------------------------- 1 | \DOC allpairs 2 | 3 | \TYPE {allpairs : ('a -> 'b -> 'c) -> 'a list -> 'b list -> 'c list} 4 | 5 | \SYNOPSIS 6 | Compute list of all results from applying function to pairs from two lists. 7 | 8 | \DESCRIBE 9 | The call {allpairs f [x1;...;xm] [y1;...;yn]} returns the list of results 10 | {[f x1 y1; f x1 y2; ... ; f x1 yn; f x2 y1; ...; f xm yn]}. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \EXAMPLE 16 | { 17 | # allpairs (fun x y -> (x,y)) [1;2;3] [4;5];; 18 | val it : (int * int) list = [(1, 4); (1, 5); (2, 4); (2, 5); (3, 4); (3, 5)] 19 | } 20 | 21 | \SEEALSO 22 | map2, zip. 23 | 24 | \ENDDOC 25 | -------------------------------------------------------------------------------- /Help/it.hlp: -------------------------------------------------------------------------------- 1 | \DOC it 2 | 3 | \TYPE {it : 'a} 4 | 5 | \SYNOPSIS 6 | Binds the value of the last expression evaluated at top level. 7 | 8 | \DESCRIBE 9 | The identifier {it} is bound to the value of the last expression evaluated 10 | at top level. Declarations do not effect the value of {it}. 11 | 12 | \EXAMPLE 13 | { 14 | # 2 + 3;; 15 | val it : int = 5 16 | # let x = 2*3;; 17 | val x : int = 6 18 | # it;; 19 | val it : int = 5 20 | # it + 12;; 21 | val it : int = 17 22 | } 23 | 24 | \USES 25 | Used in evaluating expressions that require the value of the last evaluated 26 | expression. 27 | 28 | \ENDDOC 29 | -------------------------------------------------------------------------------- /Help/setify.hlp: -------------------------------------------------------------------------------- 1 | \DOC setify 2 | 3 | \TYPE {setify : 'a list -> 'a list} 4 | 5 | \SYNOPSIS 6 | Removes repeated elements from a list. Makes a list into a `set'. 7 | 8 | \DESCRIBE 9 | {setify l} removes repeated elements from {l}, leaving the last occurrence of 10 | each duplicate in the list. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \EXAMPLE 16 | { 17 | # setify [1;2;3;1;4;3];; 18 | val it : int list = [1; 2; 3; 4] 19 | } 20 | 21 | \COMMENTS 22 | The current implementation will in fact return a sorted list according to the 23 | basic OCaml polymorphic ordering. 24 | 25 | \SEEALSO 26 | uniq. 27 | 28 | \ENDDOC 29 | -------------------------------------------------------------------------------- /Help/UNDISCH_TAC.hlp: -------------------------------------------------------------------------------- 1 | \DOC UNDISCH_TAC 2 | 3 | \TYPE {UNDISCH_TAC : term -> tactic} 4 | 5 | \SYNOPSIS 6 | Undischarges an assumption. 7 | 8 | \KEYWORDS 9 | tactic, discharge. 10 | 11 | \DESCRIBE 12 | { 13 | A ?- t 14 | ==================== UNDISCH_TAC `v` 15 | A - {{v}} ?- v ==> t 16 | } 17 | 18 | \FAILURE 19 | {UNDISCH_TAC} will fail if {`v`} is not an assumption. 20 | 21 | \COMMENTS 22 | {UNDISCH}arging {`v`} will remove all assumptions that are alpha-equivalent to 23 | {`v`}. 24 | 25 | \SEEALSO 26 | DISCH, DISCH_ALL, DISCH_TAC, DISCH_THEN, STRIP_TAC, UNDISCH, UNDISCH_ALL, 27 | UNDISCH_THEN. 28 | 29 | \ENDDOC 30 | -------------------------------------------------------------------------------- /Help/assoc.hlp: -------------------------------------------------------------------------------- 1 | \DOC assoc 2 | 3 | \TYPE {assoc : 'a -> ('a * 'b) list -> 'b} 4 | 5 | \SYNOPSIS 6 | Searches a list of pairs for a pair whose first component equals a specified 7 | value. 8 | 9 | \KEYWORDS 10 | list. 11 | 12 | \DESCRIBE 13 | {assoc x [(x1,y1);...;(xn,yn)]} returns the first {yi} in the list such 14 | that {xi} equals {x}. 15 | 16 | \FAILURE 17 | Fails if no matching pair is found. This will always be the case if the list 18 | is empty. 19 | 20 | \EXAMPLE 21 | { 22 | # assoc 2 [1,4; 3,2; 2,5; 2,6];; 23 | val it : int = 5 24 | } 25 | \SEEALSO 26 | rev_assoc, find, mem, tryfind, exists, forall. 27 | 28 | \ENDDOC 29 | -------------------------------------------------------------------------------- /Help/dest_string.hlp: -------------------------------------------------------------------------------- 1 | \DOC dest_string 2 | 3 | \TYPE {dest_string : term -> string} 4 | 5 | \SYNOPSIS 6 | Produces OCaml string corresponding to object-level string. 7 | 8 | \DESCRIBE 9 | {dest_string t} where {t} is a literal string in the HOL object logic of type 10 | {string} (which is an abbreviation for {char list}), produces the corresponding 11 | OCaml string. 12 | 13 | \FAILURE 14 | Fails if the term is not a literal term of type {string} 15 | 16 | \EXAMPLE 17 | { 18 | # dest_string `"HOL"`;; 19 | val it : string = "HOL" 20 | } 21 | 22 | \SEEALSO 23 | dest_char, dest_list, mk_char, mk_list, mk_string. 24 | 25 | \ENDDOC 26 | -------------------------------------------------------------------------------- /Help/mk_fun_ty.hlp: -------------------------------------------------------------------------------- 1 | \DOC mk_fun_ty 2 | 3 | \TYPE {mk_fun_ty : hol_type -> hol_type -> hol_type} 4 | 5 | \SYNOPSIS 6 | Construct a function type. 7 | 8 | \DESCRIBE 9 | The call {mk_fun_ty ty1 ty2} gives the function type {ty1->ty2}. This is an 10 | exact synonym of {mk_type("fun",[ty1; ty2])}, but a little more convenient. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \EXAMPLE 16 | { 17 | # mk_fun_ty `:num` `:num`;; 18 | val it : hol_type = `:num->num` 19 | 20 | # itlist mk_fun_ty [`:A`; `:B`; `:C`] `:bool`;; 21 | val it : hol_type = `:A->B->C->bool` 22 | } 23 | 24 | \SEEALSO 25 | dest_type, mk_type. 26 | 27 | \ENDDOC 28 | -------------------------------------------------------------------------------- /Help/strip_abs.hlp: -------------------------------------------------------------------------------- 1 | \DOC strip_abs 2 | 3 | \TYPE {strip_abs : term -> term list * term} 4 | 5 | \SYNOPSIS 6 | Iteratively breaks apart abstractions. 7 | 8 | \DESCRIBE 9 | {strip_abs `\x1 ... xn. t`} returns {([`x1`;...;`xn`],`t`)}. Note that 10 | { 11 | strip_abs(list_mk_abs([`x1`;...;`xn`],`t`)) 12 | } 13 | \noindent will not return {([`x1`;...;`xn`],`t`)} if {t} is an abstraction. 14 | 15 | \FAILURE 16 | Never fails. 17 | 18 | \EXAMPLE 19 | { 20 | # strip_abs `\x y z. x /\ y /\ z`;; 21 | val it : term list * term = ([`x`; `y`; `z`], `x /\ y /\ z`) 22 | } 23 | 24 | \SEEALSO 25 | list_mk_abs, dest_abs. 26 | 27 | \ENDDOC 28 | -------------------------------------------------------------------------------- /Help/subset.hlp: -------------------------------------------------------------------------------- 1 | \DOC subset 2 | 3 | \TYPE {subset : 'a list -> 'a list -> bool} 4 | 5 | \SYNOPSIS 6 | Tests if one list is a subset of another. 7 | 8 | \DESCRIBE 9 | The call {subset l1 l2} returns {true} if every element of {l1} also occurs in 10 | {l2}, regardless of whether an element appears once or more than once in each 11 | list. So when {l1} and {l2} are regarded as sets, this is a subset test. 12 | 13 | \FAILURE 14 | Never fails. 15 | 16 | \EXAMPLE 17 | { 18 | # subset [1;1;2;2] [1;2;3];; 19 | val it : bool = true 20 | } 21 | 22 | \SEEALSO 23 | insert, intersect, set_eq, setify, subtract, union. 24 | 25 | \ENDDOC 26 | -------------------------------------------------------------------------------- /Help/try_user_parser.hlp: -------------------------------------------------------------------------------- 1 | \DOC try_user_parser 2 | 3 | \TYPE {try_user_parser : lexcode list -> preterm * lexcode list} 4 | 5 | \SYNOPSIS 6 | Try all user parsing functions. 7 | 8 | \DESCRIBE 9 | HOL Light allows user parsing functions to be installed, and will try them on 10 | all terms during parsing before the usual parsers. The call 11 | {try_user_parser l} attempts to parse the list of tokens {l} using all the user 12 | parsers, taking the results from whichever one succeeds first. 13 | 14 | \FAILURE 15 | Fails if all user parsers fail. 16 | 17 | \SEEALSO 18 | delete_parser, install_parser, installed_parsers. 19 | 20 | \ENDDOC 21 | -------------------------------------------------------------------------------- /Help/tryfind.hlp: -------------------------------------------------------------------------------- 1 | \DOC tryfind 2 | 3 | \TYPE {tryfind : ('a -> 'b) -> 'a list -> 'b} 4 | 5 | \SYNOPSIS 6 | Returns the result of the first successful application of a function to the 7 | elements of a list. 8 | 9 | \KEYWORDS 10 | list. 11 | 12 | \DESCRIBE 13 | {tryfind f [x1;...;xn]} returns {(f xi)} for the first {xi} in the list for 14 | which application of {f} succeeds. 15 | 16 | \FAILURE 17 | Fails with {tryfind} if the application of the function fails for all elements 18 | in the list. This will always be the case if the list is empty. 19 | 20 | \SEEALSO 21 | find, mem, exists, forall, assoc, rev_assoc. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Help/is_type.hlp: -------------------------------------------------------------------------------- 1 | \DOC is_type 2 | 3 | \TYPE {is_type : hol_type -> bool} 4 | 5 | \SYNOPSIS 6 | Tests whether a type is an instance of a type constructor. 7 | 8 | \DESCRIBE 9 | {is_type ty} returns {true} if {ty} is a base type or constructed by an outer 10 | type constructor, and {false} if it is a type variable. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \EXAMPLE 16 | { 17 | # is_type `:bool`;; 18 | val it : bool = true 19 | 20 | # is_type `:bool->int`;; 21 | val it : bool = true 22 | 23 | # is_type `:Tyvar`;; 24 | val it : bool = false 25 | } 26 | 27 | \SEEALSO 28 | get_type_arity, is_vartype. 29 | 30 | \ENDDOC 31 | -------------------------------------------------------------------------------- /Help/p.hlp: -------------------------------------------------------------------------------- 1 | \DOC p 2 | 3 | \TYPE {p : unit -> goalstack} 4 | 5 | \SYNOPSIS 6 | Prints the top level of the subgoal package goal stack. 7 | 8 | \DESCRIBE 9 | The function {p} is part of the subgoal package, and prints the current 10 | goalstate. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \USES 16 | Examining the proof state during an interactive proof session. 17 | 18 | \COMMENTS 19 | Strictly speaking this function is side-effect-free. It simply {{\em returns}} 20 | the current goalstate. However, automatic printing will normally then print it, 21 | so that is the net effect. 22 | 23 | \SEEALSO 24 | b, e, g, r. 25 | 26 | \ENDDOC 27 | -------------------------------------------------------------------------------- /Help/print_term.hlp: -------------------------------------------------------------------------------- 1 | \DOC print_term 2 | 3 | \TYPE {print_term : term -> unit} 4 | 5 | \SYNOPSIS 6 | Prints a HOL term (without quotes) to the standard output. 7 | 8 | \DESCRIBE 9 | The call {print_term tm} prints the usual textual representation of the 10 | term {tm} to the standard output. The string is just {tm} not {`tm`}. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \USES 16 | Producing debugging output in complex rules. Note that terms are already 17 | printed at the toplevel anyway, so it is not needed to examine results 18 | interactively. 19 | 20 | \SEEALSO 21 | pp_print_qterm, pp_print_term, print_qterm. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Help/reserve_words.hlp: -------------------------------------------------------------------------------- 1 | \DOC reserve_words 2 | 3 | \TYPE {reserve_words : string list -> unit} 4 | 5 | \SYNOPSIS 6 | Add given strings to the set of reserved words. 7 | 8 | \DESCRIBE 9 | Certain identifiers in HOL are reserved, e.g. `{if}', `{let}' and `{|}', 10 | meaning that they are special to the parser and cannot be used as ordinary 11 | identifiers. A call {reserve_words l} adds all strings in {l} to the list of 12 | reserved identifiers. 13 | 14 | \FAILURE 15 | Never fails, regardless of whether the given strings were already reserved. 16 | 17 | \SEEALSO 18 | is_reserved_word, reserved_words, unreserve_words. 19 | 20 | \ENDDOC 21 | -------------------------------------------------------------------------------- /Help/the_interface.hlp: -------------------------------------------------------------------------------- 1 | \DOC the_interface 2 | 3 | \TYPE {the_interface : (string * (string * hol_type)) list ref} 4 | 5 | \SYNOPSIS 6 | List of active interface mappings. 7 | 8 | \DESCRIBE 9 | HOL Light allows the same identifier to map to one or more underlying constants 10 | using an overloading mechanism with resolution based on type. The reference 11 | variable {the_interface} stores the current list of all interface mappings. 12 | 13 | \SEEALSO 14 | make_overloadable, overload_interface, override_interface, prioritize_overload, 15 | reduce_interface, remove_interface, the_implicit_types, the_overload_skeletons. 16 | 17 | \ENDDOC 18 | -------------------------------------------------------------------------------- /Help/loaded_files.hlp: -------------------------------------------------------------------------------- 1 | \DOC loaded_files 2 | 3 | \TYPE {loaded_files : (string * Digest.t) list ref} 4 | 5 | \SYNOPSIS 6 | List of files loaded so far. 7 | 8 | \DESCRIBE 9 | This reference variable stores a list of previously loaded files together with 10 | MD5 digests. It is updated by all the main loading functions {load_on_path}, 11 | {loads}, {loadt} and {needs}, and is used by {needs} to avoid reloading the 12 | same file multiple times. 13 | 14 | \FAILURE 15 | Not applicable. 16 | 17 | \USES 18 | Not really intended for average users to examine or modify. 19 | 20 | \SEEALSO 21 | load_on_path, loads, loadt, needs. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Help/print_type.hlp: -------------------------------------------------------------------------------- 1 | \DOC print_type 2 | 3 | \TYPE {print_type : hol_type -> unit} 4 | 5 | \SYNOPSIS 6 | Prints a type (without colon or quotes) to standard output. 7 | 8 | \DESCRIBE 9 | The call {print_type ty} prints the usual textual representation of the 10 | type {ty} to the standard output. The string is just {ty} not {`:ty`}. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \USES 16 | Producing debugging output in complex rules. Note that terms are already 17 | printed at the toplevel anyway, so it is not needed to examine results 18 | interactively. 19 | 20 | \SEEALSO 21 | pp_print_qtype, pp_print_type, print_qtype. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Tutorial/Wellfounded_induction.ml: -------------------------------------------------------------------------------- 1 | let NSQRT_2 = prove 2 | (`!p q. p * p = 2 * q * q ==> q = 0`, 3 | MATCH_MP_TAC num_WF THEN REWRITE_TAC[RIGHT_IMP_FORALL_THM] THEN 4 | REPEAT STRIP_TAC THEN FIRST_ASSUM(MP_TAC o AP_TERM `EVEN`) THEN 5 | REWRITE_TAC[EVEN_MULT; ARITH] THEN REWRITE_TAC[EVEN_EXISTS] THEN 6 | DISCH_THEN(X_CHOOSE_THEN `m:num` SUBST_ALL_TAC) THEN 7 | FIRST_X_ASSUM(MP_TAC o SPECL [`q:num`; `m:num`]) THEN 8 | ASM_REWRITE_TAC[ARITH_RULE 9 | `q < 2 * m ==> q * q = 2 * m * m ==> m = 0 <=> 10 | (2 * m) * 2 * m = 2 * q * q ==> 2 * m <= q`] THEN 11 | ASM_MESON_TAC[LE_MULT2; MULT_EQ_0; ARITH_RULE `2 * x <= x <=> x = 0`]);; 12 | -------------------------------------------------------------------------------- /Help/itlist.hlp: -------------------------------------------------------------------------------- 1 | \DOC itlist 2 | 3 | \TYPE {itlist : ('a -> 'b -> 'b) -> 'a list -> 'b -> 'b} 4 | 5 | \SYNOPSIS 6 | List iteration function. Applies a binary function between adjacent elements 7 | of a list. 8 | 9 | \KEYWORDS 10 | list. 11 | 12 | \DESCRIBE 13 | {itlist f [x1;...;xn] y} returns 14 | { 15 | f x1 (f x2 ... (f xn y)...) 16 | } 17 | \noindent It returns {y} if list is empty. 18 | 19 | \FAILURE 20 | Never fails. 21 | 22 | \EXAMPLE 23 | { 24 | # itlist (+) [1;2;3;4;5] 0;; 25 | val it : int = 15 26 | # itlist (+) [1;2;3;4;5] 6;; 27 | val it : int = 21 28 | } 29 | 30 | \SEEALSO 31 | rev_itlist, end_itlist. 32 | 33 | \ENDDOC 34 | -------------------------------------------------------------------------------- /Help/list_mk_gabs.hlp: -------------------------------------------------------------------------------- 1 | \DOC list_mk_gabs 2 | 3 | \TYPE {list_mk_gabs : term list * term -> term} 4 | 5 | \SYNOPSIS 6 | Iteratively makes a generalized abstraction. 7 | 8 | \DESCRIBE 9 | The call {list_mk_gabs([vs1; ...; vsn],t)} constructs an interated generalized 10 | abstraction {\vs1. \vs2. ... \vsn. t}. See {mk_gabs} for more details on 11 | constructing generalized abstractions. 12 | 13 | \FAILURE 14 | Never fails. 15 | 16 | \EXAMPLE 17 | { 18 | # list_mk_gabs([`(x:num,y:num)`; `(w:num,z:num)`],`x + w + 1`);; 19 | val it : term = `\(x,y). \(w,z). x + w + 1` 20 | } 21 | 22 | \SEEALSO 23 | dest_gabs, is_gabs, mk_gabs. 24 | 25 | \ENDDOC 26 | -------------------------------------------------------------------------------- /Help/mk_iff.hlp: -------------------------------------------------------------------------------- 1 | \DOC mk_iff 2 | 3 | \TYPE {mk_iff : term * term -> term} 4 | 5 | \SYNOPSIS 6 | Constructs a logical equivalence (Boolean equation). 7 | 8 | \DESCRIBE 9 | {mk_iff(`t1`,`t2`)} returns {`t1 <=> t2`}. 10 | 11 | \FAILURE 12 | Fails with unless {t1} and {t2} both have Boolean type. 13 | 14 | \EXAMPLE 15 | { 16 | # mk_iff(`x = 1`,`1 = x`);; 17 | val it : term = `x = 1 <=> 1 = x` 18 | } 19 | 20 | \COMMENTS 21 | Simply {mk_eq} has the same effect on successful calls. However {mk_iff} is 22 | slightly more efficient, and will fail if the terms do not have Boolean type. 23 | 24 | \SEEALSO 25 | dest_iff, is_iff,mk_eq. 26 | 27 | \ENDDOC 28 | -------------------------------------------------------------------------------- /Help/CONTRAPOS_CONV.hlp: -------------------------------------------------------------------------------- 1 | \DOC CONTRAPOS_CONV 2 | 3 | \TYPE {CONTRAPOS_CONV : term -> thm} 4 | 5 | \SYNOPSIS 6 | Proves the equivalence of an implication and its contrapositive. 7 | 8 | \KEYWORDS 9 | conversion, contrapositive, implication. 10 | 11 | \DESCRIBE 12 | When applied to an implication {`p ==> q`}, the conversion {CONTRAPOS_CONV} 13 | returns the theorem: 14 | { 15 | |- (p ==> q) <=> (~q ==> ~p) 16 | } 17 | 18 | \FAILURE 19 | Fails if applied to a term that is not an implication. 20 | 21 | \COMMENTS 22 | The same effect can be had by {GEN_REWRITE_CONV I [GSYM CONTRAPOS_THM]} 23 | 24 | \SEEALSO 25 | CCONTR, CONTR_TAC. 26 | 27 | \ENDDOC 28 | -------------------------------------------------------------------------------- /Help/install_parser.hlp: -------------------------------------------------------------------------------- 1 | \DOC install_parser 2 | 3 | \TYPE {install_parser : string * (lexcode list -> preterm * lexcode list) -> unit} 4 | 5 | \SYNOPSIS 6 | Install a user parser. 7 | 8 | \DESCRIBE 9 | HOL Light allows user parsing functions to be installed, and will try them on 10 | all terms during parsing before the usual parsers. The call 11 | {install_parser(s,p)} installs the parser {p} among the user parsers to try in 12 | this way. The string {s} is there so that the parser can conveniently be 13 | deleted again. 14 | 15 | \FAILURE 16 | Never fails. 17 | 18 | \SEEALSO 19 | delete_parser, installed_parsers, try_user_parser. 20 | 21 | \ENDDOC 22 | -------------------------------------------------------------------------------- /Help/pp_print_num.hlp: -------------------------------------------------------------------------------- 1 | \DOC pp_print_num 2 | 3 | \TYPE {pp_print_num : Format.formatter -> num -> unit} 4 | 5 | \SYNOPSIS 6 | Print an arbitrary-precision number to the given formatter. 7 | 8 | \DESCRIBE 9 | This function prints an arbitrary-precision (type {num}) number to the 10 | formatter given as first argument. It is automatically invoked on anything of 11 | type {num} at the toplevel anyway, but it may sometimes be useful to issue it 12 | under user control. 13 | 14 | \FAILURE 15 | Never fails. 16 | 17 | \COMMENTS 18 | The usual case where the formatter is the standard output is {print_num}. 19 | 20 | \SEEALSO 21 | print_num. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Proofrecording/diffs/proofobjects_init.ml: -------------------------------------------------------------------------------- 1 | let (use_proofobjects, use_extended_proofobjects, use_coq) = 2 | try 3 | let n = Sys.getenv "HOLPROOFOBJECTS" in 4 | if n = "BASIC" then 5 | (true, false, false) 6 | else if n = "EXTENDED" then 7 | (true, true, false) 8 | else if n = "COQ" then 9 | (true, true, true) 10 | else 11 | (false, false, false) 12 | with Not_found -> (false, false, false);; 13 | 14 | let _ = 15 | if use_proofobjects then 16 | if use_coq then 17 | loads "proofobjects_coq.ml" 18 | else 19 | loads "proofobjects_trt.ml" 20 | else 21 | loads "proofobjects_dummy.ml";; 22 | -------------------------------------------------------------------------------- /Help/ASSUME.hlp: -------------------------------------------------------------------------------- 1 | \DOC ASSUME 2 | 3 | \TYPE {ASSUME : term -> thm} 4 | 5 | \SYNOPSIS 6 | Introduces an assumption. 7 | 8 | \KEYWORDS 9 | rule, assumption. 10 | 11 | \DESCRIBE 12 | When applied to a term {t}, which must have type {bool}, the inference rule 13 | {ASSUME} returns the theorem {t |- t}. 14 | { 15 | -------- ASSUME `t` 16 | t |- t 17 | } 18 | \FAILURE 19 | Fails unless the term {t} has type {bool}. 20 | 21 | \EXAMPLE 22 | { 23 | # ASSUME `p /\ q`;; 24 | val it : thm = p /\ q |- p /\ q 25 | } 26 | 27 | \COMMENTS 28 | This is one of HOL Light's 10 primitive inference rules. 29 | 30 | \SEEALSO 31 | ADD_ASSUM, REFL. 32 | 33 | \ENDDOC 34 | -------------------------------------------------------------------------------- /Help/CONJ_PAIR.hlp: -------------------------------------------------------------------------------- 1 | \DOC CONJ_PAIR 2 | 3 | \TYPE {CONJ_PAIR : thm -> thm * thm} 4 | 5 | \SYNOPSIS 6 | Extracts both conjuncts of a conjunction. 7 | 8 | \KEYWORDS 9 | rule, conjunction. 10 | 11 | \DESCRIBE 12 | { 13 | A |- t1 /\ t2 14 | ---------------------- CONJ_PAIR 15 | A |- t1 A |- t2 16 | } 17 | \noindent The two resultant theorems are returned as a pair. 18 | 19 | \FAILURE 20 | Fails if the input theorem is not a conjunction. 21 | 22 | \EXAMPLE 23 | { 24 | # CONJ_PAIR(ASSUME `p /\ q`);; 25 | val it : thm * thm = (p /\ q |- p, p /\ q |- q) 26 | } 27 | 28 | \SEEALSO 29 | CONJUNCT1, CONJUNCT2, CONJ, CONJUNCTS. 30 | 31 | \ENDDOC 32 | -------------------------------------------------------------------------------- /Help/refine.hlp: -------------------------------------------------------------------------------- 1 | \DOC refine 2 | 3 | \TYPE {refine : refinement -> goalstack} 4 | 5 | \SYNOPSIS 6 | Applies a refinement to the current goalstack. 7 | 8 | \DESCRIBE 9 | The call {refine r} applies the refinement {r} to the current goalstate, adding 10 | the resulting goalstate at the head of the current goalstack list. (A goalstate 11 | consists of a list of subgoals as well as justification and metavariable 12 | information.) 13 | 14 | \FAILURE 15 | Fails if the refinement fails. 16 | 17 | \COMMENTS 18 | Most users will not want to handle refinements explicitly. Usually one just 19 | applies a tactic to the first goal in a goalstate. 20 | 21 | \ENDDOC 22 | -------------------------------------------------------------------------------- /Help/null_inst.hlp: -------------------------------------------------------------------------------- 1 | \DOC null_inst 2 | 3 | \TYPE {null_inst : instantiation} 4 | 5 | \SYNOPSIS 6 | Empty instantiation. 7 | 8 | \DESCRIBE 9 | Several functions use objects of type {instantiation}, consisting of type and 10 | term instantiations and higher-order matching information. This instantiation 11 | {null_inst} is the trivial instantiation that does nothing. 12 | 13 | \FAILURE 14 | Not applicable. 15 | 16 | \EXAMPLE 17 | Instantiating a term with it has no effect: 18 | { 19 | # instantiate null_inst `x + 1 = 2`;; 20 | val it : term = `x + 1 = 2` 21 | } 22 | 23 | \SEEALSO 24 | instantiate, INSTANTIATE, INSTANTIATE_ALL, term_match. 25 | 26 | \ENDDOC 27 | -------------------------------------------------------------------------------- /Help/rev_assoc.hlp: -------------------------------------------------------------------------------- 1 | \DOC rev_assoc 2 | 3 | \TYPE {rev_assoc : 'a -> ('b * 'a) list -> 'b} 4 | 5 | \SYNOPSIS 6 | Searches a list of pairs for a pair whose second component equals a specified 7 | value. 8 | 9 | \KEYWORDS 10 | list. 11 | 12 | \DESCRIBE 13 | {rev_assoc y [(x1,y1);...;(xn,yn)]} returns the first {xi} in the list 14 | such that {yi} equals {y}. 15 | 16 | \FAILURE 17 | Fails if no matching pair is found. This will always be the case if the list 18 | is empty. 19 | 20 | \EXAMPLE 21 | { 22 | # rev_assoc 2 [(1,4);(3,2);(2,5);(2,6)];; 23 | val it : int = 3 24 | } 25 | 26 | \SEEALSO 27 | assoc, find, mem, tryfind, exists, forall. 28 | 29 | \ENDDOC 30 | -------------------------------------------------------------------------------- /Help/equals_goal.hlp: -------------------------------------------------------------------------------- 1 | \DOC equals_goal 2 | 3 | \TYPE {equals_goal : goal -> goal -> bool} 4 | 5 | \SYNOPSIS 6 | Equality test on goals. 7 | 8 | \DESCRIBE 9 | The relation {equals_goal} tests if two goals have exactly the same structure, 10 | with the same assumptions, conclusions and even labels, with the assumptions in 11 | the same order. The only respect in which this differs from a pure equality 12 | tests is that the various term components are tested modulo alpha-conversion. 13 | 14 | \FAILURE 15 | Never fails. 16 | 17 | \COMMENTS 18 | Probably not generally useful. Used inside {CHANGED_TAC}. 19 | 20 | \SEEALSO 21 | CHANGED_TAC, equals_thm. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Help/is_eq.hlp: -------------------------------------------------------------------------------- 1 | \DOC is_eq 2 | 3 | \TYPE {is_eq : term -> bool} 4 | 5 | \SYNOPSIS 6 | Tests a term to see if it is an equation. 7 | 8 | \DESCRIBE 9 | {is_eq `t1 = t2`} returns {true}. If the term is not an equation the result 10 | is {false}. Note that logical equivalence is just equality on type {:bool}, 11 | even though it is printed as {<=>}. 12 | 13 | \FAILURE 14 | Never fails. 15 | 16 | \EXAMPLE 17 | { 18 | # is_eq `2 + 2 = 4`;; 19 | val it : bool = true 20 | 21 | # is_eq `p /\ q <=> q /\ p`;; 22 | val it : bool = true 23 | 24 | # is_eq `p ==> p`;; 25 | val it : bool = false 26 | } 27 | 28 | \SEEALSO 29 | dest_eq, is_beq, mk_eq. 30 | 31 | \ENDDOC 32 | -------------------------------------------------------------------------------- /Help/REDEPTH_SQCONV.hlp: -------------------------------------------------------------------------------- 1 | \DOC REDEPTH_SQCONV 2 | 3 | \TYPE {REDEPTH_SQCONV : strategy} 4 | 5 | \SYNOPSIS 6 | Applies simplification bottom-up to all subterms, retraversing changed ones. 7 | 8 | \DESCRIBE 9 | HOL Light's simplification functions (e.g. {SIMP_TAC}) have their traversal 10 | algorithm controlled by a ``strategy''. {REDEPTH_SQCONV} is a strategy 11 | corresponding to {REDEPTH_CONV} for ordinary conversions: simplification is 12 | applied bottom-up to all subterms, retraversing changed ones. 13 | 14 | \FAILURE 15 | Not applicable. 16 | 17 | \SEEALSO 18 | DEPTH_SQCONV, ONCE_DEPTH_SQCONV, REDEPTH_CONV, TOP_DEPTH_SQCONV, 19 | TOP_SWEEP_SQCONV. 20 | 21 | \ENDDOC 22 | -------------------------------------------------------------------------------- /Help/TRY_CONV.hlp: -------------------------------------------------------------------------------- 1 | \DOC TRY_CONV 2 | 3 | \TYPE {TRY_CONV : conv -> conv} 4 | 5 | \SYNOPSIS 6 | Attempts to apply a conversion; applies identity conversion in case of failure. 7 | 8 | \KEYWORDS 9 | conversion, failure. 10 | 11 | \DESCRIBE 12 | {TRY_CONV c `t`} attempts to apply the conversion {c} to the term {`t`}; if 13 | this fails, then the identity conversion is applied instead giving the 14 | reflexive theorem {|- t = t}. 15 | 16 | \FAILURE 17 | Never fails. 18 | 19 | \EXAMPLE 20 | { 21 | # num_CONV `0`;; 22 | Exception: Failure "num_CONV". 23 | # TRY_CONV num_CONV `0`;; 24 | val it : thm = |- 0 = 0 25 | } 26 | 27 | \SEEALSO 28 | ALL_CONV. 29 | 30 | \ENDDOC 31 | -------------------------------------------------------------------------------- /Help/implode.hlp: -------------------------------------------------------------------------------- 1 | \DOC implode 2 | 3 | \TYPE {implode : string list -> string} 4 | 5 | \SYNOPSIS 6 | Concatenates a list of strings into one string. 7 | 8 | \DESCRIBE 9 | {implode [s1;...;sn]} returns the string formed by concatenating the 10 | strings {s1 ... sn}. If {n} is zero (the list is empty), then the empty string 11 | is returned. 12 | 13 | \FAILURE 14 | Never fails; accepts empty or multi-character component strings. 15 | 16 | \EXAMPLE 17 | { 18 | # implode ["e";"x";"a";"m";"p";"l";"e"];; 19 | val it : string = "example" 20 | # implode ["ex";"a";"mpl";"";"e"];; 21 | val it : string = "example" 22 | } 23 | 24 | \SEEALSO 25 | explode. 26 | 27 | \ENDDOC 28 | -------------------------------------------------------------------------------- /Help/insert.hlp: -------------------------------------------------------------------------------- 1 | \DOC insert 2 | 3 | \TYPE {insert : 'a -> 'a list -> 'a list} 4 | 5 | \SYNOPSIS 6 | Adds element to the head of a list if not already present. 7 | 8 | \DESCRIBE 9 | The call {insert x l} returns just {l} if {x} is already in the list, and 10 | otherwise returns {x::l}. 11 | 12 | \EXAMPLE 13 | { 14 | # insert 5 (1--10);; 15 | val it : int list = [1; 2; 3; 4; 5; 6; 7; 8; 9; 10] 16 | # insert 15 (1--10);; 17 | val it : int list = [15; 1; 2; 3; 4; 5; 6; 7; 8; 9; 10] 18 | } 19 | 20 | \USES 21 | An analog to the basic list constructor {::} but treating the list more like a 22 | set. 23 | 24 | \SEEALSO 25 | union, intersect, subtract. 26 | 27 | \ENDDOC 28 | -------------------------------------------------------------------------------- /Help/lcm_num.hlp: -------------------------------------------------------------------------------- 1 | \DOC lcm_num 2 | 3 | \TYPE {lcm_num : num -> num -> num} 4 | 5 | \SYNOPSIS 6 | Computes lowest common multiple of two unlimited-precision integers. 7 | 8 | \DESCRIBE 9 | The call {lcm_num m n} for two unlimited-precision (type {num}) integers {m} and 10 | {n} returns the (positive) lowest common multiple of {m} and {n}. If either {m} 11 | or {n} (or both) are both zero, it returns zero. 12 | 13 | \FAILURE 14 | Fails if either number is not an integer (the type {num} supports arbitrary 15 | rationals). 16 | 17 | \EXAMPLE 18 | { 19 | # lcm_num (Int 35) (Int(-77));; 20 | val it : num = 385 21 | } 22 | 23 | \SEEALSO 24 | gcd, gcd_num. 25 | 26 | \ENDDOC 27 | -------------------------------------------------------------------------------- /Help/set_eq.hlp: -------------------------------------------------------------------------------- 1 | \DOC set_eq 2 | 3 | \TYPE {set_eq : 'a list -> 'a list -> bool} 4 | 5 | \SYNOPSIS 6 | Tests two `sets' for equality. 7 | 8 | \DESCRIBE 9 | {set_eq l1 l2} returns {true} if every element of {l1} appears in {l2} and 10 | every element of {l2} appears in {l1}. Otherwise it returns {false}. In other 11 | words, it tests if the lists are the same considered as sets, i.e. ignoring 12 | duplicates. 13 | 14 | \FAILURE 15 | Never fails. 16 | 17 | \EXAMPLE 18 | { 19 | # set_eq [1;2] [2;1;2];; 20 | val it : bool = true 21 | # set_eq [1;2] [1;3];; 22 | val it : bool = false 23 | } 24 | 25 | \SEEALSO 26 | setify, union, intersect, subtract. 27 | 28 | \ENDDOC 29 | -------------------------------------------------------------------------------- /Help/AP_TERM_TAC.hlp: -------------------------------------------------------------------------------- 1 | \DOC AP_TERM_TAC 2 | 3 | \TYPE {AP_TERM_TAC : tactic} 4 | 5 | \SYNOPSIS 6 | Strips a function application from both sides of an equational goal. 7 | 8 | \KEYWORDS 9 | tactic. 10 | 11 | \DESCRIBE 12 | {AP_TERM_TAC} reduces a goal of the form {A ?- f x = f y} by stripping away 13 | the function applications, giving the new goal {A ?- x = y}. 14 | { 15 | A ?- f x = f y 16 | ================ AP_TERM_TAC 17 | A ?- x = y 18 | } 19 | 20 | \FAILURE 21 | Fails unless the goal is equational, with both sides being applications 22 | of the same function. 23 | 24 | \SEEALSO 25 | ABS_TAC, AP_TERM, AP_THM_TAC, BINOP_TAC, MK_COMB_TAC. 26 | 27 | \ENDDOC 28 | -------------------------------------------------------------------------------- /Help/remove_interface.hlp: -------------------------------------------------------------------------------- 1 | \DOC remove_interface 2 | 3 | \TYPE {remove_interface : string -> unit} 4 | 5 | \SYNOPSIS 6 | Remove all overload/interface mappings for an identifier. 7 | 8 | \DESCRIBE 9 | HOL Light allows an identifier to map to a specific constant (see 10 | {override_interface}) or be overloaded to several depending on type (see 11 | {overload_interface}). A call to {remove_interface "ident"} removes all such 12 | mappings for the identifier {ident}. 13 | 14 | \FAILURE 15 | Never fails, whether or not there were any interface mappings in effect. 16 | 17 | \SEEALSO 18 | overload_interface, override_interface, reduce_interface, the_interface. 19 | 20 | \ENDDOC 21 | -------------------------------------------------------------------------------- /miz3/Samples/wishes.ml: -------------------------------------------------------------------------------- 1 | let EXAMPLE = prove 2 | (`!n. nsum(0..n) (\i. i) = (n*(n + 1)) DIV 2`, 3 | INDUCT_TAC THEN ASM_REWRITE_TAC[NSUM_CLAUSES_NUMSEG] THEN ARITH_TAC);; 4 | 5 | let EXAMPLE_IN_MIZAR_LIGHT = thm `; 6 | !n. nsum (0..n) (\i. i) = (n * (n + 1)) DIV 2 [1] 7 | proof 8 | nsum (0..0) (\i. i) = 0 [2] by NSUM_CLAUSES_NUMSEG; 9 | ... = (0 * (0 + 1)) DIV 2 [3] by ARITH_TAC; 10 | !n. nsum (0..n) (\i. i) = (n * (n + 1)) DIV 2 11 | ==> nsum (0..SUC n) (\i. i) = (SUC n * (SUC n + 1)) DIV 2 [4] 12 | proof 13 | let n be num; 14 | assume nsum (0..n) (\i. i) = (n * (n + 1)) DIV 2 [5]; 15 | qed by #; 16 | qed by INDUCT_TAC from 3,4`;; 17 | -------------------------------------------------------------------------------- /Help/CHEAT_TAC.hlp: -------------------------------------------------------------------------------- 1 | \DOC CHEAT_TAC 2 | 3 | \TYPE {CHEAT_TAC : tactic} 4 | 5 | \SYNOPSIS 6 | Proves goal by asserting it as an axiom. 7 | 8 | \DESCRIBE 9 | Given any goal {A ?- p}, the tactic {CHEAT_TAC} solves it by using {mk_thm}, 10 | which in turn involves essentially asserting the goal as a new axiom. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \USES 16 | Temporarily plugging boring parts of a proof to deal with the interesting 17 | parts. 18 | 19 | \COMMENTS 20 | Needless to say, this should be used with caution since once new axioms are 21 | asserted there is no guarantee that logical consistency is preserved. 22 | 23 | \SEEALSO 24 | new_axiom, mk_thm. 25 | 26 | \ENDDOC 27 | -------------------------------------------------------------------------------- /Help/LAND_CONV.hlp: -------------------------------------------------------------------------------- 1 | \DOC LAND_CONV 2 | 3 | \TYPE {LAND_CONV : conv -> conv} 4 | 5 | \SYNOPSIS 6 | Apply a conversion to left-hand argument of binary operator. 7 | 8 | \DESCRIBE 9 | If {c} is a conversion where {c `l`} gives {|- l = l'}, then 10 | {LAND_CONV c `op l r`} gives {|- op l r = op l' r}. 11 | 12 | \FAILURE 13 | Fails if the underlying conversion does or returns an inappropriate theorem 14 | (i.e. is not really a conversion). 15 | 16 | \EXAMPLE 17 | { 18 | # LAND_CONV NUM_ADD_CONV `(2 + 2) + (2 + 2)`;; 19 | val it : thm = |- (2 + 2) + 2 + 2 = 4 + 2 + 2 20 | } 21 | 22 | \SEEALSO 23 | ABS_CONV, COMB_CONV, COMB_CONV2, RAND_CONV, RATOR_CONV, SUB_CONV. 24 | 25 | \ENDDOC 26 | -------------------------------------------------------------------------------- /Help/TOP_DEPTH_SQCONV.hlp: -------------------------------------------------------------------------------- 1 | \DOC TOP_DEPTH_SQCONV 2 | 3 | \TYPE {TOP_DEPTH_SQCONV : strategy} 4 | 5 | \SYNOPSIS 6 | Applies simplification top-down to all subterms, retraversing changed ones. 7 | 8 | \DESCRIBE 9 | HOL Light's simplification functions (e.g. {SIMP_TAC}) have their traversal 10 | algorithm controlled by a ``strategy''. {TOP_DEPTH_SQCONV} is a strategy 11 | corresponding to {TOP_DEPTH_CONV} for ordinary conversions: simplification is 12 | applied top-down to all subterms, retraversing changed ones. 13 | 14 | \FAILURE 15 | Not applicable. 16 | 17 | \SEEALSO 18 | DEPTH_SQCONV, ONCE_DEPTH_SQCONV, REDEPTH_SQCONV, TOP_DEPTH_CONV, 19 | TOP_SWEEP_SQCONV. 20 | 21 | \ENDDOC 22 | -------------------------------------------------------------------------------- /Help/USE_THEN.hlp: -------------------------------------------------------------------------------- 1 | \DOC USE_THEN 2 | 3 | \TYPE {USE_THEN : string -> thm_tactic -> tactic} 4 | 5 | \SYNOPSIS 6 | Apply a theorem tactic to named assumption. 7 | 8 | \DESCRIBE 9 | The tactic {USE_THEN "name" ttac} applies the theorem-tactic {ttac} to the 10 | assumption labelled {name} (or the first in the list if there is more than 11 | one). 12 | 13 | \FAILURE 14 | Fails if there is no assumption of that name or if the theorem-tactic fails 15 | when applied to it. 16 | 17 | \EXAMPLE 18 | See {LABEL_TAC} for an extended example. 19 | 20 | \USES 21 | Using an assumption identified by name. 22 | 23 | \SEEALSO 24 | ASSUME, FIND_ASSUM, HYP, LABEL_TAC, REMOVE_THEN. 25 | 26 | \ENDDOC 27 | -------------------------------------------------------------------------------- /Help/frees.hlp: -------------------------------------------------------------------------------- 1 | \DOC frees 2 | 3 | \TYPE {frees : term -> term list} 4 | 5 | \SYNOPSIS 6 | Returns a list of the variables free in a term. 7 | 8 | \DESCRIBE 9 | When applied to a term, {frees} returns a list of the free variables in 10 | that term. There are no repetitions in the list produced even if there are 11 | multiple free instances of some variables. 12 | 13 | \FAILURE 14 | Never fails. 15 | 16 | \EXAMPLE 17 | Clearly in the following term, {x} and {y} are free, whereas {z} is bound: 18 | { 19 | # frees `x = 1 /\ y = 2 /\ !z. z >= 0`;; 20 | val it : term list = [`x`; `y`] 21 | } 22 | 23 | \SEEALSO 24 | atoms, freesl, free_in, thm_frees, variables. 25 | 26 | \ENDDOC 27 | -------------------------------------------------------------------------------- /Help/index.hlp: -------------------------------------------------------------------------------- 1 | \DOC index 2 | 3 | \TYPE {index : 'a -> 'a list -> int} 4 | 5 | \SYNOPSIS 6 | Returns position of given element in list. 7 | 8 | \DESCRIBE 9 | The call {index x l} where l is a list returns the position number of the first 10 | instance of x in the list, failing if there is none. The indices start at zero, 11 | corresponding to {el}. 12 | 13 | \EXAMPLE 14 | { 15 | # index "j" (explode "abcdefghijklmnopqrstuvwxyz");; 16 | val it : int = 9 17 | } 18 | This is a sort of inverse to the indexing into a string by {el}: 19 | { 20 | # el 9 (explode "abcdefghijklmnopqrstuvwxyz");; 21 | val it : string = "j" 22 | } 23 | 24 | \SEEALSO 25 | el, find. 26 | 27 | \ENDDOC 28 | -------------------------------------------------------------------------------- /Help/inst_goal.hlp: -------------------------------------------------------------------------------- 1 | \DOC inst_goal 2 | 3 | \TYPE {inst_goal : instantiation -> goal -> goal} 4 | 5 | \SYNOPSIS 6 | Apply higher-order instantiation to a goal. 7 | 8 | \DESCRIBE 9 | The call {inst_goal i g} where {i} is an instantiation (as returned by 10 | {term_match} for example), will perform the instantiation indicated by {i} in 11 | both assumptions and conclusion of the goal {g}. 12 | 13 | \FAILURE 14 | Should never fail on a valid instantiation. 15 | 16 | \COMMENTS 17 | Probably only of specialist interest to those writing tactics from scratch. 18 | 19 | \SEEALSO 20 | compose_insts, instantiate, INSTANTIATE, INSTANTIATE_ALL, 21 | PART_MATCH, term_match. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Help/basic_ss.hlp: -------------------------------------------------------------------------------- 1 | \DOC basic_ss 2 | 3 | \TYPE {basic_ss : thm list -> simpset} 4 | 5 | \SYNOPSIS 6 | Construct a straightforward simpset from a list of theorems. 7 | 8 | \DESCRIBE 9 | In their maximal generality, simplification operations in HOL Light (as invoked 10 | by {SIMP_TAC}) are controlled by a `simpset'. A call {basic_ss thl} gives a 11 | straightforward simpset used by the default simplifier instances like 12 | {SIMP_TAC}, which has the given theorems as well as the basic rewrites and 13 | conversions, and no other provers. 14 | 15 | \FAILURE 16 | Never fails. 17 | 18 | \SEEALSO 19 | basic_convs, basic_rewrites, empty_ss, SIMP_CONV, SIMP_RULE, SIMP_TAC. 20 | 21 | \ENDDOC 22 | -------------------------------------------------------------------------------- /Help/dest_finty.hlp: -------------------------------------------------------------------------------- 1 | \DOC dest_finty 2 | 3 | \TYPE {dest_finty : hol_type -> num} 4 | 5 | \SYNOPSIS 6 | Converts a standard finite type to corresponding integer. 7 | 8 | \DESCRIBE 9 | Finite types parsed and printed as numerals are provided, and this operation 10 | when applied to such a type gives the corresponding number. 11 | 12 | \FAILURE 13 | Fails if the type is not a standard finite type. 14 | 15 | \EXAMPLE 16 | Here we use a 32-element type, perhaps useful for indexing the bits of a 17 | word: 18 | { 19 | # dest_finty `:32`;; 20 | val it : num = 32 21 | } 22 | 23 | \SEEALSO 24 | dest_type, DIMINDEX_CONV, DIMINDEX_TAC, HAS_SIZE_DIMINDEX_RULE, mk_finty. 25 | 26 | \ENDDOC 27 | -------------------------------------------------------------------------------- /Help/list_mk_forall.hlp: -------------------------------------------------------------------------------- 1 | \DOC list_mk_forall 2 | 3 | \TYPE {list_mk_forall : term list * term -> term} 4 | 5 | \SYNOPSIS 6 | Iteratively constructs a universal quantification. 7 | 8 | \DESCRIBE 9 | {list_mk_forall([`x1`;...;`xn`],`t`)} returns {`!x1 ... xn. t`}. 10 | 11 | \FAILURE 12 | Fails if any term in the list is not a variable or if {t} is not of type 13 | {`:bool`} and the list of terms is non-empty. If the list of terms is empty the 14 | type of {t} can be anything. 15 | 16 | \EXAMPLE 17 | { 18 | # list_mk_forall([`x:num`; `y:num`],`x + y + 1 = SUC z`);; 19 | val it : term = `!x y. x + y + 1 = SUC z` 20 | } 21 | 22 | \SEEALSO 23 | mk_forall, strip_forall. 24 | 25 | \ENDDOC 26 | -------------------------------------------------------------------------------- /Help/parses_as_binder.hlp: -------------------------------------------------------------------------------- 1 | \DOC parses_as_binder 2 | 3 | \TYPE {parses_as_binder : string -> bool} 4 | 5 | \SYNOPSIS 6 | Tests if a string has binder status in the parser. 7 | 8 | \DESCRIBE 9 | Certain identifiers {c} have binder status, meaning that {`c x. y`} is parsed 10 | as a shorthand for {`(c) (\x. y)'}. The call {parses_as_binder "c"} tests if 11 | {c} is one of the identifiers with binder status. 12 | 13 | \FAILURE 14 | Never fails. 15 | 16 | \EXAMPLE 17 | { 18 | # parses_as_binder "!";; 19 | val it : bool = true 20 | # parses_as_binder "==>";; 21 | val it : bool = false 22 | } 23 | 24 | \SEEALSO 25 | binders, parses_as_binder, unparse_as_binder. 26 | 27 | \ENDDOC 28 | -------------------------------------------------------------------------------- /Help/reduce_interface.hlp: -------------------------------------------------------------------------------- 1 | \DOC reduce_interface 2 | 3 | \TYPE {reduce_interface : string * term -> unit} 4 | 5 | \SYNOPSIS 6 | Remove a specific overload/interface mapping for an identifier. 7 | 8 | \DESCRIBE 9 | HOL Light allows an identifier to map to a specific constant (see 10 | {override_interface}) or be overloaded to several depending on type (see 11 | {overload_interface}). A call to {remove_interface "ident"} removes all such 12 | mappings for the identifier {ident}. 13 | 14 | \FAILURE 15 | Never fails, whether or not there were any interface mappings in effect. 16 | 17 | \SEEALSO 18 | overload_interface, override_interface, remove_interface, the_interface. 19 | 20 | \ENDDOC 21 | -------------------------------------------------------------------------------- /Help/thm_frees.hlp: -------------------------------------------------------------------------------- 1 | \DOC thm_frees 2 | 3 | \TYPE {thm_frees : thm -> term list} 4 | 5 | \SYNOPSIS 6 | Returns a list of the variables free in a theorem's assumptions and conclusion. 7 | 8 | \DESCRIBE 9 | When applied to a theorem, {A |- t}, the function {thm_frees} returns a list, 10 | without repetitions, of those variables which are free either in {t} or in 11 | some member of the assumption list {A}. 12 | 13 | \FAILURE 14 | Never fails. 15 | 16 | \EXAMPLE 17 | { 18 | # let th = CONJUNCT1 (ASSUME `p /\ q`);; 19 | val th : thm = p /\ q |- p 20 | 21 | # thm_frees th;; 22 | val it : term list = [`q`; `p`] 23 | } 24 | 25 | \SEEALSO 26 | frees, freesl, free_in. 27 | 28 | \ENDDOC 29 | -------------------------------------------------------------------------------- /Help/type_unify.hlp: -------------------------------------------------------------------------------- 1 | \DOC type_unify 2 | 3 | \TYPE {type_unify : hol_type -> hol_type -> (hol_type * hol_type) list -> (hol_type * hol_type) list} 4 | 5 | \SYNOPSIS 6 | Unify two types by instantiating their type variables 7 | 8 | \DESCRIBE 9 | Given two types {ty1} and {ty2} and an existing instantiation {i} of type 10 | variables, a call {type_unify vars ty1 ty2 i} attempts to find an augmented 11 | instantiation of the type variables to make the two types equal. 12 | 13 | \FAILURE 14 | Fails if the two types are not first-order unifiable by instantiating the given 15 | type variables. 16 | 17 | \SEEALSO 18 | instantiate, term_match, term_type_unify, term_unify. 19 | 20 | \ENDDOC 21 | -------------------------------------------------------------------------------- /Help/DEPTH_SQCONV.hlp: -------------------------------------------------------------------------------- 1 | \DOC DEPTH_SQCONV 2 | 3 | \TYPE {DEPTH_SQCONV : strategy} 4 | 5 | \SYNOPSIS 6 | Applies simplification repeatedly to all the sub-terms of a term, in bottom-up 7 | order. 8 | 9 | \DESCRIBE 10 | HOL Light's simplification functions (e.g. {SIMP_TAC}) have their traversal 11 | algorithm controlled by a ``strategy''. {DEPTH_SQCONV} is a strategy 12 | corresponding to {DEPTH_CONV} for ordinary conversions: simplification is 13 | applied repeatedly to all the sub-terms of a term, in bottom-up 14 | order. 15 | 16 | \FAILURE 17 | Not applicable. 18 | 19 | \SEEALSO 20 | DEPTH_CONV, ONCE_DEPTH_SQCONV, REDEPTH_SQCONV, TOP_DEPTH_SQCONV, 21 | TOP_SWEEP_SQCONV. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Help/is_intconst.hlp: -------------------------------------------------------------------------------- 1 | \DOC is_intconst 2 | 3 | \TYPE {is_intconst : term -> bool} 4 | 5 | \SYNOPSIS 6 | Tests if a term is an integer literal of type {:int}. 7 | 8 | \DESCRIBE 9 | The call {is_intconst t} tests whether the term {t} is a canonical integer 10 | literal of type {:int}, i.e. either `{&n}' for a numeral {n} or `{-- &n}' for 11 | a nonzero numeral {n}. If so it returns {true}, otherwise {false}. 12 | 13 | \FAILURE 14 | Never fails. 15 | 16 | \EXAMPLE 17 | { 18 | # is_intconst `-- &3 :int`;; 19 | val it : bool = true 20 | # is_intconst `-- &0 :int`;; 21 | val it : bool = false 22 | } 23 | 24 | \SEEALSO 25 | dest_intconst, is_realintconst, mk_intconst. 26 | 27 | \ENDDOC 28 | -------------------------------------------------------------------------------- /Help/is_setenum.hlp: -------------------------------------------------------------------------------- 1 | \DOC is_setenum 2 | 3 | \TYPE {is_setenum : term -> bool} 4 | 5 | \SYNOPSIS 6 | Tests if a term is a set enumeration. 7 | 8 | \DESCRIBE 9 | When applied to a term that is an explicit set enumeration {`{{t1,...,tn}}`}, 10 | the function {is_setenum} returns {true}; otherwise it returns {false}. 11 | 12 | \FAILURE 13 | Never fails. 14 | 15 | \EXAMPLE 16 | { 17 | # is_setenum `1 INSERT 2 INSERT {{}}`;; 18 | val it : bool = true 19 | 20 | # is_setenum `{{1,2,3,4,1,2,3,4}}`;; 21 | val it : bool = true 22 | 23 | # is_setenum `1 INSERT 2 INSERT s`;; 24 | val it : bool = false 25 | } 26 | 27 | \SEEALSO 28 | dest_setenum, mk_fset, mk_setenum. 29 | 30 | \ENDDOC 31 | -------------------------------------------------------------------------------- /Help/repeat.hlp: -------------------------------------------------------------------------------- 1 | \DOC repeat 2 | 3 | \TYPE {repeat : ('a -> 'a) -> 'a -> 'a} 4 | 5 | \SYNOPSIS 6 | Repeatedly apply a function until it fails. 7 | 8 | \DESCRIBE 9 | The call {repeat f x} successively applies {f} over and over again starting 10 | with {x}, and stops at the first point when a {Failure _} exception occurs. 11 | 12 | \FAILURE 13 | Never fails. If {f} fails at once it returns {x}. 14 | 15 | \EXAMPLE 16 | { 17 | # repeat (snd o dest_forall) `!x y z. x + y + z < 1`;; 18 | val it : term = `x + y + z < 1` 19 | } 20 | 21 | \COMMENTS 22 | If you know exactly how many times you want to apply it, you may prefer 23 | {funpow}. 24 | 25 | \SEEALSO 26 | funpow, fail. 27 | 28 | \ENDDOC 29 | -------------------------------------------------------------------------------- /Help/ASM.hlp: -------------------------------------------------------------------------------- 1 | \DOC ASM 2 | 3 | \TYPE {ASM : (thm list -> tactic) -> thm list -> tactic} 4 | 5 | \SYNOPSIS 6 | Augments a tactic's theorem list with the assumptions. 7 | 8 | \DESCRIBE 9 | If {tac} is a tactic that expects a list of theorems as its arguments, e.g. 10 | {MESON_TAC}, {REWRITE_TAC} or {SET_TAC}, then {ASM tac} converts it to a tactic 11 | where that list is augmented by the goal's assumptions. 12 | 13 | \FAILURE 14 | Never fails (though the resulting tactic may do). 15 | 16 | \EXAMPLE 17 | { 18 | The inbuilt {ASM_REWRITE_TAC} is in fact defined as just {ASM REWRITE_TAC}. 19 | } 20 | 21 | \SEEALSO 22 | ASSUM_LIST, FREEZE_THEN, HYP, MESON_TAC, REWRITE_TAC, SET_TAC. 23 | 24 | \ENDDOC 25 | -------------------------------------------------------------------------------- /Help/ORELSE_TCL.hlp: -------------------------------------------------------------------------------- 1 | \DOC ORELSE_TCL 2 | 3 | \TYPE {(ORELSE_TCL) : thm_tactical -> thm_tactical -> thm_tactical} 4 | 5 | \SYNOPSIS 6 | Applies a theorem-tactical, and if it fails, tries a second. 7 | 8 | \KEYWORDS 9 | theorem-tactical. 10 | 11 | \DESCRIBE 12 | When applied to two theorem-tacticals, {ttl1} and {ttl2}, a theorem-tactic 13 | {ttac}, and a theorem {th}, if {ttl1 ttac th} succeeds, that gives the 14 | result. If it fails, the result is {ttl2 ttac th}, which may itself fail. 15 | 16 | \FAILURE 17 | {ORELSE_TCL} fails if both the theorem-tacticals fail when applied to the 18 | given theorem-tactic and theorem. 19 | 20 | \SEEALSO 21 | EVERY_TCL, FIRST_TCL, THEN_TCL. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /100/arithmetic.ml: -------------------------------------------------------------------------------- 1 | (* ========================================================================= *) 2 | (* Sum of an arithmetic series. *) 3 | (* ========================================================================= *) 4 | 5 | let ARITHMETIC_PROGRESSION_LEMMA = prove 6 | (`!n. nsum(0..n) (\i. a + d * i) = ((n + 1) * (2 * a + n * d)) DIV 2`, 7 | INDUCT_TAC THEN ASM_REWRITE_TAC[NSUM_CLAUSES_NUMSEG] THEN ARITH_TAC);; 8 | 9 | let ARITHMETIC_PROGRESSION = prove 10 | (`!n. 1 <= n 11 | ==> nsum(0..n-1) (\i. a + d * i) = (n * (2 * a + (n - 1) * d)) DIV 2`, 12 | INDUCT_TAC THEN REWRITE_TAC[ARITHMETIC_PROGRESSION_LEMMA; SUC_SUB1] THEN 13 | ARITH_TAC);; 14 | -------------------------------------------------------------------------------- /Formal_ineqs/README.md: -------------------------------------------------------------------------------- 1 | A tool for verification of nonlinear inequalities in HOL Light 2 | ============ 3 | 4 | The most recent version of the tool is available 5 | [here](https://github.com/monadius/formal_ineqs). 6 | 7 | Part of the Flyspeck project: 8 | https://github.com/flyspeck/flyspeck 9 | 10 | See [docs/FormalVerifier.pdf](docs/FormalVerifier.pdf) for additional information. 11 | 12 | 13 | Copyright / License 14 | ------------------- 15 | 16 | Except lib/ipow.hl, all files in this directory are copyright Alexey Solovyev 2014-2017. The file lib/ipow.hl is copyright Alexey Solovyev and the University of Utah 2014-2017. 17 | 18 | Distributed under the same license as HOL Light (BSD-2-clause). 19 | -------------------------------------------------------------------------------- /Help/ABS.hlp: -------------------------------------------------------------------------------- 1 | \DOC ABS 2 | 3 | \TYPE {ABS : term -> thm -> thm} 4 | 5 | \SYNOPSIS 6 | Abstracts both sides of an equation. 7 | 8 | \KEYWORDS 9 | rule, abstraction. 10 | 11 | \DESCRIBE 12 | { 13 | A |- t1 = t2 14 | ------------------------ ABS `x` [Where x is not free in A] 15 | A |- (\x.t1) = (\x.t2) 16 | } 17 | 18 | \FAILURE 19 | If the theorem is not an equation, or if the variable {x} is free in the 20 | assumptions {A}. 21 | 22 | \EXAMPLE 23 | { 24 | # ABS `m:num` (REFL `m:num`);; 25 | val it : thm = |- (\m. m) = (\m. m) 26 | } 27 | 28 | \COMMENTS 29 | This is one of HOL Light's 10 primitive inference rules. 30 | 31 | \SEEALSO 32 | ETA_CONV. 33 | 34 | \ENDDOC 35 | -------------------------------------------------------------------------------- /Help/binders.hlp: -------------------------------------------------------------------------------- 1 | \DOC binders 2 | 3 | \TYPE {binders : unit -> string list} 4 | 5 | \SYNOPSIS 6 | Lists the binders. 7 | 8 | \DESCRIBE 9 | The function {binders} returns a list of all the binders declared so far. A 10 | binder {b} is then parsed in constructs like {b x. t[x]} as an abbreviation for 11 | {(b) (\x. t[x])}. The set of binders can be changed with {parse_as_binder} and 12 | {unparse_as_binder}. 13 | 14 | \FAILURE 15 | Never fails 16 | 17 | \EXAMPLE 18 | { 19 | # binders();; 20 | val it : string list = ["\\"; "!"; "?"; "?!"; "@"; "minimal"; "lambda"] 21 | } 22 | \SEEALSO 23 | parse_as_binder, parses_as_binder, parse_as_infix, parse_as_prefix, 24 | unparse_as_binder. 25 | 26 | \ENDDOC 27 | -------------------------------------------------------------------------------- /Help/exists.hlp: -------------------------------------------------------------------------------- 1 | \DOC exists 2 | 3 | \TYPE {exists : ('a -> bool) -> 'a list -> bool} 4 | 5 | \SYNOPSIS 6 | Tests a list to see if some element satisfy a predicate. 7 | 8 | \KEYWORDS 9 | list. 10 | 11 | \DESCRIBE 12 | {exists p [x1;...;xn]} returns {true} if {(p xi)} is true for some {xi} in the 13 | list. Otherwise, for example if the list is empty, it returns {false}. 14 | 15 | \FAILURE 16 | Never fails. 17 | 18 | \EXAMPLE 19 | { 20 | # exists (fun n -> n mod 2 = 0) [2;3;5;7;11;13;17];; 21 | val it : bool = true 22 | # exists (fun n -> n mod 2 = 0) [3;5;7;9;11;13;15];; 23 | val it : bool = false 24 | } 25 | 26 | \SEEALSO 27 | find, forall, tryfind, mem, assoc, rev_assoc. 28 | 29 | \ENDDOC 30 | -------------------------------------------------------------------------------- /Help/forall.hlp: -------------------------------------------------------------------------------- 1 | \DOC forall 2 | 3 | \TYPE {forall : ('a -> bool) -> 'a list -> bool} 4 | 5 | \SYNOPSIS 6 | Tests a list to see if all its elements satisfy a predicate. 7 | 8 | \KEYWORDS 9 | list. 10 | 11 | \DESCRIBE 12 | {forall p [x1;...;xn]} returns {true} if {(p xi)} is true for all {xi} in the 13 | list. Otherwise it returns {false}. If the list is empty, this function always 14 | returns true. 15 | 16 | \FAILURE 17 | Never fails. 18 | 19 | \EXAMPLE 20 | { 21 | # forall (fun x -> x <= 2) [0;1;2];; 22 | val it : bool = true 23 | # forall (fun x -> x <= 2) [1;2;3];; 24 | val it : bool = false 25 | } 26 | 27 | \SEEALSO 28 | exists, find, tryfind, mem, assoc, rev_assoc. 29 | 30 | \ENDDOC 31 | -------------------------------------------------------------------------------- /Help/gcd.hlp: -------------------------------------------------------------------------------- 1 | \DOC gcd 2 | 3 | \TYPE {gcd : int -> int -> int} 4 | 5 | \SYNOPSIS 6 | Computes greatest common divisor of two integers. 7 | 8 | \DESCRIBE 9 | The call {gcd m n} for two integers {m} and {n} returns the (nonnegative) 10 | greatest common divisor of {m} and {n}. If {m} and {n} are both zero, it 11 | returns zero. 12 | 13 | \FAILURE 14 | Never fails. 15 | 16 | \EXAMPLE 17 | { 18 | # gcd 10 12;; 19 | val it : int = 2 20 | 21 | # gcd 11 27;; 22 | val it : int = 1 23 | 24 | # gcd (-33) 76;; 25 | val it : int = 1 26 | 27 | # gcd 0 99;; 28 | val it : int = 99 29 | 30 | # gcd 0 0;; 31 | val it : int = 0 32 | } 33 | 34 | \SEEALSO 35 | gcd_num, lcm_num. 36 | 37 | \ENDDOC 38 | -------------------------------------------------------------------------------- /Help/preterm_of_term.hlp: -------------------------------------------------------------------------------- 1 | \DOC preterm_of_term 2 | 3 | \TYPE {preterm_of_term : term -> preterm} 4 | 5 | \SYNOPSIS 6 | Converts a term into a preterm. 7 | 8 | \DESCRIBE 9 | HOL Light uses ``pretypes'' and ``preterms'' as intermediate structures for 10 | parsing and typechecking, which are later converted to types and terms. A call 11 | {preterm_of_term `tm`} converts in the other direction, from a normal HOL term 12 | back to a preterm. 13 | 14 | \FAILURE 15 | Never fails. 16 | 17 | \USES 18 | User manipulation of preterms is not usually necessary, unless you seek to 19 | radically change aspects of parsing and typechecking. 20 | 21 | \SEEALSO 22 | pretype_of_type, term_of_preterm. 23 | 24 | \ENDDOC 25 | -------------------------------------------------------------------------------- /Help/quotexpander.hlp: -------------------------------------------------------------------------------- 1 | \DOC quotexpander 2 | 3 | \TYPE {quotexpander : string -> string} 4 | 5 | \SYNOPSIS 6 | Quotation expander. 7 | 8 | \DESCRIBE 9 | This function determines how anything in {`backquotes`} is expanded on input. 10 | 11 | \FAILURE 12 | Never fails. 13 | 14 | \EXAMPLE 15 | { 16 | # quotexpander "1 + 1";; 17 | val it : string = "parse_term \"1 + 1\"" 18 | # quotexpander ":num";; 19 | val it : string = "parse_type \"num\"" 20 | } 21 | 22 | \COMMENTS 23 | Not intended for general use, but automatically invoked when anything is typed 24 | in backquotes {`like this`}. May be of some interest for users wishing to 25 | change the behavior of the quotation parser. 26 | 27 | \ENDDOC 28 | -------------------------------------------------------------------------------- /Help/decreasing.hlp: -------------------------------------------------------------------------------- 1 | \DOC decreasing 2 | 3 | \TYPE {decreasing : ('a -> 'b) -> 'a -> 'a -> bool} 4 | 5 | \SYNOPSIS 6 | When applied to a ``measure'' function {f}, the call {increasing f} returns a 7 | binary function ordering elements in a call {increasing f x y} by 8 | {f(y) string} 4 | 5 | \SYNOPSIS 6 | Gets the name of a constant or variable. 7 | 8 | \DESCRIBE 9 | When applied to a term that is either a constant or a variable, {name_of} 10 | returns its name (its true name, even if there is an interface mapping for it 11 | in effect). When applied to any other term, it returns the empty string. 12 | 13 | \FAILURE 14 | Never fails. 15 | 16 | \EXAMPLE 17 | { 18 | # name_of `x:int`;; 19 | val it : string = "x" 20 | 21 | # name_of `SUC`;; 22 | val it : string = "SUC" 23 | 24 | # name_of `x + 1`;; 25 | val it : string = "" 26 | 27 | } 28 | 29 | \SEEALSO 30 | dest_const, dest_var. 31 | 32 | \ENDDOC 33 | -------------------------------------------------------------------------------- /Help/pretype_of_type.hlp: -------------------------------------------------------------------------------- 1 | \DOC pretype_of_type 2 | 3 | \TYPE {pretype_of_type : hol_type -> pretype} 4 | 5 | \SYNOPSIS 6 | Converts a type into a pretype. 7 | 8 | \DESCRIBE 9 | HOL Light uses ``pretypes'' and ``preterms'' as intermediate structures for 10 | parsing and typechecking, which are later converted to types and terms. A call 11 | {preterm_of_term `tm`} converts in the other direction, from a normal HOL term 12 | back to a preterm. 13 | 14 | \FAILURE 15 | Never fails. 16 | 17 | \USES 18 | User manipulation of pretypes is not usually necessary, unless you seek to 19 | radically change aspects of parsing and typechecking. 20 | 21 | \SEEALSO 22 | preterm_of_term, type_of_pretype. 23 | 24 | \ENDDOC 25 | -------------------------------------------------------------------------------- /Help/AP_THM_TAC.hlp: -------------------------------------------------------------------------------- 1 | \DOC AP_THM_TAC 2 | 3 | \TYPE {AP_THM_TAC : tactic} 4 | 5 | \SYNOPSIS 6 | Strips identical operands from functions on both sides of an equation. 7 | 8 | \KEYWORDS 9 | tactic. 10 | 11 | \DESCRIBE 12 | When applied to a goal of the form {A ?- f x = g x}, the tactic {AP_THM_TAC} 13 | strips away the operands of the function application: 14 | { 15 | A ?- f x = g x 16 | ================ AP_THM_TAC 17 | A ?- f = g 18 | } 19 | 20 | \FAILURE 21 | Fails unless the goal has the above form, namely an equation both sides of 22 | which consist of function applications to the same argument. 23 | 24 | \SEEALSO 25 | ABS_TAC, AP_TERM_TAC, AP_THM, BINOP_TAC, MK_COMB_TAC. 26 | 27 | \ENDDOC 28 | -------------------------------------------------------------------------------- /Help/GEN_SIMPLIFY_CONV.hlp: -------------------------------------------------------------------------------- 1 | \DOC GEN_SIMPLIFY_CONV 2 | 3 | \TYPE {GEN_SIMPLIFY_CONV : strategy -> simpset -> int -> thm list -> conv} 4 | 5 | \SYNOPSIS 6 | General simplification with given strategy and simpset and theorems. 7 | 8 | \DESCRIBE 9 | The call {GEN_SIMPLIFY_CONV strat ss n thl} incorporates the rewrites and 10 | conditional rewrites derived from {thl} into the simpset {ss}, then simplifies 11 | using that simpset, controlling the traversal of the term by {strat}, and 12 | starting at level {n}. 13 | 14 | \FAILURE 15 | Never fails unless some component is malformed. 16 | 17 | \SEEALSO 18 | GEN_REWRITE_CONV, ONCE_SIMPLIFY_CONV, SIMPLIFY_CONV, SIMP_CONV, SIMP_RULE, 19 | SIMP_TAC. 20 | 21 | \ENDDOC 22 | -------------------------------------------------------------------------------- /Help/ONCE_DEPTH_SQCONV.hlp: -------------------------------------------------------------------------------- 1 | \DOC ONCE_DEPTH_SQCONV 2 | 3 | \TYPE {ONCE_DEPTH_SQCONV : strategy} 4 | 5 | \SYNOPSIS 6 | Applies simplification to the first suitable sub-term(s) encountered in 7 | top-down order. 8 | 9 | \DESCRIBE 10 | HOL Light's simplification functions (e.g. {SIMP_TAC}) have their traversal 11 | algorithm controlled by a ``strategy''. {ONCE_DEPTH_SQCONV} is a strategy 12 | corresponding to {ONCE_DEPTH_CONV} for ordinary conversions: simplification is 13 | applied to the first suitable subterm(s) encountered in top-down order. 14 | 15 | \FAILURE 16 | Not applicable. 17 | 18 | \SEEALSO 19 | DEPTH_SQCONV, ONCE_DEPTH_CONV, REDEPTH_SQCONV, TOP_DEPTH_SQCONV, 20 | TOP_SWEEP_SQCONV. 21 | 22 | \ENDDOC 23 | -------------------------------------------------------------------------------- /Help/dest_let.hlp: -------------------------------------------------------------------------------- 1 | \DOC dest_let 2 | 3 | \TYPE {dest_let : term -> (term * term) list * term} 4 | 5 | \SYNOPSIS 6 | Breaks apart a let-expression. 7 | 8 | \DESCRIBE 9 | {dest_let} is a term destructor for general let-expressions: 10 | {dest_let `let x1 = e1 and ... and xn = en in E`} returns a pair 11 | of the list {[`x1`,`e1`; ... ; `xn`,`en`]} and {`E`}. 12 | 13 | \FAILURE 14 | Fails with {dest_let} if term is not a {let}-expression. 15 | 16 | \EXAMPLE 17 | { 18 | # dest_let `let m = 256 and n = 65536 in (x MOD m + y MOD m) MOD n`;; 19 | val it : (term * term) list * term = 20 | ([(`m`, `256`); (`n`, `65536`)], `(x MOD m + y MOD m) MOD n`) 21 | } 22 | \SEEALSO 23 | mk_let, is_let. 24 | 25 | \ENDDOC 26 | -------------------------------------------------------------------------------- /Help/help_path.hlp: -------------------------------------------------------------------------------- 1 | \DOC help_path 2 | 3 | \TYPE {help_path : string list ref} 4 | 5 | \SYNOPSIS 6 | Path where HOL Light tries to find help files. 7 | 8 | \DESCRIBE 9 | The reference variable {help_path} gives a list of directories. When using the 10 | online {help} function, HOL Light will search in these places for help files. 11 | An initial dollar sign {$} in each path is interpreted as a reference to the 12 | current setting of {hol_dir}. To get an actual {$} at the start of the 13 | filename, actually use two dollar signs {$$}. 14 | 15 | \FAILURE 16 | Not applicable. 17 | 18 | \SEEALSO 19 | file_on_path, help, hol_dir, hol_expand_directory, load_on_path, load_path, 20 | loads, loadt. 21 | 22 | \ENDDOC 23 | -------------------------------------------------------------------------------- /Help/subtract.hlp: -------------------------------------------------------------------------------- 1 | \DOC subtract 2 | 3 | \TYPE {subtract : 'a list -> 'a list -> 'a list} 4 | 5 | \SYNOPSIS 6 | Computes the set-theoretic difference of two `sets'. 7 | 8 | \KEYWORDS 9 | list, set. 10 | 11 | \DESCRIBE 12 | {subtract l1 l2} returns a list consisting of those elements of {l1} that do 13 | not appear in {l2}. If both lists are initially free of repetitions, this can 14 | be considered a set difference operation. 15 | 16 | \FAILURE 17 | Never fails. 18 | 19 | \EXAMPLE 20 | { 21 | # subtract [1;2;3] [3;5;4;1];; 22 | val it : int list = [2] 23 | # subtract [1;2;4;1] [4;5];; 24 | val it : int list = [1; 2; 1] 25 | } 26 | 27 | \SEEALSO 28 | setify, set_eq, union, intersect. 29 | 30 | \ENDDOC 31 | -------------------------------------------------------------------------------- /Help/union.hlp: -------------------------------------------------------------------------------- 1 | \DOC union 2 | 3 | \TYPE {union : 'a list -> 'a list -> 'a list} 4 | 5 | \SYNOPSIS 6 | Computes the union of two `sets'. 7 | 8 | \KEYWORDS 9 | list, set. 10 | 11 | \DESCRIBE 12 | {union l1 l2} returns a list consisting of the elements of {l1} not already in 13 | {l2} concatenated with {l2}. If {l1} and {l2} are initially free from 14 | duplicates, this gives a set-theoretic union operation. 15 | 16 | \FAILURE 17 | Never fails. 18 | 19 | \EXAMPLE 20 | { 21 | # union [1;2;3] [1;5;4;3];; 22 | val it : int list = [2; 1; 5; 4; 3] 23 | # union [1;1;1] [1;2;3;2];; 24 | val it : int list = [1; 2; 3; 2] 25 | } 26 | 27 | \SEEALSO 28 | setify, set_equal, intersect, subtract. 29 | 30 | \ENDDOC 31 | -------------------------------------------------------------------------------- /Help/EQ_TAC.hlp: -------------------------------------------------------------------------------- 1 | \DOC EQ_TAC 2 | 3 | \TYPE {EQ_TAC : tactic} 4 | 5 | \SYNOPSIS 6 | Reduces goal of equality of boolean terms to forward and backward implication. 7 | 8 | \KEYWORDS 9 | tactic, equality, implication. 10 | 11 | \DESCRIBE 12 | When applied to a goal {A ?- t1 <=> t2}, where {t1} and {t2} have type {bool}, 13 | the tactic {EQ_TAC} returns the subgoals {A ?- t1 ==> t2} and 14 | {A ?- t2 ==> t1}. 15 | { 16 | A ?- t1 <=> t2 17 | ================================= EQ_TAC 18 | A ?- t1 ==> t2 A ?- t2 ==> t1 19 | } 20 | 21 | \FAILURE 22 | Fails unless the conclusion of the goal is an equation between boolean terms. 23 | 24 | \SEEALSO 25 | EQ_IMP_RULE, IMP_ANTISYM_RULE. 26 | 27 | \ENDDOC 28 | -------------------------------------------------------------------------------- /Help/SIMP_RULE.hlp: -------------------------------------------------------------------------------- 1 | \DOC SIMP_RULE 2 | 3 | \TYPE {SIMP_RULE : thm list -> thm -> thm} 4 | 5 | \SYNOPSIS 6 | Simplify conclusion of a theorem repeatedly by conditional contextual 7 | rewriting. 8 | 9 | \DESCRIBE 10 | A call {SIMP_CONV thl (|- tm)} will return {|- tm'} where {tm'} results from 11 | applying the theorems in {thl} as (conditional) rewrite rules, as well as 12 | built-in simplifications (see {basic_rewrites} and {basic_convs}). For more 13 | details on this kind of conditional rewriting, see {SIMP_CONV}. 14 | 15 | \FAILURE 16 | Never fails, but may return the input theorem unchanged if no simplifications 17 | were applicable. 18 | 19 | \SEEALSO 20 | ONCE_SIMP_RULE, SIMP_CONV, SIMP_TAC. 21 | 22 | \ENDDOC 23 | -------------------------------------------------------------------------------- /Help/list_mk_conj.hlp: -------------------------------------------------------------------------------- 1 | \DOC list_mk_conj 2 | 3 | \TYPE {list_mk_conj : term list -> term} 4 | 5 | \SYNOPSIS 6 | Constructs the conjunction of a list of terms. 7 | 8 | \DESCRIBE 9 | {list_mk_conj([`t1`;...;`tn`])} returns {`t1 /\ ... /\ tn`}. 10 | 11 | \FAILURE 12 | Fails with {list_mk_conj} if the list is empty or if the list has more than 13 | one element, one or more of which are not of type {`:bool`}. 14 | 15 | \EXAMPLE 16 | { 17 | # list_mk_conj [`T`;`F`;`T`];; 18 | val it : term = `T /\ F /\ T` 19 | 20 | # list_mk_conj [`T`;`1`;`F`];; 21 | Exception: Failure "mk_binary". 22 | 23 | # list_mk_conj [`1`];; 24 | val it : term = `1` 25 | } 26 | 27 | \SEEALSO 28 | conjuncts, mk_conj. 29 | 30 | \ENDDOC 31 | -------------------------------------------------------------------------------- /Help/mk_char.hlp: -------------------------------------------------------------------------------- 1 | \DOC mk_char 2 | 3 | \TYPE {mk_char : char -> term} 4 | 5 | \SYNOPSIS 6 | Constructs object-level character from OCaml character. 7 | 8 | \DESCRIBE 9 | {mk_char 'c'} produces the HOL term of type {char} corresponding to the OCaml 10 | character {c}. 11 | 12 | \FAILURE 13 | Never fails 14 | 15 | \EXAMPLE 16 | { 17 | # mk_char 'c';; 18 | val it : term = `ASCII F T T F F F T T` 19 | } 20 | 21 | \COMMENTS 22 | There is no particularly convenient parser/printer support for the HOL {char} 23 | type, but when combined into lists they are considered as strings and provided 24 | with more intuitive parser/printer support. 25 | 26 | \SEEALSO 27 | dest_char, dest_string, mk_string. 28 | 29 | \ENDDOC 30 | -------------------------------------------------------------------------------- /Help/REPLICATE_TAC.hlp: -------------------------------------------------------------------------------- 1 | \DOC REPLICATE_TAC 2 | 3 | \TYPE {REPLICATE_TAC : int -> tactic -> tactic} 4 | 5 | \SYNOPSIS 6 | Apply a tactic a specific number of times. 7 | 8 | \DESCRIBE 9 | The call {REPLICATE n tac} gives a new tactic that it equivalent to an {n}-fold 10 | repetition of {tac}, i.e. {tac THEN tac THEN ... THEN tac}. 11 | 12 | \FAILURE 13 | The call {REPLICATE n tac} never fails, but when applied to a goal it will fail 14 | if the tactic does. 15 | 16 | \EXAMPLE 17 | We might conceivably want to strip off exactly three universal quantifiers from 18 | a goal that contains more than three. We can use {REPLICATE_TAC 3 GEN_TAC} to 19 | do that. 20 | 21 | \SEEALSO 22 | EVERY, MAP_EVERY, THEN. 23 | 24 | \ENDDOC 25 | -------------------------------------------------------------------------------- /Help/extend_basic_rewrites.hlp: -------------------------------------------------------------------------------- 1 | \DOC extend_basic_rewrites 2 | 3 | \TYPE {extend_basic_rewrites : thm list -> unit} 4 | 5 | \SYNOPSIS 6 | Extend the set of default rewrites used by rewriting and simplification. 7 | 8 | \DESCRIBE 9 | The HOL Light rewriter ({REWRITE_TAC} etc.) and simplifier ({SIMP_TAC} etc.) 10 | have default sets of (conditional) equations and other conversions that are 11 | applied by default, except in the {PURE_} variants. A call to 12 | {extend_basic_rewrites thl} extends the former with the list of theorems {thl}, 13 | so they will thereafter happen by default. 14 | 15 | \FAILURE 16 | Never fails. 17 | 18 | \SEEALSO 19 | basic_rewrites, extend_basic_convs, set_basic_rewrites. 20 | 21 | \ENDDOC 22 | -------------------------------------------------------------------------------- /Help/load_path.hlp: -------------------------------------------------------------------------------- 1 | \DOC load_path 2 | 3 | \TYPE {load_path : string list ref} 4 | 5 | \SYNOPSIS 6 | Path where HOL Light tries to find files to load. 7 | 8 | \DESCRIBE 9 | The reference variable {load_path} gives a list of directories. When HOL loads 10 | files with {loadt}, it will try these places in order on all non-absolute 11 | filenames. An initial dollar sign {$} in each path is interpreted as a 12 | reference to the current setting of {hol_dir}. To get an actual {$} character 13 | at the start of the filename, use two dollar signs {$$}. 14 | 15 | \FAILURE 16 | Not applicable. 17 | 18 | \SEEALSO 19 | file_on_path, help_path, hol_dir, hol_expand_directory, load_on_path, loads, 20 | loadt, needs. 21 | 22 | \ENDDOC 23 | -------------------------------------------------------------------------------- /Help/mk_comb.hlp: -------------------------------------------------------------------------------- 1 | \DOC mk_comb 2 | 3 | \TYPE {mk_comb : term * term -> term} 4 | 5 | \SYNOPSIS 6 | Constructs a combination. 7 | 8 | \DESCRIBE 9 | Given two terms {f} and {x}, the call {mk_comb(f,x)} returns the combination or 10 | application {f x}. It is necessary that {f} has a function type with domain 11 | type the same as {x}'s type. 12 | 13 | \FAILURE 14 | Fails if the types of the terms are not compatible as specified above. 15 | 16 | \EXAMPLE 17 | { 18 | # mk_comb(`SUC`,`0`);; 19 | val it : term = `SUC 0` 20 | 21 | # mk_comb(`SUC`,`T`);; 22 | Exception: Failure "mk_comb: types do not agree". 23 | } 24 | 25 | \SEEALSO 26 | dest_comb, is_comb, list_mk_comb, list_mk_icomb, mk_icomb. 27 | 28 | \ENDDOC 29 | -------------------------------------------------------------------------------- /Help/parse_preterm.hlp: -------------------------------------------------------------------------------- 1 | \DOC parse_preterm 2 | 3 | \TYPE {parse_preterm : lexcode list -> preterm * lexcode list} 4 | 5 | \SYNOPSIS 6 | Parses a preterm. 7 | 8 | \DESCRIBE 9 | The call {parse_preterm t}, where {t} is a list of lexical tokens (as produced 10 | by {lex}), parses the tokens and returns a preterm as well as the unparsed 11 | tokens. 12 | 13 | \FAILURE 14 | Fails if there is a syntax error in the token list. 15 | 16 | \USES 17 | This is mostly an internal function; pretypes and preterms are used as an 18 | intermediate representation for typechecking and overload resolution and are 19 | not normally of concern to users. 20 | 21 | \SEEALSO 22 | lex, parse_pretype, parse_term, parse_type. 23 | 24 | \ENDDOC 25 | -------------------------------------------------------------------------------- /Help/parse_pretype.hlp: -------------------------------------------------------------------------------- 1 | \DOC parse_pretype 2 | 3 | \TYPE {parse_pretype : lexcode list -> pretype * lexcode list} 4 | 5 | \SYNOPSIS 6 | Parses a pretype. 7 | 8 | \DESCRIBE 9 | The call {parse_pretype t}, where {t} is a list of lexical tokens (as produced 10 | by {lex}), parses the tokens and returns a pretype as well as the unparsed 11 | tokens. 12 | 13 | \FAILURE 14 | Fails if there is a syntax error in the token list. 15 | 16 | \USES 17 | This is mostly an internal function; pretypes and preterms are used as an 18 | intermediate representation for typechecking and overload resolution and are 19 | not normally of concern to users. 20 | 21 | \SEEALSO 22 | lex, parse_preterm, parse_term, parse_type. 23 | 24 | \ENDDOC 25 | -------------------------------------------------------------------------------- /Help/pp_print_fpf.hlp: -------------------------------------------------------------------------------- 1 | \DOC pp_print_fpf 2 | 3 | \TYPE {pp_print_fpf : Format.formatter -> ('a, 'b) func -> unit} 4 | 5 | \SYNOPSIS 6 | Print a finite partial function to a formatter. 7 | 8 | \DESCRIBE 9 | This prints a finite partial function but only as a trivial string `{}', 10 | to the given formatter. Installed automatically at the top level and probably 11 | not useful for most users. 12 | 13 | \FAILURE 14 | Never fails. 15 | 16 | \COMMENTS 17 | The usual case where the formatter is the standard output is {print_fpf}. 18 | 19 | \SEEALSO 20 | |->, |=>, apply, applyd, choose, combine, defined, dom, foldl, foldr, 21 | graph, is_undefined, mapf, print_fpf, ran, tryapplyd, undefine, undefined. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Help/prefixes.hlp: -------------------------------------------------------------------------------- 1 | \DOC prefixes 2 | 3 | \TYPE {prefixes : unit -> string list} 4 | 5 | \SYNOPSIS 6 | Certain identifiers {c} have prefix status, meaning that combinations of the 7 | form {c f x} will be parsed as {c (f x)} rather than the usual {(c f) x}. The 8 | call {prefixes()} returns the list of all such identifiers. 9 | 10 | \FAILURE 11 | Never fails. 12 | 13 | \EXAMPLE 14 | In the default HOL state: 15 | { 16 | # prefixes();; 17 | val it : string list = ["~"; "--"; "mod"] 18 | } 19 | This explains, for example, why `{~ ~ p}' parses as `{~(~p)}' rather than 20 | parsing as `{(~ ~) p}' and generating a typechecking error. 21 | 22 | \SEEALSO 23 | is_prefix, parse_as_prefix, unparse_as_prefix. 24 | 25 | \ENDDOC 26 | -------------------------------------------------------------------------------- /bignum_num.ml: -------------------------------------------------------------------------------- 1 | (* ------------------------------------------------------------------------- *) 2 | (* Load in the bignum library. *) 3 | (* ------------------------------------------------------------------------- *) 4 | 5 | include Num;; 6 | 7 | let num = Num.num_of_int;; 8 | 9 | module NumExt = struct 10 | let numdom (r:num):num * num = 11 | let r' = Ratio.normalize_ratio (ratio_of_num r) in 12 | num_of_big_int(Ratio.numerator_ratio r'), 13 | num_of_big_int(Ratio.denominator_ratio r');; 14 | 15 | let gcd_num (n1:num) (n2:num): num = 16 | num_of_big_int(Big_int.gcd_big_int (big_int_of_num n1) (big_int_of_num n2));; 17 | end;; 18 | 19 | include NumExt;; 20 | -------------------------------------------------------------------------------- /Help/COMB_CONV.hlp: -------------------------------------------------------------------------------- 1 | \DOC COMB_CONV 2 | 3 | \TYPE {COMB_CONV : conv -> conv} 4 | 5 | \SYNOPSIS 6 | Applies a conversion to the two sides of an application. 7 | 8 | \DESCRIBE 9 | If {c} is a conversion such that {c `f`} returns {|- f = f'} and 10 | {c `x`} returns {|- x = x'}, then {COMB_CONV c `f x`} returns 11 | {|- f x = f' x'}. That is, the conversion {c} is applied to the two 12 | immediate subterms. 13 | 14 | \FAILURE 15 | Never fails when applied to the initial conversion. On application to the term, 16 | it fails if conversion given as the argument does, or if the theorem returned 17 | by it is inappropriate. 18 | 19 | \SEEALSO 20 | BINOP_CONV, BINOP2_CONV, COMB2_CONV, LAND_CONV, RAND_CONV, RATOR_CONV 21 | 22 | \ENDDOC 23 | -------------------------------------------------------------------------------- /Help/NOT_ELIM.hlp: -------------------------------------------------------------------------------- 1 | \DOC NOT_ELIM 2 | 3 | \TYPE {NOT_ELIM : thm -> thm} 4 | 5 | \SYNOPSIS 6 | Transforms {|- ~t} into {|- t ==> F}. 7 | 8 | \KEYWORDS 9 | rule, implication, negation. 10 | 11 | \DESCRIBE 12 | When applied to a theorem {A |- ~t}, the inference rule {NOT_ELIM} returns the 13 | theorem {A |- t ==> F}. 14 | { 15 | A |- ~t 16 | -------------- NOT_ELIM 17 | A |- t ==> F 18 | } 19 | 20 | \FAILURE 21 | Fails unless the theorem has a negated conclusion. 22 | 23 | \EXAMPLE 24 | { 25 | # let th = UNDISCH(TAUT `p ==> ~ ~p`);; 26 | val th : thm = p |- ~ ~p 27 | 28 | # NOT_ELIM th;; 29 | val it : thm = p |- ~p ==> F 30 | } 31 | 32 | \SEEALSO 33 | EQF_ELIM, EQF_INTRO, NOT_INTRO. 34 | 35 | \ENDDOC 36 | -------------------------------------------------------------------------------- /Help/TAC_PROOF.hlp: -------------------------------------------------------------------------------- 1 | \DOC TAC_PROOF 2 | 3 | \TYPE {TAC_PROOF : goal * tactic -> thm} 4 | 5 | \SYNOPSIS 6 | Attempts to prove a goal using a given tactic. 7 | 8 | \DESCRIBE 9 | When applied to a goal-tactic pair {(A ?- t,tac)}, the {TAC_PROOF} function 10 | attempts to prove the goal {A ?- t}, using the tactic {tac}. If it succeeds, it 11 | returns the theorem {A' |- t} corresponding to the goal, where the assumption 12 | list {A'} may be a proper superset of {A} unless the tactic is valid; there 13 | is no inbuilt validity checking. 14 | 15 | \FAILURE 16 | Fails unless the goal has hypotheses and conclusions all of type {bool}, 17 | and the tactic can solve the goal. 18 | 19 | \SEEALSO 20 | prove, VALID. 21 | 22 | \ENDDOC 23 | -------------------------------------------------------------------------------- /Help/THEN_TCL.hlp: -------------------------------------------------------------------------------- 1 | \DOC THEN_TCL 2 | 3 | \TYPE {(THEN_TCL) : thm_tactical -> thm_tactical -> thm_tactical} 4 | 5 | \SYNOPSIS 6 | Composes two theorem-tacticals. 7 | 8 | \KEYWORDS 9 | theorem-tactical. 10 | 11 | \DESCRIBE 12 | If {ttl1} and {ttl2} are two theorem-tacticals, {ttl1 THEN_TCL ttl2} is 13 | a theorem-tactical which composes their effect; that is, if: 14 | { 15 | ttl1 ttac th1 = ttac th2 16 | } 17 | \noindent and 18 | { 19 | ttl2 ttac th2 = ttac th3 20 | } 21 | \noindent then 22 | { 23 | (ttl1 THEN_TCL ttl2) ttac th1 = ttac th3 24 | } 25 | 26 | \FAILURE 27 | The application of {THEN_TCL} to a pair of theorem-tacticals never fails. 28 | 29 | \SEEALSO 30 | EVERY_TCL, FIRST_TCL, ORELSE_TCL. 31 | 32 | \ENDDOC 33 | -------------------------------------------------------------------------------- /Help/dest_binary.hlp: -------------------------------------------------------------------------------- 1 | \DOC dest_binary 2 | 3 | \TYPE {dest_binary : string -> term -> term * term} 4 | 5 | \SYNOPSIS 6 | Breaks apart an instance of a binary operator with given name. 7 | 8 | \DESCRIBE 9 | The call {dest_binary s tm} will, if {tm} is a binary operator application 10 | {(op l) r} where {op} is a constant with name {s}, return the two arguments to 11 | which it is applied as a pair {l,r}. Otherwise, it fails. Note that {op} is 12 | required to be a constant. 13 | 14 | \FAILURE 15 | Never fails. 16 | 17 | \EXAMPLE 18 | This one succeeds: 19 | { 20 | # dest_binary "+" `1 + 2`;; 21 | val it : term * term = (`1`, `2`) 22 | } 23 | 24 | \SEEALSO 25 | dest_binop, is_binary, is_comb, mk_binary. 26 | 27 | \ENDDOC 28 | -------------------------------------------------------------------------------- /Help/dest_type.hlp: -------------------------------------------------------------------------------- 1 | \DOC dest_type 2 | 3 | \TYPE {dest_type : hol_type -> string * hol_type list} 4 | 5 | \SYNOPSIS 6 | Breaks apart a type (other than a variable type). 7 | 8 | \DESCRIBE 9 | {dest_type(`:(ty1,...,tyn)op`)} returns {("op",[`:ty1`;...;`:tyn`])}. 10 | 11 | \FAILURE 12 | Fails with {dest_type} if the type is a type variable. 13 | 14 | \EXAMPLE 15 | { 16 | # dest_type `:bool`;; 17 | val it : string * hol_type list = ("bool", []) 18 | 19 | # dest_type `:(bool)list`;; 20 | val it : string * hol_type list = ("list", [`:bool`]) 21 | 22 | # dest_type `:num -> bool`;; 23 | val it : string * hol_type list = ("fun", [`:num`; `:bool`]) 24 | } 25 | 26 | \SEEALSO 27 | mk_type, dest_vartype. 28 | 29 | \ENDDOC 30 | -------------------------------------------------------------------------------- /Help/list_mk_disj.hlp: -------------------------------------------------------------------------------- 1 | \DOC list_mk_disj 2 | 3 | \TYPE {list_mk_disj : term list -> term} 4 | 5 | \SYNOPSIS 6 | Constructs the disjunction of a list of terms. 7 | 8 | \DESCRIBE 9 | {list_mk_disj([`t1`;...;`tn`])} returns {`t1 \/ ... \/ tn`}. 10 | 11 | \FAILURE 12 | Fails with {list_mk_disj} if the list is empty or if the list has more than 13 | one element, one or more of which are not of type {`:bool`}. 14 | 15 | \EXAMPLE 16 | { 17 | # list_mk_disj [`T`;`F`;`T`];; 18 | val it : term = `T \/ F \/ T` 19 | 20 | # list_mk_disj [`T`;`1`;`F`];; 21 | Exception: Failure "mk_binary". 22 | 23 | # list_mk_disj [`1`];; 24 | val it : term = `1` 25 | } 26 | 27 | \SEEALSO 28 | disjuncts, is_disj, mk_disj. 29 | 30 | \ENDDOC 31 | -------------------------------------------------------------------------------- /Help/nsplit.hlp: -------------------------------------------------------------------------------- 1 | \DOC nsplit 2 | 3 | \TYPE {nsplit : ('a -> 'b * 'a) -> 'c list -> 'a -> 'b list * 'a} 4 | 5 | \SYNOPSIS 6 | Applies a destructor in right-associative mode a specified number of times. 7 | 8 | \DESCRIBE 9 | If {d} is an inverse to a binary constructor {f}, then 10 | { 11 | nsplit d l (f(x1,f(x2,...f(xn,y)))) 12 | } 13 | \noindent where the list {l} has length {k}, returns 14 | { 15 | ([x1;...;xk],f(x(k+1),...f(xn,y)) 16 | } 17 | 18 | \FAILURE 19 | Never fails. 20 | 21 | \EXAMPLE 22 | { 23 | # nsplit dest_conj [1;2;3] `a /\ b /\ c /\ d /\ e /\ f`;; 24 | val it : term list * term = ([`a`; `b`; `c`], `d /\ e /\ f`) 25 | } 26 | 27 | \SEEALSO 28 | splitlist, rev_splitlist, striplist. 29 | 30 | \ENDDOC 31 | -------------------------------------------------------------------------------- /Help/top_goal.hlp: -------------------------------------------------------------------------------- 1 | \DOC top_goal 2 | 3 | \TYPE {top_goal : unit -> term list * term} 4 | 5 | \SYNOPSIS 6 | Returns the current goal of the subgoal package. 7 | 8 | \DESCRIBE 9 | The function {top_goal} is part of the subgoal package. It returns the top goal 10 | of the goal stack in the current proof state. For a description of the subgoal 11 | package, see {set_goal}. 12 | 13 | \FAILURE 14 | A call to {top_goal} will fail if there are no unproven goals. This could be 15 | because no goal has been set using {set_goal} or because the last goal set has 16 | been completely proved. 17 | 18 | \USES 19 | Examining the proof state after a proof fails. 20 | 21 | \SEEALSO 22 | b, e, g, p, r, set_goal, top_thm. 23 | 24 | \ENDDOC 25 | -------------------------------------------------------------------------------- /Help/dest_iff.hlp: -------------------------------------------------------------------------------- 1 | \DOC dest_iff 2 | 3 | \TYPE {dest_iff : term -> term * term} 4 | 5 | \SYNOPSIS 6 | Term destructor for logical equivalence. 7 | 8 | \DESCRIBE 9 | {dest_iff(`t1 <=> t2`)} returns {(`t1`,`t2`)}. 10 | 11 | \FAILURE 12 | Fails with if term is not a logical equivalence, i.e. an equation between terms 13 | of Boolean type. 14 | 15 | \EXAMPLE 16 | { 17 | # dest_iff `x = y <=> y = 1`;; 18 | val it : term * term = (`x = y`, `y = 1`) 19 | } 20 | 21 | \COMMENTS 22 | The function {dest_eq} has the same effect, but the present function checks 23 | that the types of the two sides are indeed Boolean, whereas {dest_eq} will 24 | break apart any equation. 25 | 26 | \SEEALSO 27 | dest_eq, is_iff, mk_iff. 28 | 29 | \ENDDOC 30 | -------------------------------------------------------------------------------- /Help/is_realintconst.hlp: -------------------------------------------------------------------------------- 1 | \DOC is_realintconst 2 | 3 | \TYPE {is_realintconst : term -> bool} 4 | 5 | \SYNOPSIS 6 | Tests if a term is an integer literal of type {:real}. 7 | 8 | \DESCRIBE 9 | The call {is_realintconst t} tests whether the term {t} is a canonical integer 10 | literal of type {:real}, i.e. either `{&n}' for a numeral {n} or `{-- &n}' for 11 | a nonzero numeral {n}. If so it returns {true}, otherwise {false}. 12 | 13 | \FAILURE 14 | Never fails. 15 | 16 | \EXAMPLE 17 | { 18 | # is_realintconst `-- &3 :real`;; 19 | val it : bool = true 20 | # is_realintconst `&1 :int`;; 21 | val it : bool = false 22 | } 23 | 24 | \SEEALSO 25 | dest_realintconst, is_intconst, is_ratconst, mk_realintconst. 26 | 27 | \ENDDOC 28 | -------------------------------------------------------------------------------- /Help/strip_comb.hlp: -------------------------------------------------------------------------------- 1 | \DOC strip_comb 2 | 3 | \TYPE {strip_comb : term -> term * term list} 4 | 5 | \SYNOPSIS 6 | Iteratively breaks apart combinations (function applications). 7 | 8 | \DESCRIBE 9 | {strip_comb `t t1 ... tn`} returns {(`t`,[`t1`;...;`tn`])}. Note that 10 | { 11 | strip_comb(list_mk_comb(`t`,[`t1`;...;`tn`])) 12 | } 13 | \noindent will not return {(`t`,[`t1`;...;`tn`])} if {t} is a combination. 14 | 15 | \FAILURE 16 | Never fails. 17 | 18 | \EXAMPLE 19 | { 20 | # strip_comb `x /\ y`;; 21 | val it : term * term list = (`(/\)`, [`x`; `y`]) 22 | 23 | # strip_comb `T`;; 24 | val it : term * term list = (`T`, []) 25 | } 26 | 27 | \SEEALSO 28 | dest_comb, list_mk_comb, splitlist, striplist. 29 | 30 | \ENDDOC 31 | -------------------------------------------------------------------------------- /Unity/make.ml: -------------------------------------------------------------------------------- 1 | (*-------------------------------------------------------------------------*) 2 | (* 3 | File: unity 4 | Description: 5 | 6 | This file loads and opens the HOL Light theory unity, called HOL_UNITY 7 | 8 | Author: Flemming Andersen 9 | Date: November 17, 2003 10 | *) 11 | (*-------------------------------------------------------------------------*) 12 | 13 | loadt "Examples/hol88.ml";; 14 | 15 | loadt "Unity/aux_definitions.ml";; 16 | loadt "Unity/mk_state_logic.ml";; 17 | loadt "Unity/mk_unless.ml";; 18 | loadt "Unity/mk_ensures.ml";; 19 | loadt "Unity/mk_gen_induct.ml";; 20 | loadt "Unity/mk_leadsto.ml";; 21 | loadt "Unity/mk_comp_unity.ml";; 22 | loadt "Unity/mk_unity_prog.ml";; 23 | -------------------------------------------------------------------------------- /Help/new_constant.hlp: -------------------------------------------------------------------------------- 1 | \DOC new_constant 2 | 3 | \TYPE {new_constant : string * hol_type -> unit} 4 | 5 | \SYNOPSIS 6 | Declares a new constant. 7 | 8 | \DESCRIBE 9 | A call {new_constant("c",`:ty`)} makes {c} a constant with most general type 10 | {ty}. 11 | 12 | \FAILURE 13 | Fails if there is already a constant of that name in the current theory. 14 | 15 | \EXAMPLE 16 | { 17 | 18 | #new_constant("graham's_number",`:num`);; 19 | val it : unit = () 20 | } 21 | 22 | \USES 23 | Can be useful for declaring some arbitrary parameter, but more usually a 24 | prelude to some new axioms about the constant introduced. Take care when using 25 | {new_axiom}! 26 | 27 | \SEEALSO 28 | constants, new_axiom, new_definition. 29 | 30 | \ENDDOC 31 | -------------------------------------------------------------------------------- /Help/unparse_as_infix.hlp: -------------------------------------------------------------------------------- 1 | \DOC unparse_as_infix 2 | 3 | \TYPE {unparse_as_infix : string -> unit} 4 | 5 | \SYNOPSIS 6 | Removes string from the list of infix operators. 7 | 8 | \DESCRIBE 9 | Certain identifiers are treated as infix operators with a given precedence and 10 | associativity (left or right). The call {unparse_as_infix "op"} removes {op} 11 | from the list of infix identifiers, if it was indeed there. 12 | 13 | \FAILURE 14 | Never fails, even if the given string did not originally have infix status. 15 | 16 | \COMMENTS 17 | Take care with applying this to some of the built-in operators, or parsing may 18 | fail in existing libraries. 19 | 20 | \SEEALSO 21 | get_infix_status, infixes, parse_as_infix. 22 | 23 | \ENDDOC 24 | -------------------------------------------------------------------------------- /Help/CHANGED_TAC.hlp: -------------------------------------------------------------------------------- 1 | \DOC CHANGED_TAC 2 | 3 | \TYPE {CHANGED_TAC : tactic -> tactic} 4 | 5 | \SYNOPSIS 6 | Makes a tactic fail if it has no effect. 7 | 8 | \KEYWORDS 9 | tactical. 10 | 11 | \DESCRIBE 12 | When applied to a tactic {t}, the tactical {CHANGED_TAC} gives a new tactic 13 | which is the same as {t} if that has any effect, and otherwise fails. 14 | 15 | \FAILURE 16 | The application of {CHANGED_TAC} to a tactic never fails. The resulting 17 | tactic fails if the basic tactic either fails or has no effect. 18 | 19 | \USES 20 | Occasionally useful in controlling complicated tactic compositions. Also 21 | sometimes convenient just to check that a step did indeed modify a goal. 22 | 23 | \SEEALSO 24 | TRY, VALID. 25 | 26 | \ENDDOC 27 | --------------------------------------------------------------------------------