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