src/Tools/isac/Specify/Input_Descript.thy
author wneuper <Walther.Neuper@jku.at>
Thu, 16 Nov 2023 08:15:46 +0100
changeset 60763 2121f1a39a64
parent 60710 21ae85b023bb
child 60764 f82fd40eb400
permissions -rw-r--r--
prepare 14: improved item_to_add

emergency-CS:
* no code cleanup
* ERROR in test/../biegelinie-3.sml outcommented
     1 (* Title:  collect all defitions for the Lucas-Interpreter
     2    Author: Walther Neuper 110226
     3    (c) due to copyright terms
     4 
     5 Descriptors for interactive input into a Specification of an Example.
     6 
     7 Since PblObj.ctxt records explicit type constraints in Formalise.model,
     8 and since we have Model_Pattern.adapt_to_type,
     9 the types of the descriptors below can be more relaxed. 
    10 *)
    11 
    12 theory Input_Descript imports "$ISABELLE_ISAC/BaseDefinitions/BaseDefinitions"
    13 begin
    14 
    15 subsection \<open>typedecl for descriptors of input in specify-phase\<close>
    16 text \<open>
    17   The typedecl below MUST ALL be registered in "Input_Descript.is_a"
    18 \<close>
    19   typedecl nam     (* named variables                                             *)
    20   typedecl una     (* unnamed variables                                           *)
    21   typedecl unl     (* unnamed variables of type list, elementwise input prohibited*)
    22   typedecl str     (* structured variables                                        *)
    23   typedecl toreal  (* var with undef real value: forces typing                    *)
    24   typedecl toreall (* var with undef real list value: forces typing               *)
    25   typedecl tobooll (* var with undef bool list value: forces typing               *)
    26   typedecl unknow  (* input without dsc in fmz=[]                                 *)
    27   typedecl cpy     (* copy-named variables identified by .._0, .._i .._' in pbt   *)
    28 
    29 subsection \<open>consts for general descriptors of input in the specify-phase\<close>
    30 
    31 consts
    32 (*TODO.180331: review for localisation of model-patterns and of method's guards*)
    33   someList       :: "'a list => unl" (*not for elementwise input, eg. inssort*)
    34 
    35   additionalRels :: "bool list => una"
    36   boundVariable    :: "real => una"
    37   FunctionVariable :: "real => una"
    38   derivative     :: "real => una"
    39   equalities     :: "bool list => tobooll" (*WN071228 see fixedValues*)
    40   equality       :: "bool => una"
    41   errorBound     :: "bool => nam"
    42   ErrorBound     :: "bool => nam"
    43   
    44   fixedValues    :: "bool list => nam"
    45   Constants      :: "bool list => nam"
    46   functionEq     :: "bool => una"     (*6.5.03: functionTerm -> functionEq*)
    47   antiDerivative :: "bool => una"
    48   functionOf     :: "real => una"
    49 (*functionTerm   :: 'a => toreal 28.11.00*)
    50   functionTerm   :: "real => una"     (*6.5.03: functionTerm -> functionEq*)
    51   functionName   :: "real => una"
    52   interval       :: "real set => una"
    53   Domain         :: "real set => una"
    54   maxArgument    :: "bool => toreal"
    55   maximum        :: "real => toreal"
    56   Maximum        :: "real => toreal"
    57   Extremum       :: "bool => toreal"
    58   
    59   relations      :: "bool list => una"
    60   SideConditions :: "bool list => una"
    61   solutions      :: "bool list => toreall"
    62 (*Gleichungen    :: "bool list => toreall" MUST be here to beknown by fun is_NObrack_list <--Biegelinie*)
    63 (*solution       :: "bool => toreal"       WN0509 bool list => toreall <--EqSystem*)
    64   solveFor       :: "real => una"
    65   differentiateFor:: "real => una"
    66   unknown        :: "'a => unknow"
    67   valuesFor        :: "real list => toreall"
    68   AdditionalValues :: "real list => toreall"
    69 
    70   intTestGiven   :: "int => una"
    71   realTestGiven  :: "real => una"
    72   realTestFind   :: "real => una"
    73   boolTestGiven  :: "bool => una"
    74   boolTestFind   :: "bool => una"
    75 
    76 subsection \<open>Functions decking for descriptions of input in specify-phase\<close>
    77 ML \<open>
    78 \<close> ML \<open>
    79 signature INPUT_DESCRIPT =
    80 sig
    81 (*type T          ..TODO rename*)
    82   type descriptor = Model_Pattern.descriptor
    83 
    84   val for_real_list: term -> bool
    85   val for_list: term -> bool
    86 \<^isac_test>\<open>
    87   val for_bool_list: term -> bool
    88   val for_fun: term -> bool
    89   val is_a: term -> bool 
    90 \<close>
    91   val split: term -> descriptor * term list
    92   val split': term * term -> term list
    93   val join: descriptor * term list -> term
    94   val join': descriptor * term list -> term
    95   val join''': descriptor * term list -> string
    96   val join'''': descriptor * term list -> term
    97 end
    98 \<close> ML \<open>
    99 
   100 
   101 (**)
   102 structure Input_Descript(**): INPUT_DESCRIPT =(**)
   103 struct
   104 (**)
   105 
   106 type descriptor = Model_Pattern.descriptor;
   107 
   108 (* distinguish input to Model (deep embedding of terms as Isac's object language) *)
   109 fun for_real_list (Const (_, Type(\<^type_name>\<open>fun\<close>, [Type(\<^type_name>\<open>list\<close>, [Type (\<^type_name>\<open>real\<close>, [])]), _]))) = true
   110   | for_real_list (Const (_, Type(\<^type_name>\<open>fun\<close>, [Type(\<^type_name>\<open>list\<close>, [Type (\<^type_name>\<open>real\<close>, [])]),_])) $ _) = true
   111   | for_real_list _ = false;
   112 fun for_bool_list (Const (_, Type(\<^type_name>\<open>fun\<close>, [Type(\<^type_name>\<open>list\<close>, [Type (\<^type_name>\<open>bool\<close>, [])]), _]))) = true
   113   | for_bool_list (Const (_, Type(\<^type_name>\<open>fun\<close>, [Type(\<^type_name>\<open>list\<close>, [Type (\<^type_name>\<open>bool\<close>, [])]),_])) $ _) = true
   114   | for_bool_list _ = false;
   115 fun for_list (Const (_, Type(\<^type_name>\<open>fun\<close>, [Type(\<^type_name>\<open>list\<close>, _), _]))) = true
   116   | for_list (Const (_, Type(\<^type_name>\<open>fun\<close>, [Type(\<^type_name>\<open>list\<close>, _), _])) $ _) = true
   117   (*WN:8.5.03: ???   TODO test/../scrtools.sml                ~~~~ ???*)
   118   | for_list _ = false;
   119 fun for_fun (Const (_, Type(\<^type_name>\<open>fun\<close>, [_, Type(\<^type_name>\<open>unl\<close>, _)]))) = true
   120   | for_fun _ = false;
   121 fun is_a (Const(_, Type(\<^type_name>\<open>fun\<close>, [_, Type(\<^type_name>\<open>nam\<close>, _)]))) = true
   122   | is_a (Const(_, Type(\<^type_name>\<open>fun\<close>, [_, Type(\<^type_name>\<open>una\<close>, _)]))) = true
   123   | is_a (Const(_, Type(\<^type_name>\<open>fun\<close>, [_, Type(\<^type_name>\<open>unl\<close>, _)]))) = true
   124   | is_a (Const(_, Type(\<^type_name>\<open>fun\<close>, [_, Type(\<^type_name>\<open>str\<close>, _)]))) = true
   125   | is_a (Const(_, Type(\<^type_name>\<open>fun\<close>, [_, Type(\<^type_name>\<open>toreal\<close>, _)]))) = true
   126   | is_a (Const(_, Type(\<^type_name>\<open>fun\<close>, [_, Type(\<^type_name>\<open>toreall\<close>, _)])))= true
   127   | is_a (Const(_, Type(\<^type_name>\<open>fun\<close>, [_, Type(\<^type_name>\<open>tobooll\<close>, _)])))= true
   128   | is_a (Const(_, Type(\<^type_name>\<open>fun\<close>, [_, Type(\<^type_name>\<open>unknow\<close>, _)])))= true
   129   | is_a (Const(_, Type(\<^type_name>\<open>fun\<close>, [_, Type(\<^type_name>\<open>cpy\<close>, _)])))= true
   130   | is_a _ = false;
   131 
   132 (* special handling for lists. ?WN:14.5.03 ??!? *)
   133 fun dest_list (d, ts) = 
   134   let fun dest t = 
   135     if for_list d andalso not (for_fun d) andalso not (is_Free t) (*..for pbt*)
   136     then TermC.isalist2list t
   137     else [t]
   138   in (flat o (map dest)) ts end;
   139 
   140 (* take list-term apart w.r.t. handling elementwise input: @{term "[a, b]"} \<rightarrow> ["[a]", "[b]"] *)
   141 fun take_apart t =
   142   let val elems = TermC.isalist2list t
   143   in map ((TermC.list2isalist (type_of (hd elems))) o single) elems end;
   144 
   145 (* decompose an input into description, terms (ev. elems of lists),
   146     and the value for the problem-environment; inv to join   *)
   147 fun split (t as d $ arg) =
   148     if is_a d
   149     then if for_list d andalso TermC.is_list arg andalso for_fun d |> not
   150       then (d, take_apart arg)
   151       else (d, [arg])
   152     else (TermC.empty, TermC.dest_list' t)
   153   | split t =
   154     let val t' as (h, _) = strip_comb t;
   155     in
   156       if is_a h
   157       then (h, dest_list t')
   158       else (TermC.empty, TermC.dest_list' t)
   159     end;
   160 
   161 (* version returning ts only *)
   162 fun split' (d, arg) = 
   163     if is_a d
   164     then if for_list d
   165       then if TermC.is_list arg
   166 	      then if for_fun d
   167 	        then ([arg])           (* e.g. someList [1,3,2]                 *)
   168 		      else (take_apart arg)  (* [a,b] --> SML[ [a], [b] ]SML          *)
   169 	      else ([arg])             (* a variable or metavariable for a list *)
   170 	     else ([arg])
   171     else (TermC.dest_list' arg)
   172 
   173 (* 
   174   take list-term apart w.r.t. handling elementwise input: @{term "[a, b]"} \<rightarrow> ["[a]", "[b]"]
   175     take_apart_inv: term list -> term
   176 *)
   177 fun take_apart_inv ts = (* t = (take_apart_inv o take_apart) t *)
   178   let val elems = (flat o (map TermC.isalist2list)) ts;
   179   in TermC.list2isalist (type_of (hd elems)) elems end;
   180 
   181 fun script_parse str =
   182   let
   183     val ctxt = @{theory BaseDefinitions} |> Proof_Context.init_global
   184   in
   185     ParseC.term_opt ctxt str |> the
   186   end;
   187 val e_listReal = script_parse "[]::(real list)";
   188 val e_listBool = script_parse "[]::(bool list)";
   189 
   190 (* revert split only for ts *)
   191 fun join'''' (d, ts) = 
   192   if for_list d
   193   then if TermC.is_list (hd ts)
   194 	  then if for_fun d
   195 	    then (hd ts)             (* e.g. someList [1,3,2] *)
   196 	    else (take_apart_inv ts) (* [ [a], [b] ] -> [a,b]  ..CONFUSING !!!*)
   197 	  else (hd ts)               (* a variable or metavariable for a list *)
   198   else (hd ts);
   199 fun join (d, []) = 
   200   	if for_real_list d
   201   	then (d $ e_listReal)
   202   	else if for_bool_list d then (d $ e_listBool) else d
   203   | join (d, ts) = 
   204     case \<^try>\<open> d $ (join'''' (d, ts))\<close> of
   205       SOME x => x
   206     | NONE => raise ERROR ("join: " ^ UnparseC.term (ContextC.for_ERROR ()) d ^ " $ " ^ UnparseC.term (ContextC.for_ERROR ()) (hd ts)); 
   207 fun join' (d, []) = 
   208     if for_real_list d
   209     then (d $ e_listReal)
   210     else if for_bool_list d then (d $ e_listBool) else d
   211   | join' (d, ts) = 
   212     case \<^try>\<open> d $ (join'''' (d, ts))\<close> of
   213       SOME x => x
   214     | NONE => raise ERROR ("join': " ^ UnparseC.term (ContextC.for_ERROR ()) d ^ " $ " ^ UnparseC.term (ContextC.for_ERROR ()) (hd ts)); 
   215 fun join''' (d, []) = 
   216     if for_real_list d
   217     then UnparseC.term (ContextC.for_ERROR ()) (d $ e_listReal)
   218     else if for_bool_list d
   219       then UnparseC.term (ContextC.for_ERROR ()) (d $ e_listBool)
   220       else UnparseC.term (ContextC.for_ERROR ()) d
   221   | join''' (d, ts) = 
   222     case \<^try>\<open> UnparseC.term (ContextC.for_ERROR ()) (d $ (join'''' (d, ts)))\<close> of
   223       SOME x => x
   224     | NONE => raise ERROR ("join''': " ^ UnparseC.term (ContextC.for_ERROR ()) d ^ " $ " ^ UnparseC.term (ContextC.for_ERROR ()) (hd ts)); 
   225 
   226 (**)end(**)
   227 \<close> ML \<open>
   228 \<close> ML \<open>
   229 \<close> 
   230 
   231 end