]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/matita/matitaInterpreter.ml
rebuilt against ocaml 3.08.3
[helm.git] / helm / matita / matitaInterpreter.ml
index 640b7406ee17f925832f864d410f7e95fb5b233c..e578a3eb039a44612b163b0f422b216c258ebbc3 100644 (file)
@@ -47,7 +47,7 @@ let uri name =
 *)
 
 let baseuri = lazy (ref ("cic:/matita/" ^ Helm_registry.get "matita.owner"))
-let basedir = ref ((Unix.getpwuid (Unix.getuid ())).Unix.pw_dir) ;;
+let basedir = lazy (ref (Helm_registry.get "matita.basedir"))
 
 let qualify name =
   let baseuri = !(Lazy.force baseuri) in
@@ -63,49 +63,52 @@ let split_obj = function
 class virtual interpreterState = 
     (* static values, shared by all states inheriting this class *)
   let loc = ref None in
-  let history = ref [] in
+  let last_item = ref None in
+  let evalAstCallback = ref None in
   fun ~(console: #MatitaTypes.console) ->
   object (self)
 
     val dbd = MatitaMisc.dbd_instance ()
     val currentProof = MatitaProof.instance ()
+    val disambiguator = MatitaDisambiguator.instance ()
 
       (** eval a toplevel phrase in the current state and return the new state
       *)
     method parsePhrase s =
-      match CicTextualParser2.parse_tactical s with
+      match disambiguator#parserr#parseTactical s with
       | (TacticAst.LocatedTactical (loc', tac)) as tactical ->
           loc := Some loc';
-          (match tac with (* update interpreter history *)
-          | TacticAst.Command (TacticAst.Qed None) ->
-              history := `Qed :: !history
-          | TacticAst.Command (TacticAst.Theorem (_, Some name, _, None)) ->
-              history := `Theorem name :: !history
-          | TacticAst.Command (TacticAst.Qed _)
-          | TacticAst.Command (TacticAst.Theorem _) -> assert false
-          | _ -> history := `Tactic :: !history);
           tactical
       | _ -> assert false
 
     method virtual evalTactical:
       (CicAst.term, string) TacticAst.tactical -> outcome
 
+    method private _evalTactical ast =
+      self#setLastItem None;
+      let res = self#evalTactical ast in
+      (match !evalAstCallback with Some f -> f ast | None -> ());
+      res
+
     method evalPhrase s =
-      debug_print (sprintf "evaluating '%s'" s);
-      self#evalTactical (self#parsePhrase (Stream.of_string s))
+      self#_evalTactical (self#parsePhrase (Stream.of_string s))
 
-    method evalAst ast = self#evalTactical ast
+    method evalAst ast = self#_evalTactical ast
 
     method endOffset =
       match !loc with
       | Some (start_pos, end_pos) -> end_pos.Lexing.pos_cnum
       | None -> failwith "MatitaInterpreter: no offset recorded"
 
+    method lastItem: script_item option = !last_item
+    method private setLastItem item = last_item := item
+
+    method setEvalAstCallback f = evalAstCallback := Some f
+
   end
 
   (** Implements phrases that should be accepted in all states *)
 class sharedState
-  ~(disambiguator: MatitaTypes.disambiguator)
   ~(console: #MatitaTypes.console)
   ?(mathViewer: MatitaTypes.mathViewer option)
   ()
@@ -128,17 +131,18 @@ class sharedState
             !(Lazy.force baseuri));
           Quiet
       | TacticAst.Command (TacticAst.Basedir (Some path)) ->
-          basedir := path;
+          Lazy.force basedir := path;
           console#echo_message (sprintf "base dir set to \"%s\"" path);
           Quiet
       | TacticAst.Command (TacticAst.Basedir None) ->
-          console#echo_message (sprintf "base dir is \"%s\"" !basedir);
+          console#echo_message (sprintf "base dir is \"%s\""
+            !(Lazy.force basedir));
           Quiet
       | TacticAst.Command (TacticAst.Check term) ->
-          let (_, _, term,ugraph) = 
+          let (_, metasenv, term,ugraph) = 
            MatitaCicMisc.disambiguate ~disambiguator ~currentProof term 
          in
-          let (context, metasenv) =
+          let (context, _) =
             MatitaCicMisc.get_context_and_metasenv currentProof
           in
 (* this is the Eval Compute
@@ -151,7 +155,7 @@ class sharedState
            (* TASSI: here ugraph1 is unused.... FIXME *)
           let expr = Cic.Cast (term, ty) in
           (match mathViewer with
-          | Some v -> v#checkTerm (`Cic expr)
+          | Some v -> v#checkTerm (`Cic (expr, metasenv))
           | _ -> ());
           Quiet
       | TacticAst.Command (TacticAst.Search_pat (search_kind, pat)) ->
@@ -253,14 +257,26 @@ let inddef_of_ast params indTypes (disambiguator:MatitaTypes.disambiguator) =
   let cicIndTypes = List.rev cicIndTypes in
   (UriManager.uri_of_string uri, (cicIndTypes, [], paramsno))
 
-let 
- save_object_to_disk uri obj 
-=
+let save_object_to_disk uri obj =
+  let ensure_path_exists path =
+    let dir = Filename.dirname path in
+    try 
+      let stats = Unix.stat dir in
+      if stats.Unix.st_kind <> Unix.S_DIR then
+        raise (Failure (dir ^ " already exists and is not a directory"))
+      else
+        ()
+    with
+      Unix.Unix_error (_,_,_) -> 
+        let pstatus = Unix.system ("mkdir -p " ^ dir) in
+        match pstatus with
+        | Unix.WEXITED n when n = 0 -> ()
+        | _ -> raise (Failure ("Unable to create " ^ dir))
+  in
   (* generate annobj, ids_to_inner_sorts and ids_to_inner_types *)
   let annobj,_,_,ids_to_inner_sorts,ids_to_inner_types,_,_ =
     Cic2acic.acic_object_of_cic_object ~eta_fix:false obj
   in 
-
   (* prepare XML *)
   let xml, bodyxml =
    Cic2Xml.print_object uri ~ids_to_inner_sorts ~ask_dtd_to_the_getter:false
@@ -270,43 +286,39 @@ let
    Cic2Xml.print_inner_types uri ~ids_to_inner_sorts ~ids_to_inner_types
     ~ask_dtd_to_the_getter:false
   in
-  
   (* prepare URIs and paths *)
   let innertypesuri = UriManager.innertypesuri_of_uri uri in
   let bodyuri = UriManager.bodyuri_of_uri uri in
   let innertypesfilename = Str.replace_first (Str.regexp "^cic:") ""
-        (UriManager.string_of_uri innertypesuri) ^ ".xml" in
-  let innertypespath = !basedir ^ "/" ^ innertypesfilename in
-  let xmlfilename = Str.replace_first (Str.regexp "^cic:") ""
-        (UriManager.string_of_uri uri) ^ ".xml" in
-  let xmlpath = !basedir ^ "/" ^ xmlfilename in
-  let xmlbodyfilename = Str.replace_first (Str.regexp "^cic:") ""
-        (UriManager.string_of_uri uri) ^ ".body.xml" in
-  let xmlbodypath = !basedir ^ "/" ^  xmlbodyfilename in
-  let path_scheme_of path = "file:/" ^ path in
-
+        (UriManager.string_of_uri innertypesuri) ^ ".xml.gz" in
+  let innertypespath = !(Lazy.force basedir) ^ "/" ^ innertypesfilename in
+  let xmlfilename = Str.replace_first (Str.regexp "^cic:/") ""
+        (UriManager.string_of_uri uri) ^ ".xml.gz" in
+  let xmlpath = !(Lazy.force basedir) ^ "/" ^ xmlfilename in
+  let xmlbodyfilename = Str.replace_first (Str.regexp "^cic:/") ""
+        (UriManager.string_of_uri uri) ^ ".body.xml.gz" in
+  let xmlbodypath = !(Lazy.force basedir) ^ "/" ^  xmlbodyfilename in
+  let path_scheme_of path = "file://" ^ path in
+  MatitaMisc.mkdirs (List.map Filename.dirname [innertypespath; xmlpath]);
    (* now write to disk *)
-    Xml.pp ~quiet:true xmlinnertypes (Some innertypespath) ;
-    Xml.pp ~quiet:true xml (Some xmlpath) ;
-
+    ensure_path_exists innertypespath;
+    Xml.pp ~gzip:true xmlinnertypes (Some innertypespath) ;
+    ensure_path_exists xmlpath;
+    Xml.pp ~gzip:true xml (Some xmlpath) ;
+    
    (* now register to the getter *)
     Http_getter.register' innertypesuri (path_scheme_of innertypespath); 
     Http_getter.register' uri (path_scheme_of xmlpath);
-   
     (* now the optional body, both write and register *)
     (match bodyxml,bodyuri with
        None,None -> ()
      | Some bodyxml,Some bodyuri->
-         Xml.pp ~quiet:true bodyxml (Some xmlbodypath) ;
+         ensure_path_exists xmlbodypath;
+         Xml.pp ~gzip:true bodyxml (Some xmlbodypath) ;
          Http_getter.register' bodyuri (path_scheme_of xmlbodypath)
      | _-> assert false) 
-;;
-
 
-
-  (* TODO Zack a lot more to be done here:
-    * - save object to disk in xml format
-    * - register uri to the getter 
+  (* TODO ZACK a lot more to be done here:
     * - save universe file *)
 let add_constant_to_world ~(console: #MatitaTypes.console)
   ~dbd ~uri ?body ~ty ?(params = []) ?(attrs = []) ~ugraph ()
@@ -319,8 +331,7 @@ let add_constant_to_world ~(console: #MatitaTypes.console)
     let obj = Cic.Constant (name, body, ty, params, attrs) in
     let ugraph = CicUnivUtils.clean_and_fill uri obj ugraph in
     CicEnvironment.add_type_checked_term uri (obj, ugraph);
-    MetadataDb.index_constant ~dbd
-      ~owner:(Helm_registry.get "matita.owner") ~uri ~body ~ty;
+    MetadataDb.index_constant ~dbd ~uri ~body ~ty;
     save_object_to_disk uri obj;  
     console#echo_message (sprintf "%s constant defined" suri)
   end
@@ -336,8 +347,7 @@ let add_inductive_def_to_world ~(console: #MatitaTypes.console)
     let obj = Cic.InductiveDefinition (indTypes, params, leftno, attrs) in
     let ugraph = CicUnivUtils.clean_and_fill uri obj ugraph in
     CicEnvironment.put_inductive_definition uri (obj, ugraph);
-    MetadataDb.index_inductive_def ~dbd
-      ~owner:(Helm_registry.get "matita.owner") ~uri ~types:indTypes;
+    MetadataDb.index_inductive_def ~dbd ~uri ~types:indTypes;
     save_object_to_disk uri obj;  
     console#echo_message (sprintf "%s inductive type defined" suri);
     let elim sort =
@@ -359,13 +369,8 @@ let add_inductive_def_to_world ~(console: #MatitaTypes.console)
   end
 
   (** Implements phrases that should be accepted only in Command state *)
-class commandState
-  ~(disambiguator: MatitaTypes.disambiguator)
-  ~(console: #MatitaTypes.console)
-  ?mathViewer
-  ()
-=
-  let shared = new sharedState ~disambiguator ~console ?mathViewer () in
+class commandState ~(console: #MatitaTypes.console) ?mathViewer () =
+  let shared = new sharedState ~console ?mathViewer () in
   object (self)
     inherit interpreterState ~console
 
@@ -378,6 +383,7 @@ class commandState
           let uri = UriManager.uri_of_string (qualify name ^ ".con") in
           let proof = MatitaProof.proof ~typ:expr ~uri ~metasenv () in
           currentProof#start proof;
+          self#setLastItem (Some `Theorem);
           New_state Proof
       | TacticAst.Command
         (TacticAst.Theorem (_, Some name, type_ast, Some body_ast)) ->
@@ -397,6 +403,7 @@ class commandState
           let body = CicMetaSubst.apply_subst subst body_cic in
           let ty = CicMetaSubst.apply_subst subst type_cic in
           add_constant_to_world ~console ~dbd ~uri ~body ~ty ~ugraph ();
+          self#setLastItem (Some (`Def uri));
           Quiet
       | TacticAst.Command (TacticAst.Inductive (params, indTypes)) ->
           
@@ -410,6 +417,7 @@ class commandState
           in
           add_inductive_def_to_world ~console
             ~dbd ~uri ~indTypes ~params ~leftno ~ugraph ();
+          self#setLastItem (Some (`Inductive uri));
           Quiet
       | TacticAst.Command TacticAst.Quit ->
           currentProof#quit ();
@@ -469,8 +477,7 @@ class commandState
                   UriManager.uri_of_string (base ^ xp)
             | Cic.Appl (he::_) -> uri_of_term he
             | t -> 
-                prerr_endline ("Fallisco a estrarre la uri di " ^ 
-                  (CicPp.ppterm t));
+                error ("can't extract uri from " ^ (CicPp.ppterm t));
                 assert false 
           in
           let ty_src,ty_tgt = extract_last_two_p coer_ty in
@@ -514,13 +521,8 @@ let namer_of names =
 
   (** Implements phrases that should be accepted only in Proof state, basically
   * tacticals *)
-class proofState
-  ~(disambiguator: MatitaTypes.disambiguator)
-  ~(console: #MatitaTypes.console)
-  ?mathViewer
-  ()
-=
-  let shared = new sharedState ~disambiguator ~console ?mathViewer () in
+class proofState ~(console: #MatitaTypes.console) ?mathViewer () =
+  let shared = new sharedState ~console ?mathViewer () in
   object (self)
     inherit interpreterState ~console
 
@@ -559,10 +561,21 @@ class proofState
           Tactics.replace ~what:(self#disambiguate what)
             ~with_what:(self#disambiguate with_what)
       | TacticAst.Auto -> Tactics.auto_new ~dbd
+      | TacticAst.Hint -> 
+          let l = List.map fst 
+            (MetadataQuery.experimental_hint ~dbd  
+             (currentProof#proof#proof,currentProof#proof#goal))
+          in
+          let u = console#choose_uri l in
+          Tactics.apply (CicUtil.term_of_uri u) 
+      | TacticAst.Change (what, with_what, _) ->
+          let what = self#disambiguate what in
+          let with_what = self#disambiguate with_what in
+          Tactics.change ~what ~with_what
     (*
       (* 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.Change of 'term * 'term * 'ident option
       | TacticAst.Decompose of 'ident * 'ident list
       | TacticAst.Discriminate of 'ident
       | TacticAst.Fold of reduction_kind * 'term
@@ -604,30 +617,35 @@ class proofState
           add_constant_to_world ~console ~dbd ~uri ~body:bo ~ty ~ugraph ();
           currentProof#abort ();
           console#echo_message (sprintf "%s defined" suri);
+          self#setLastItem (Some (`Qed uri));
           New_state Command
       | TacticAst.Seq tacticals ->
-          (* TODO Zack check for proof completed at each step? *)
+          (* TODO ZACK check for proof completed at each step? *)
+          (* TODO ZACK code completely broken here: we must build logic level
+          * tacticals instead of iterating interpreter evaluation *)
+          if (List.length tacticals > 1) then
+            warning "tacticals are broken: see matitaInterpreter.ml";
           List.iter (fun t -> ignore (self#evalTactical t)) tacticals;
+          self#setLastItem (Some `Tactic);
           New_state Proof
       | TacticAst.Tactic tactic_phrase ->
           let tactic = self#lookup_tactic tactic_phrase in
           currentProof#proof#apply_tactic tactic;
+          self#setLastItem (Some `Tactic);
           New_state Proof
       | tactical -> shared#evalTactical tactical
   end
 
-class interpreter
-  ~(disambiguator: MatitaTypes.disambiguator)
-  ~(console: #MatitaTypes.console)
-  ?mathViewer
-  ()
-=
-  let commandState = new commandState ~disambiguator ~console ?mathViewer () in
-  let proofState = new proofState ~disambiguator ~console ?mathViewer () in
+class interpreter ~(console: #MatitaTypes.console) ?mathViewer () =
+  let commandState = new commandState ~console ?mathViewer () in
+  let proofState = new proofState ~console ?mathViewer () in
   object (self)
     val mutable state = commandState
 
-    method reset = state <- commandState
+    method setState (tag: [`Proof | `Command]) =
+      match tag with
+      | `Proof -> (state <- proofState)
+      | `Command -> (state <- commandState)
 
     method endOffset = state#endOffset
 
@@ -647,13 +665,14 @@ class interpreter
 
     method evalPhrase s = self#eval (fun () -> state#evalPhrase s)
     method evalAst ast = self#eval (fun () -> state#evalAst ast)
+
+      (** {2 methods delegated to current state} *)
+
+    method endOffset = state#endOffset
+    method lastItem = state#lastItem
+    method setEvalAstCallback = state#setEvalAstCallback
   end
 
-let interpreter
-  ~(disambiguator: MatitaTypes.disambiguator)
-  ~(console: #MatitaTypes.console)
-  ?mathViewer
-  ()
-=
-  new interpreter ~disambiguator ~console ?mathViewer ()
+let interpreter ~(console: #MatitaTypes.console) ?mathViewer () =
+  new interpreter ~console ?mathViewer ()