]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/matita/matitaInterpreter.ml
snapshot, notably:
[helm.git] / helm / matita / matitaInterpreter.ml
index 78d30c4cd062a16898dc16eff268fce257940a91..640b7406ee17f925832f864d410f7e95fb5b233c 100644 (file)
@@ -64,9 +64,12 @@ class virtual interpreterState =
     (* static values, shared by all states inheriting this class *)
   let loc = ref None in
   let history = ref [] in
-  fun ~(console: MatitaTypes.console) ->
+  fun ~(console: #MatitaTypes.console) ->
   object (self)
 
+    val dbd = MatitaMisc.dbd_instance ()
+    val currentProof = MatitaProof.instance ()
+
       (** eval a toplevel phrase in the current state and return the new state
       *)
     method parsePhrase s =
@@ -103,10 +106,8 @@ class virtual interpreterState =
   (** Implements phrases that should be accepted in all states *)
 class sharedState
   ~(disambiguator: MatitaTypes.disambiguator)
-  ~(currentProof: MatitaTypes.currentProof)
-  ~(console: MatitaTypes.console)
+  ~(console: #MatitaTypes.console)
   ?(mathViewer: MatitaTypes.mathViewer option)
-  ~(dbd: Mysql.dbd)
   ()
 =
   object (self)
@@ -307,7 +308,7 @@ let
     * - save object to disk in xml format
     * - register uri to the getter 
     * - save universe file *)
-let add_constant_to_world ~(console:MatitaTypes.console)
+let add_constant_to_world ~(console: #MatitaTypes.console)
   ~dbd ~uri ?body ~ty ?(params = []) ?(attrs = []) ~ugraph ()
 =
   let suri = UriManager.string_of_uri uri in
@@ -324,7 +325,7 @@ let add_constant_to_world ~(console:MatitaTypes.console)
     console#echo_message (sprintf "%s constant defined" suri)
   end
 
-let add_inductive_def_to_world ~(console:MatitaTypes.console)
+let add_inductive_def_to_world ~(console: #MatitaTypes.console)
   ~dbd ~uri ~indTypes ?(params = []) ?(leftno = 0) ?(attrs = []) ~ugraph ()
 =
   let suri = UriManager.string_of_uri uri in
@@ -360,15 +361,11 @@ let add_inductive_def_to_world ~(console:MatitaTypes.console)
   (** Implements phrases that should be accepted only in Command state *)
 class commandState
   ~(disambiguator: MatitaTypes.disambiguator)
-  ~(currentProof: MatitaTypes.currentProof)
-  ~(console: MatitaTypes.console)
+  ~(console: #MatitaTypes.console)
   ?mathViewer
-  ~(dbd: Mysql.dbd)
   ()
 =
-  let shared =
-    new sharedState ~disambiguator ~currentProof ~console ?mathViewer ~dbd ()
-  in
+  let shared = new sharedState ~disambiguator ~console ?mathViewer () in
   object (self)
     inherit interpreterState ~console
 
@@ -519,67 +516,64 @@ let namer_of names =
   * tacticals *)
 class proofState
   ~(disambiguator: MatitaTypes.disambiguator)
-  ~(currentProof: MatitaTypes.currentProof)
-  ~(console: MatitaTypes.console)
+  ~(console: #MatitaTypes.console)
   ?mathViewer
-  ~(dbd: Mysql.dbd)
   ()
 =
-  let disambiguate ast =
-    let (_, _, term, _) =
-      MatitaCicMisc.disambiguate ~disambiguator ~currentProof ast
-    in
-    term
-  in
-    (** tactic AST -> ProofEngineTypes.tactic *)
-  let rec lookup_tactic = function
-    | TacticAst.LocatedTactic (_, tactic) -> lookup_tactic tactic
-    | TacticAst.Intros (_, names) ->  (* TODO Zack implement intros length *)
-        PrimitiveTactics.intros_tac ~mk_fresh_name_callback:(namer_of names) ()
-    | TacticAst.Reflexivity -> Tactics.reflexivity
-    | TacticAst.Assumption -> Tactics.assumption
-    | TacticAst.Contradiction -> Tactics.contradiction
-    | TacticAst.Exists -> Tactics.exists
-    | TacticAst.Fourier -> Tactics.fourier
-    | TacticAst.Left -> Tactics.left
-    | TacticAst.Right -> Tactics.right
-    | TacticAst.Ring -> Tactics.ring
-    | TacticAst.Split -> Tactics.split
-    | TacticAst.Symmetry -> Tactics.symmetry
-    | TacticAst.Transitivity term -> Tactics.transitivity (disambiguate term)
-    | TacticAst.Apply term -> Tactics.apply (disambiguate term)
-    | TacticAst.Absurd term -> Tactics.absurd (disambiguate term)
-    | TacticAst.Exact term -> Tactics.exact (disambiguate term)
-    | TacticAst.Cut term -> Tactics.cut (disambiguate term)
-    | TacticAst.Elim (term, _) -> (* TODO Zack implement "using" argument *)
-        Tactics.elim_intros_simpl (disambiguate term)
-    | TacticAst.ElimType term -> Tactics.elim_type (disambiguate term)
-    | TacticAst.Replace (what, with_what) ->
-        Tactics.replace ~what:(disambiguate what)
-          ~with_what:(disambiguate with_what)
-    | TacticAst.Auto -> Tactics.auto_new ~dbd
-  (*
-    (* TODO Zack a lot more of tactics to be implemented here ... *)
-    | TacticAst.Change of 'term * 'term * 'ident option
-    | TacticAst.Change_pattern of 'term pattern * 'term * 'ident option
-    | TacticAst.Decompose of 'ident * 'ident list
-    | TacticAst.Discriminate of 'ident
-    | TacticAst.Fold of reduction_kind * 'term
-    | TacticAst.Injection of 'ident
-    | TacticAst.LetIn of 'term * 'ident
-    | TacticAst.Reduce of reduction_kind * 'term pattern * 'ident option
-    | TacticAst.Replace_pattern of 'term pattern * 'term
-    | TacticAst.Rewrite of direction * 'term * 'ident option
-  *)
-    | _ ->
-        MatitaTypes.not_implemented "some tactic"
-  in
-  let shared =
-    new sharedState ~disambiguator ~currentProof ~console ?mathViewer ~dbd ()
-  in
+  let shared = new sharedState ~disambiguator ~console ?mathViewer () in
   object (self)
     inherit interpreterState ~console
 
+    method private disambiguate ast =
+      let (_, _, term, _) =
+        MatitaCicMisc.disambiguate ~disambiguator ~currentProof ast
+      in
+      term
+
+    (** tactic AST -> ProofEngineTypes.tactic *)
+    method private lookup_tactic = function
+      | TacticAst.LocatedTactic (_, tactic) -> self#lookup_tactic tactic
+      | TacticAst.Intros (_, names) ->  (* TODO Zack implement intros length *)
+          PrimitiveTactics.intros_tac ~mk_fresh_name_callback:(namer_of names)
+            ()
+      | TacticAst.Reflexivity -> Tactics.reflexivity
+      | TacticAst.Assumption -> Tactics.assumption
+      | TacticAst.Contradiction -> Tactics.contradiction
+      | TacticAst.Exists -> Tactics.exists
+      | TacticAst.Fourier -> Tactics.fourier
+      | TacticAst.Left -> Tactics.left
+      | TacticAst.Right -> Tactics.right
+      | TacticAst.Ring -> Tactics.ring
+      | TacticAst.Split -> Tactics.split
+      | TacticAst.Symmetry -> Tactics.symmetry
+      | TacticAst.Transitivity term ->
+          Tactics.transitivity (self#disambiguate term)
+      | TacticAst.Apply term -> Tactics.apply (self#disambiguate term)
+      | TacticAst.Absurd term -> Tactics.absurd (self#disambiguate term)
+      | TacticAst.Exact term -> Tactics.exact (self#disambiguate term)
+      | TacticAst.Cut term -> Tactics.cut (self#disambiguate term)
+      | TacticAst.Elim (term, _) -> (* TODO Zack implement "using" argument *)
+          Tactics.elim_intros_simpl (self#disambiguate term)
+      | TacticAst.ElimType term -> Tactics.elim_type (self#disambiguate term)
+      | TacticAst.Replace (what, with_what) ->
+          Tactics.replace ~what:(self#disambiguate what)
+            ~with_what:(self#disambiguate with_what)
+      | TacticAst.Auto -> Tactics.auto_new ~dbd
+    (*
+      (* TODO Zack a lot more of tactics to be implemented here ... *)
+      | TacticAst.Change of 'term * 'term * 'ident option
+      | TacticAst.Change_pattern of 'term pattern * 'term * 'ident option
+      | TacticAst.Decompose of 'ident * 'ident list
+      | TacticAst.Discriminate of 'ident
+      | TacticAst.Fold of reduction_kind * 'term
+      | TacticAst.Injection of 'ident
+      | TacticAst.LetIn of 'term * 'ident
+      | TacticAst.Reduce of reduction_kind * 'term pattern * 'ident option
+      | TacticAst.Replace_pattern of 'term pattern * 'term
+      | TacticAst.Rewrite of direction * 'term * 'ident option
+    *)
+      | _ -> MatitaTypes.not_implemented "some tactic"
+
     method evalTactical = function
       | TacticAst.LocatedTactical (_, tactical) -> self#evalTactical tactical
       | TacticAst.Command TacticAst.Abort ->
@@ -616,7 +610,7 @@ class proofState
           List.iter (fun t -> ignore (self#evalTactical t)) tacticals;
           New_state Proof
       | TacticAst.Tactic tactic_phrase ->
-          let tactic = lookup_tactic tactic_phrase in
+          let tactic = self#lookup_tactic tactic_phrase in
           currentProof#proof#apply_tactic tactic;
           New_state Proof
       | tactical -> shared#evalTactical tactical
@@ -624,18 +618,12 @@ class proofState
 
 class interpreter
   ~(disambiguator: MatitaTypes.disambiguator)
-  ~(currentProof: MatitaTypes.currentProof)
-  ~(console: MatitaTypes.console)
+  ~(console: #MatitaTypes.console)
   ?mathViewer
-  ~(dbd: Mysql.dbd)
   ()
 =
-  let commandState =
-    new commandState ~disambiguator ~currentProof ~console ?mathViewer ~dbd ()
-  in
-  let proofState =
-    new proofState ~disambiguator ~currentProof ~console ?mathViewer ~dbd ()
-  in
+  let commandState = new commandState ~disambiguator ~console ?mathViewer () in
+  let proofState = new proofState ~disambiguator ~console ?mathViewer () in
   object (self)
     val mutable state = commandState
 
@@ -661,3 +649,11 @@ class interpreter
     method evalAst ast = self#eval (fun () -> state#evalAst ast)
   end
 
+let interpreter
+  ~(disambiguator: MatitaTypes.disambiguator)
+  ~(console: #MatitaTypes.console)
+  ?mathViewer
+  ()
+=
+  new interpreter ~disambiguator ~console ?mathViewer ()
+