--- a/src/Pure/IsaMakefile Fri Jun 04 22:12:33 1999 +0200
+++ b/src/Pure/IsaMakefile Sat Jun 05 20:27:53 1999 +0200
@@ -29,24 +29,24 @@
General/pretty.ML General/scan.ML General/seq.ML General/source.ML \
General/symbol.ML General/table.ML General/url.ML Interface/ROOT.ML \
Interface/isamode.ML Interface/proof_general.ML Isar/ROOT.ML \
- Isar/args.ML Isar/attrib.ML Isar/calculation.ML Isar/comment.ML \
- Isar/isar.ML Isar/isar_cmd.ML Isar/isar_syn.ML Isar/isar_thy.ML \
- Isar/method.ML Isar/outer_lex.ML Isar/outer_parse.ML \
- Isar/outer_syntax.ML Isar/proof.ML Isar/proof_context.ML \
- Isar/proof_data.ML Isar/proof_history.ML Isar/session.ML \
- Isar/toplevel.ML ML-Systems/mlworks.ML ML-Systems/polyml.ML \
- ML-Systems/smlnj-0.93.ML ML-Systems/smlnj.ML ROOT.ML Syntax/ROOT.ML \
- Syntax/ast.ML Syntax/lexicon.ML Syntax/mixfix.ML Syntax/parser.ML \
- Syntax/printer.ML Syntax/syn_ext.ML Syntax/syn_trans.ML \
- Syntax/syntax.ML Syntax/token_trans.ML Syntax/type_ext.ML \
- Thy/ROOT.ML Thy/browser_info.ML Thy/html.ML Thy/present.ML \
- Thy/thm_database.ML Thy/thy_info.ML Thy/thy_load.ML Thy/thy_parse.ML \
- Thy/thy_scan.ML Thy/thy_syn.ML axclass.ML basis.ML context.ML \
- deriv.ML display.ML drule.ML envir.ML goals.ML install_pp.ML \
- library.ML locale.ML logic.ML net.ML object_logic.ML pattern.ML \
- pure.ML pure_thy.ML search.ML sign.ML sorts.ML tactic.ML tctical.ML \
- term.ML theory.ML theory_data.ML thm.ML type.ML type_infer.ML \
- unify.ML
+ Isar/args.ML Isar/attrib.ML Isar/auto_bind.ML Isar/calculation.ML \
+ Isar/comment.ML Isar/isar.ML Isar/isar_cmd.ML Isar/isar_syn.ML \
+ Isar/isar_thy.ML Isar/method.ML Isar/outer_lex.ML \
+ Isar/outer_parse.ML Isar/outer_syntax.ML Isar/proof.ML \
+ Isar/proof_context.ML Isar/proof_data.ML Isar/proof_history.ML \
+ Isar/session.ML Isar/toplevel.ML ML-Systems/mlworks.ML \
+ ML-Systems/polyml.ML ML-Systems/smlnj-0.93.ML ML-Systems/smlnj.ML \
+ ROOT.ML Syntax/ROOT.ML Syntax/ast.ML Syntax/lexicon.ML \
+ Syntax/mixfix.ML Syntax/parser.ML Syntax/printer.ML \
+ Syntax/syn_ext.ML Syntax/syn_trans.ML Syntax/syntax.ML \
+ Syntax/token_trans.ML Syntax/type_ext.ML Thy/ROOT.ML \
+ Thy/browser_info.ML Thy/html.ML Thy/present.ML Thy/thm_database.ML \
+ Thy/thy_info.ML Thy/thy_load.ML Thy/thy_parse.ML Thy/thy_scan.ML \
+ Thy/thy_syn.ML axclass.ML basis.ML context.ML deriv.ML display.ML \
+ drule.ML envir.ML goals.ML install_pp.ML library.ML locale.ML \
+ logic.ML net.ML pattern.ML pure.ML pure_thy.ML search.ML sign.ML \
+ sorts.ML tactic.ML tctical.ML term.ML theory.ML theory_data.ML \
+ thm.ML type.ML type_infer.ML unify.ML
@./mk
--- a/src/Pure/Isar/ROOT.ML Fri Jun 04 22:12:33 1999 +0200
+++ b/src/Pure/Isar/ROOT.ML Sat Jun 05 20:27:53 1999 +0200
@@ -6,6 +6,7 @@
*)
(*proof engine*)
+use "auto_bind.ML";
use "proof_context.ML";
use "proof.ML";
use "proof_data.ML";
@@ -37,6 +38,7 @@
structure PureIsar =
struct
+ structure AutoBind = AutoBind;
structure ProofContext = ProofContext;
structure Proof = Proof;
structure ProofHistory = ProofHistory;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Isar/auto_bind.ML Sat Jun 05 20:27:53 1999 +0200
@@ -0,0 +1,36 @@
+(* Title: Pure/Isar/auto_bind.ML
+ ID: $Id$
+ Author: Markus Wenzel, TU Muenchen
+
+Automatic term bindings -- logic specific patterns.
+
+TODO:
+ - logic specific theory data instead of hardwired operations (!!);
+*)
+
+signature AUTO_BIND =
+sig
+ val goal: term -> (indexname * term) list
+ val facts: term list -> (indexname * term) list
+end;
+
+structure AutoBind: AUTO_BIND =
+struct
+
+val thesisN = "thesis";
+fun thesis_name sfx = (Syntax.binding (thesisN ^ sfx), 0);
+
+fun goal prop =
+ let val concl = Logic.strip_imp_concl prop in
+ [(thesis_name "_prop", prop), (thesis_name "_concl", concl)] @
+ (case concl of Const ("Trueprop", _) $ t => [(thesis_name "", t)] | _ => [])
+ end;
+
+fun facts [] = []
+ | facts props =
+ (case Logic.strip_imp_concl (Library.last_elem props) of
+ Const ("Trueprop", _) $ (_ $ t) => [(Syntax.dddot_indexname, t)]
+ | _ => []);
+
+
+end;
--- a/src/Pure/ROOT.ML Fri Jun 04 22:12:33 1999 +0200
+++ b/src/Pure/ROOT.ML Sat Jun 05 20:27:53 1999 +0200
@@ -39,7 +39,6 @@
use "theory.ML";
use "theory_data.ML";
use "context.ML";
-use "object_logic.ML";
use "thm.ML";
use "display.ML";
use "pure_thy.ML";
--- a/src/Pure/object_logic.ML Fri Jun 04 22:12:33 1999 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,37 +0,0 @@
-(* Title: Pure/object_logic.ML
- ID: $Id$
- Author: Markus Wenzel, TU Muenchen
-
-Object logic specific operations.
-
-TODO:
- - theory data instead of hardwired operations (!!);
-*)
-
-signature OBJECT_LOGIC =
-sig
- val dest_statement: string * term -> (string * term) list
- val dest_main_statement: term -> term
- val setup: (theory -> theory) list
-end;
-
-structure ObjectLogic: OBJECT_LOGIC =
-struct
-
-
-fun dest_statement (name, prop) =
- let val concl = Logic.strip_imp_concl prop in
- [(name ^ "_prop", prop), (name ^ "_concl", concl)] @
- (case concl of Const ("Trueprop", _) $ t => [(name, t)] | _ => [])
- end;
-
-fun dest_main_statement t =
- (case Logic.strip_imp_concl t of
- _ $ t => t
- | _ => raise TERM ("dest_main", [t]));
-
-
-(* FIXME *)
-val setup = [];
-
-end;