src/Tools/isac/BaseDefinitions/termC.sml
author wneuper <Walther.Neuper@jku.at>
Mon, 01 Jan 2024 11:31:16 +0100
changeset 60789 8fa678b678e8
parent 60766 2e0603ca18a4
permissions -rw-r--r--
Doc/Specify_Phase 4: start use antiquotations from isar-ref
     1 (* Title: extends Isabelle/src/Pure/term.ML
     2    Author: Walther Neuper 1999, Mathias Lehnfeld
     3    (c) due to copyright terms
     4 *)
     5 infix contains_one_of
     6 
     7 (* TERM_C extends Isabelle's naming conventions: "C" indicates Isac add-ons to an Isabelle module *)
     8 signature TERM_ISAC =
     9 sig
    10   type as_string
    11   val empty: term
    12   val typ_empty: typ
    13 
    14   datatype lrd = D | L | R
    15   type path
    16   val string_of_path: path -> string
    17   val sub_at: path -> term -> term
    18   val go_up: Proof.context -> path -> term -> term
    19 
    20   val perm: term -> term -> bool (*old original Isabelle code*)
    21   val contains_Var: term -> bool
    22   val dest_binop_typ: typ -> typ * typ * typ
    23   val dest_equals: term -> term * term
    24 
    25   type id
    26   val cut_longid: string -> id
    27   val free2str: term -> string
    28   val str_of_free_opt: term -> string option
    29   val str_of_int: int -> string
    30   val id_of: term -> string
    31   val ids_of: term -> string list
    32   val ids2str: term -> string list (*double?*)
    33 
    34   val ins_concl: term -> term -> term
    35   val inst_abs: term -> term
    36   val inst_bdv: LibraryC.subst -> term -> term
    37 
    38   val lhs: term -> term
    39   val rhs: term -> term
    40   val the_single: term -> term
    41 
    42   val mk_frac: typ -> int * (int * int) -> term
    43   val numerals_to_Free: term -> term
    44   val term_of_num: typ -> int -> term
    45   val num_of_term: term -> int
    46   val to_string: term -> string
    47   val int_of_str: string -> int
    48   val isastr_of_int: int -> string
    49   val int_opt_of_string: string -> int option
    50 
    51   val isalist2list: term -> term list
    52   val list2isalist: typ -> term list -> term
    53   val list2isalist': term list -> term
    54   val wrap_HOL_list: term -> term
    55   val single_to_list: term -> term
    56   val isapair2pair: term -> term * term (* rename to dest_pair, compare HOLogic.dest_string *)
    57 
    58   val is_atom: term -> bool
    59   val string_of_atom: term -> string
    60   val is_variable: term -> bool
    61   val is_bdv: string -> bool
    62   val is_bdv_subst: term -> bool
    63   val guess_bdv_typ: term -> typ
    64   val is_equality: term -> bool
    65   val is_expliceq: term -> bool
    66   val is_f_x: term -> bool
    67   val is_list: term -> bool
    68   val is_bool_list: term -> bool
    69   val dest_listT: typ -> typ
    70   val is_num: term -> bool
    71   val is_num': string -> bool
    72   val string_of_num: term -> string
    73   val variable_constant_pair: term * term -> bool
    74 
    75   val mk_add: term -> term -> term
    76   val mk_free: typ -> string -> term
    77   val mk_equality: term * term -> term
    78   val mk_factroot: string -> typ -> int -> int -> term
    79   val mk_Free: string * typ -> term
    80   val mk_thmid: string -> string -> string -> string
    81   val mk_num_op_num: typ -> typ -> string * typ -> int -> int -> term
    82   val mk_num_op_var: term -> string -> typ -> typ -> int -> term
    83   val mk_var_op_num: term -> string -> typ -> typ -> int -> term
    84 
    85   val matches: theory -> term -> term -> bool
    86 
    87   val strip_imp_prems': term -> term option
    88   val subst_atomic_all: LibraryC.subst -> term -> bool * term
    89 
    90   val pairt: term -> term -> term
    91   val pairT: typ -> typ -> typ
    92   val raise_type_conflicts: term list -> unit
    93   val strip_trueprop: term -> term
    94 
    95   val var2free: term -> term
    96   val vars: term -> term list  (* recognises numerals, should replace "fun vars_of" TODOO*)
    97   val vars': term list -> term list
    98   val vars_of: term -> term list   (* deprecated TODOO: see differences in test/../termC.sml*)
    99   val dest_list': term -> term list
   100   val negates: term -> term -> bool
   101 
   102   val contains_one_of: thm * (string * typ) list -> bool
   103   val contains_Const_typeless: term list -> term -> bool
   104   val sym_trm : term -> term (* unused code, kept as hints to design ideas *)
   105 
   106   val string_of_detail: Proof.context -> term -> string
   107 (*from isac_test for Minisubpbl*)
   108   val atom_typ: Proof.context -> typ -> unit
   109   val atom_trace_detail: Proof.context -> term -> unit
   110 
   111 \<^isac_test>\<open>
   112   val mk_negative: typ -> term -> term
   113   val mk_Var: term -> term
   114   val scala_of_term: term -> string
   115 
   116   val atom_write: Proof.context -> term -> unit
   117   val atom_trace: Proof.context -> term -> unit
   118 
   119   val atom_write_detail: Proof.context -> term -> unit
   120 (*val atom_trace_detail: Proof.context -> term -> unit*)
   121 \<close>
   122 end
   123 
   124 (**)
   125 structure TermC(**): TERM_ISAC(**) =
   126 struct
   127 (**)
   128 
   129 type as_string = UnparseC.term_as_string
   130 val empty = UnparseC.term_empty
   131 val typ_empty = UnparseC.typ_empty
   132 
   133 datatype lrd = L (*t1 in "t1$t2"*)
   134              | R (*t2 in "t1$t2"*) | D; (*b in Abs(_,_,b*)
   135 type path = lrd list; 
   136 fun ldr2str L = "L"
   137   | ldr2str R = "R"
   138   | ldr2str D = "D";
   139 fun string_of_path k = (strs2str' o (map ldr2str)) k;
   140 (*go to a location in a term and fetch the resective sub-term*)
   141 fun sub_at [] t = t
   142   | sub_at (D :: p) (Abs(_, _, body)) = sub_at p body
   143   | sub_at (L :: p) (t1 $ _) = sub_at p t1
   144   | sub_at (R :: p) (_ $ t2) = sub_at p t2
   145   | sub_at l t = raise TERM ("sub_at: no " ^ string_of_path l ^ " for ", [t]);
   146 fun go_up ctxt l t =
   147   if length l > 1 then sub_at (drop_last l) t else raise ERROR ("go_up [] " ^ UnparseC.term ctxt t)
   148 
   149 fun isastr_of_int i = if i >= 0 then string_of_int i else "-" ^ string_of_int (abs i)
   150 
   151 fun matches thy tm pa = 
   152     (Pattern.match thy (pa, tm) (Vartab.empty, Vartab.empty); true)
   153     handle Pattern.MATCH => false
   154 
   155 (** transform  typ / term to a String to be parsed by Scala after transport via libisabelle **)
   156 fun string_of_detail ctxt t =
   157   let
   158     fun ato (Const (a, T)) n = "\n*** " ^ indent n ^ "Const (" ^ a ^ ", " ^ UnparseC.typ ctxt T ^ ")"
   159       | ato (Free (a, T)) n = "\n*** " ^ indent n ^ "Free (" ^ a ^ ", " ^ UnparseC.typ ctxt T ^ ")"
   160       | ato (Var ((a, i), T)) n =
   161         "\n*** " ^ indent n ^ "Var ((" ^ a ^ ", " ^ string_of_int i ^ "), " ^ UnparseC.typ ctxt T ^ ")"
   162       | ato (Bound i) n = "\n*** " ^ indent n ^ "Bound " ^ string_of_int i
   163       | ato (Abs(a, T, body))  n = 
   164         "\n*** " ^ indent n ^ "Abs (" ^ a ^ ", " ^ UnparseC.typ ctxt T ^ ",.." ^ ato body (n + 1)
   165       | ato (f $ t) n = ato f n ^ ato t (n + 1)
   166   in "\n*** " ^ ato t 0 ^ "\n***" end;
   167 
   168 
   169 \<^isac_test>\<open>
   170 fun scala_of_typ (Type (s, typs)) =
   171     enclose "Type(" ")" (quote s ^ ", " ^
   172       (typs |> map scala_of_typ |> commas |> enclose "List(" ")"))
   173   | scala_of_typ (TFree (s, sort)) =
   174     enclose "TFree(" ")" (quote s ^ ", " ^ (sort |> map quote |> commas |> enclose "List(" ")"))
   175   | scala_of_typ (TVar ((s, i), sort)) =
   176     enclose "TVar(" ")" (
   177       enclose "(" ")," (quote s ^ ", " ^ quote (string_of_int i)) ^ 
   178       (sort |> map quote |> commas |> enclose "List(" ")"))
   179 
   180 fun scala_of_term (Const (s, T)) =
   181     enclose "Const (" ")" (quote s ^ ", " ^ scala_of_typ T)
   182   | scala_of_term (Free (s, T)) =
   183     enclose "Free(" ")" (quote s ^ ", " ^ scala_of_typ T)
   184   | scala_of_term (Var ((s, i), T)) =
   185     enclose "TVar(" ")" (
   186       enclose "(" ")," (quote s ^ ", " ^ quote (string_of_int i)) ^ 
   187       scala_of_typ T)
   188   | scala_of_term (Bound i) = enclose "Bound(" ")" (string_of_int i)
   189   | scala_of_term (Abs (s, T, t)) =
   190     enclose "Abs(" ")" (
   191       quote s ^ ", " ^
   192       scala_of_typ T ^ ", " ^
   193       scala_of_term t)
   194   | scala_of_term (t1 $ t2) =
   195     enclose "App(" ")" (scala_of_term t1 ^ ", " ^ scala_of_term t2)
   196 \<close>
   197 
   198 (* see structure's bare bones.
   199    for Isabelle standard output compare 2017 "structure ML_PP" *)
   200 fun atom_typ _ t =
   201   let
   202     fun ato n (Type (s, [])) = "\n*** " ^ indent n ^ "Type (" ^ s ^",[])"
   203       | ato n (Type (s, Ts)) = "\n*** " ^ indent n ^ "Type (" ^ s ^ ",[" ^ atol (n + 1) Ts
   204       | ato n (TFree (s, sort)) = "\n*** " ^ indent n ^ "TFree (" ^ s ^ ", " ^ strs2str' sort
   205       | ato n (TVar ((s, i), sort)) =
   206         "\n*** " ^ indent n ^ "TVar ((" ^ s ^ ", " ^ string_of_int i ^ strs2str' sort
   207     and atol n [] = "\n*** " ^ indent n ^ "]"
   208       | atol n (T :: Ts) = (ato n T ^ atol n Ts)
   209 in tracing (ato 0 t ^ "\n") end;
   210 
   211 
   212 \<^isac_test>\<open>
   213 local 
   214   fun ato (Const (a, _)) n = "\n*** " ^ indent n ^ "Const (" ^ a ^ ", _)"
   215 	  | ato (Free (a, _)) n = "\n*** " ^ indent n ^ "Free (" ^ a ^ ", _)"
   216 	  | ato (Var ((a, i), _)) n =
   217 	    "\n*** " ^ indent n ^ "Var (" ^ a ^ ", " ^ string_of_int i ^ "), _)"
   218 	  | ato (Bound i) n = "\n*** " ^ indent n ^ "Bound " ^ string_of_int i
   219 	  | ato (Abs (a, _, body)) n = "\n*** " ^ indent n ^ "Abs(" ^ a ^ ", _" ^ ato body (n+1)
   220 	  | ato (f $ t) n = (ato f n ^ ato t (n + 1))
   221 in
   222   fun atom_write _ t = writeln ("\n*** -------------" ^ ato t 0 ^ "\n***");
   223   fun atom_trace _ t = tracing ("\n*** -------------" ^ ato t 0 ^ "\n***");
   224 end;
   225 
   226 fun atom_write_detail ctxt t = (writeln o (string_of_detail ctxt)) t;
   227 \<close>
   228 fun atom_trace_detail ctxt t = (tracing o (string_of_detail ctxt)) t;
   229 
   230 (* contains the term a VAR(("*",_),_) ? *)
   231 fun contains_Var (Abs(_,_,body)) = contains_Var body
   232   | contains_Var (f $ f') = contains_Var f orelse contains_Var f'
   233   | contains_Var (Var _) = true
   234   | contains_Var _ = false;
   235 
   236 fun str_of_int n = 
   237   if n < 0 then "-" ^ ((string_of_int o abs) n)
   238   else string_of_int n;
   239 val int_of_str = Value.parse_int;
   240 
   241 val int_opt_of_string = ThmC_Def.int_opt_of_string
   242 fun is_num' str = case int_opt_of_string str of SOME _ => true | NONE => false;
   243 
   244 fun is_num (Const (\<^const_name>\<open>numeral\<close>, _) $ _) = true
   245   | is_num (Const (\<^const_name>\<open>uminus\<close>, _) $
   246     (Const (\<^const_name>\<open>numeral\<close>, _) $ _)) = true
   247   | is_num (Const (\<^const_name>\<open>one_class.one\<close>, _)) = true
   248   | is_num (Const (\<^const_name>\<open>uminus\<close>, _) $ Const (\<^const_name>\<open>one_class.one\<close>, _)) = true
   249   | is_num (Const (\<^const_name>\<open>zero_class.zero\<close>, _)) = true
   250   | is_num (Const (\<^const_name>\<open>uminus\<close>, _) $ Const (\<^const_name>\<open>zero_class.zero\<close>, _)) = true
   251   | is_num _ = false;
   252 
   253 fun string_of_num n = (n |> HOLogic.dest_number |> snd |> string_of_int)
   254 
   255 fun lhs (Const (\<^const_name>\<open>HOL.eq\<close>,_) $ l $ _) = l
   256   | lhs t = raise TERM ("lhs called with ", [t]);
   257 fun rhs (Const (\<^const_name>\<open>HOL.eq\<close>,_) $ _ $ r) = r
   258   | rhs t = raise TERM ("rhs called with ", [t]);
   259 fun the_single (Const ("List.list.Cons", _) $ t $ Const ("List.list.Nil", _)) = t
   260   | the_single t = raise TERM ("the_single", [t])
   261 
   262 
   263 fun mk_negative T t = Const (\<^const_name>\<open>uminus\<close>, T --> T) $ t
   264 fun mk_frac T (sg, (i1, i2)) =
   265   if sg = 1 then
   266     if i2 = 1 then HOLogic.mk_number T i1
   267     else Const (\<^const_name>\<open>divide_class.divide\<close>, T --> T --> T) $
   268       HOLogic.mk_number T i1 $ HOLogic.mk_number T i2
   269   else (*take negative*)
   270     if i2 = 1 then mk_negative T (HOLogic.mk_number T i1)
   271     else Const (\<^const_name>\<open>divide_class.divide\<close>, T --> T --> T) $
   272       mk_negative T (HOLogic.mk_number T i1) $ HOLogic.mk_number T i2
   273 
   274 val numerals_to_Free = (* Makarius 100308 *)
   275   let
   276     fun dest_num t =
   277       (case try HOLogic.dest_number t of
   278         SOME (T, i) => SOME (Free (signed_string_of_int i, T))
   279       | NONE => NONE);
   280     fun to_str (Abs (x, T, b)) = Abs (x, T, to_str b)
   281       | to_str (t as (u1 $ u2)) =
   282           (case dest_num t of SOME t' => t' | NONE => to_str u1 $ to_str u2)
   283       | to_str t = perhaps dest_num t;
   284   in to_str end
   285 
   286 val term_of_num = HOLogic.mk_number;
   287 fun num_of_term t = t |> HOLogic.dest_number |> snd;
   288 (* accomodate string-representation for int to term-orders *)
   289 fun to_string t = t |> num_of_term |> signed_string_of_int
   290 
   291 fun is_variable (t as Free _) = not (is_num t)
   292   | is_variable _ = false;
   293 fun is_Free (Free _) = true | is_Free _ = false;
   294 fun is_fun_id (Const _) = true
   295   | is_fun_id (Free _) = true
   296   | is_fun_id _ = false;
   297 fun is_f_x (f $ x) = is_fun_id f andalso is_Free x
   298   | is_f_x _ = false;
   299 (* precondition: TermC.is_atom v andalso TermC.is_atom c *)
   300 fun is_const (Const _) = true | is_const _ = false;
   301 fun variable_constant_pair (v, c) =
   302   if (is_variable v andalso (is_const c orelse is_num c)) orelse
   303      (is_variable c andalso (is_const v orelse is_num v))
   304   then true
   305   else false
   306 
   307 fun vars t =
   308   let
   309     fun scan vs (Const _) = vs
   310       | scan vs (t as Free _) = (*if is_num' s then vs else*) t :: vs
   311       | scan vs (t as Var _) = t :: vs
   312       | scan vs (Bound _) = vs 
   313       | scan vs (Abs (_, _, t)) = scan vs t
   314       | scan vs (t1 $ t2) = (scan vs t1) @ (scan vs t2)
   315   in ((distinct op =) o (scan [])) t end;
   316 fun vars' ts = ts |> map vars |> flat |> distinct op =
   317 
   318 (* bypass Isabelle's Pretty, which requires ctxt *)
   319 fun ids2str t =
   320   let
   321     fun scan vs (t as Const (s, _) $ arg) =
   322         if is_num t then vs else scan (s :: vs) arg
   323       | scan vs (Const (s as "Partial_Fractions.AA", _)) = s :: vs (*how get rid of spec.case?*)
   324       | scan vs (Const (s as "Partial_Fractions.BB", _)) = s :: vs (*how get rid of spec.case?*)
   325       | scan vs (Const _) = vs
   326       | scan vs (Free (s, _)) = if is_num' s then vs else s :: vs
   327       | scan vs (Var ((s, i), _)) = (s ^ "_" ^ string_of_int i) :: vs
   328       | scan vs (Bound _) = vs 
   329       | scan vs (Abs (s, _, t)) = scan (s :: vs) t
   330       | scan vs (t1 $ t2) = (scan vs t1) @ (scan vs t2)
   331   in ((distinct op =) o (scan [])) t
   332   end;
   333 fun is_bdv str = case Symbol.explode str of "b"::"d"::"v"::_ => true | _ => false;
   334 (* instantiate #prop thm with bound vars_of (as Free) *)
   335 fun inst_bdv [] t = t
   336   | inst_bdv (instl: (term*term) list) t =
   337     let
   338       fun subst (v as Var((s, _), T)) = 
   339           (case Symbol.explode s of
   340             "b"::"d"::"v"::_ => if_none (assoc(instl,Free(s,T))) (Free(s,T))
   341           | _ => v)
   342         | subst (Abs(a, T, body)) = Abs(a, T, subst body)
   343         | subst (f $ t') = subst f $ subst t'
   344         | subst t = if_none (assoc (instl, t)) t
   345     in  subst t  end;
   346 
   347 (* is a term a substitution for a bdv as found in programs and tactics *)
   348 fun is_bdv_subst (Const (\<^const_name>\<open>Cons\<close>, _) $
   349       (Const (\<^const_name>\<open>Pair\<close>, _) $ str $ _) $ _) = is_bdv (HOLogic.dest_string str)
   350   | is_bdv_subst _ = false;
   351 
   352 (* this shall be improved due to future requirements *)
   353 fun guess_bdv_typ t = t |> vars |> hd |> type_of
   354 
   355 type id = string (*the shortest significant*)
   356 val cut_longid = ThmC.cut_longid
   357 fun free2str (Free (s, _)) = s
   358   | free2str t = raise TERM ("free2str not for ", [t]);
   359 fun str_of_free_opt (Free (s, _)) = SOME s
   360   | str_of_free_opt _ = NONE
   361 fun id_of (Var ((id,ix),_)) = if ix= 0 then id else id ^ string_of_int ix
   362   | id_of (Free (id    ,_)) = id
   363   | id_of (Const(id    ,_)) = id
   364   | id_of _                 = ""; (* never such an identifier *)
   365 
   366 fun ids_of t =
   367   let fun con ss (Const (s,_)) = s::ss
   368 	| con ss (Free (s,_)) = s::ss
   369 	| con ss (Abs (s,_,b)) = s::(con ss b)
   370 	| con ss (t1 $ t2) = (con ss t1) @ (con ss t2)
   371 	| con ss _ = ss
   372   in map cut_longid (((distinct op =) o (con [])) t) end;
   373 
   374 (* compare Logic.unvarify_global, which rejects Free *)
   375 fun var2free (t as Const _) = t
   376   | var2free (t as Free _) = t
   377   | var2free (Var((s, _), T)) = Free (s,T)
   378   | var2free (t as Bound _) = t 
   379   | var2free (Abs(s, T, t)) = Abs(s, T, var2free t)
   380   | var2free (t1 $ t2) = (var2free t1) $ (var2free t2);
   381   
   382 
   383 \<^isac_test>\<open> (*TODO: check with new numerals --vv*)
   384 (* Logic.varify does NOT take care of 'Free ("1", _)'*)
   385 fun mk_Var (t as Const _) = t
   386   | mk_Var (t as Free (s, T)) = if is_num' s then t else Var ((s, 0), T)
   387   | mk_Var (t as Var _) = t
   388   | mk_Var (t as Bound _) = t 
   389   | mk_Var (Abs (s, T, t)) = Abs (s, T, mk_Var t)
   390   | mk_Var (t1 $ t2) = (mk_Var t1) $ (mk_Var t2);
   391 \<close>
   392 
   393 fun mk_listT T = Type (\<^type_name>\<open>list\<close>, [T]);
   394 fun list_const T = Const (\<^const_name>\<open>Cons\<close>, [T, mk_listT T] ---> mk_listT T);
   395 fun list2isalist T [] = Const (\<^const_name>\<open>Nil\<close>, mk_listT T)
   396   | list2isalist T (t :: ts) = (list_const T) $ t $ (list2isalist T ts);
   397 fun list2isalist' ts = list2isalist (type_of (hd ts)) ts
   398 fun wrap_HOL_list t = [t] |> list2isalist'
   399 
   400 (*maps   [<a>, <b>, <c>]  to  [<[a]>, <[b]>, <[c]>]*)
   401 fun single_to_list t = t |> single |> list2isalist (type_of t)
   402 
   403 
   404 fun isapair2pair (Const (\<^const_name>\<open>Pair\<close>,_) $ a $ b) = (a, b)
   405   | isapair2pair t = 
   406     raise TERM ("isapair2pair called with ", [t]);
   407 fun isalist2list ls =
   408   let
   409     fun get es (Const(\<^const_name>\<open>Cons\<close>, _) $ t $ ls) = get (t :: es) ls
   410       | get es (Const(\<^const_name>\<open>Nil\<close>, _)) = es
   411       | get _ t = raise TERM ("isalist2list applied to NON-list: ", [t])
   412   in (rev o (get [])) ls end;
   413 
   414 fun is_list ((Const (\<^const_name>\<open>Cons\<close>, _)) $ _ $ _) = true
   415   | is_list _ = false;
   416 fun dest_listT (Type (\<^type_name>\<open>list\<close>, [T])) = T
   417   | dest_listT T = raise TYPE ("dest_listT: list type expected", [T], []);
   418 fun is_bool_list t =
   419   (if dest_listT (Term.type_of t) = HOLogic.boolT then true else false)
   420   handle TYPE _ => false
   421 
   422 
   423 fun dest_binop_typ (Type (\<^type_name>\<open>fun\<close>, [range, Type (\<^type_name>\<open>fun\<close>, [arg2, arg1])])) = (arg1, arg2, range)
   424   | dest_binop_typ _ = raise ERROR "dest_binop_typ: not binary";
   425 fun dest_equals (Const(\<^const_name>\<open>HOL.eq\<close>, _) $ t $ u)  =  (t, u) (* Pure/logic.ML: Const ("==", ..*)
   426   | dest_equals t = raise TERM ("dest_equals'", [t]);
   427 fun is_equality (Const(\<^const_name>\<open>HOL.eq\<close>,_) $ _ $ _)  =  true  (* logic.ML: Const ("=="*)
   428   | is_equality _ = false;
   429 fun mk_equality (t, u) = (Const(\<^const_name>\<open>HOL.eq\<close>, [type_of t, type_of u] ---> HOLogic.boolT) $ t $ u); 
   430 fun is_expliceq (Const(\<^const_name>\<open>HOL.eq\<close>,_) $ (Free _) $ _)  =  true
   431   | is_expliceq _ = false;
   432 fun strip_trueprop (Const (\<^const_name>\<open>Trueprop\<close>, _) $ t) = t
   433   | strip_trueprop t = t;
   434 
   435 (* (A1==>...An==>B) goes to (A1==>...An==>)   Pure/logic.ML: term -> term list*)
   436 fun strip_imp_prems' (Const (\<^const_name>\<open>Pure.imp\<close>, _) $ A $ t) = 
   437     let
   438       fun coll_prems As (Const(\<^const_name>\<open>Pure.imp\<close>, _) $ A $ t) = 
   439           coll_prems (As $ (Logic.implies $ A)) t
   440         | coll_prems As _ = SOME As
   441     in coll_prems (Logic.implies $ A) t end
   442   | strip_imp_prems' _ = NONE;  (* *)
   443 
   444 (* (A1==>...An==>) (B) goes to (A1==>...An==>B), where B is lowest branch, 2002 Pure/thm.ML *)
   445 fun ins_concl (Const (\<^const_name>\<open>Pure.imp\<close>, _) $ A $ t) B = Logic.implies $ A $ (ins_concl t B)
   446   | ins_concl (Const (\<^const_name>\<open>Pure.imp\<close>, _) $ A    ) B = Logic.implies $ A $ B
   447   | ins_concl t B =  raise TERM ("ins_concl", [t, B]);
   448 
   449 fun vperm (Var _, Var _) = true  (* 2002 Pure/thm.ML *)
   450   | vperm (Abs (_, _, s), Abs (_, _, t)) = vperm (s, t)
   451   | vperm (t1 $ t2, u1 $ u2) = vperm (t1, u1) andalso vperm (t2, u2)
   452   | vperm (t, u) = (t = u);
   453 
   454 (*2002 cp from Pure/term.ML --- since 2009 in Pure/old_term.ML*)
   455 fun mem_term (_, []) = false
   456   | mem_term (t, t' :: ts) = t aconv t' orelse mem_term (t, ts);
   457 fun subset_term ([], _) = true
   458   | subset_term (x :: xs, ys) = mem_term (x, ys) andalso subset_term (xs, ys);
   459 fun eq_set_term (xs, ys) =
   460     xs = ys orelse (subset_term (xs, ys) andalso subset_term (ys, xs));
   461 (*a total, irreflexive ordering on index names*)
   462 fun xless ((a, i), (b, j): indexname) = i<j  orelse  (i = j andalso a < b);
   463 (*a partial ordering (not reflexive) for atomic terms*)
   464 fun atless (Const (a, _), Const (b, _)) = a < b
   465   | atless (Free (a, _), Free (b, _)) = a < b
   466   | atless (Var (v, _), Var (w, _)) = xless (v, w)
   467   | atless (Bound i, Bound j) =  i < j
   468   | atless _ = false;
   469 (*insert atomic term into partially sorted list, suppressing duplicates (?)*)
   470 fun insert_aterm (t,us) =
   471   let fun inserta [] = [t]
   472         | inserta (us as u::us') =
   473               if atless(t,u) then t::us
   474               else if t=u then us (*duplicate*)
   475               else u :: inserta us'
   476   in inserta us end;
   477 
   478 (* Accumulates the Vars in the term, suppressing duplicates *)
   479 fun add_term_vars (t, vars: term list) = case t of
   480     Var   _ => insert_aterm (t, vars)
   481   | Abs (_, _, body) => add_term_vars (body, vars)
   482   | f$t =>  add_term_vars (f, add_term_vars (t, vars))
   483   | _ => vars;
   484 fun term_vars t = add_term_vars (t, []);
   485 
   486 (*2002 Pure/thm.ML *)
   487 fun var_perm (t, u) = vperm (t, u) andalso eq_set_term (term_vars t, term_vars u);
   488 (*2002 fun decomp_simp, Pure/thm.ML *)
   489 fun perm lhs rhs = var_perm (lhs, rhs) andalso not (lhs aconv rhs) andalso not (is_Var lhs);
   490 
   491 
   492 fun pairT T1 T2 = Type (\<^type_name>\<open>prod\<close>, [T1, T2]);
   493 fun PairT T1 T2 = ([T1, T2] ---> Type (\<^type_name>\<open>prod\<close>, [T1, T2]));
   494 fun pairt t1 t2 = Const (\<^const_name>\<open>Pair\<close>, PairT (type_of t1) (type_of t2)) $ t1 $ t2;
   495 
   496 fun mk_factroot op_(*=thy.sqrt*) T fact root = 
   497   Const (\<^const_name>\<open>times\<close>, [T, T] ---> T) $ (term_of_num T fact) $
   498     (Const (op_, T --> T) $ term_of_num T root);
   499 fun mk_var_op_num v op_ optype ntyp n = Const (op_, optype) $ v $ HOLogic.mk_number ntyp n;
   500 fun mk_num_op_var v op_ optype ntyp n = Const (op_, optype) $ HOLogic.mk_number ntyp n $ v;
   501 fun mk_num_op_num T1 T2 (op_, Top) n1 n2 =
   502   Const (op_, Top) $ HOLogic.mk_number T1 n1 $ HOLogic.mk_number T2 n2;
   503 fun mk_thmid thmid n1 n2 = 
   504   thmid ^ (cut_longid n1) ^ "_" ^ (cut_longid n2);
   505 fun mk_add t1 t2 =
   506   let
   507     val (T1, T2) = (type_of t1, type_of t2)
   508   in
   509     if T1 <> T2 then raise TYPE ("mk_add gets ", [T1, T2], [t1,t2])
   510     else (Const (\<^const_name>\<open>plus\<close>, [T1, T2] ---> T1) $ t1 $ t2)
   511   end;
   512 
   513 (** transform binary numeralsstrings **)
   514 fun mk_Free (s,T) = Free (s, T);
   515 fun mk_free T s =  Free (s, T);
   516 
   517 (*Special case: one argument cp from Isabelle2002/src/Pure/term.ML*)
   518 fun subst_bound (arg, t) =
   519   let
   520     fun subst (t as Bound i, lev) =
   521         if i < lev then t (*var is locally bound*)
   522         else if i = lev then incr_boundvars lev arg
   523         else Bound (i - 1) (*loose: change it*)
   524       | subst (Abs(a, T, body), lev) = Abs (a, T, subst (body, lev + 1))
   525       | subst (f$t, lev) =  subst(f, lev)  $  subst(t, lev)
   526       | subst (t, _) = t
   527   in subst (t, 0)  end;
   528 
   529 (* instantiate let; necessary for scan_up1 *)
   530 fun inst_abs (Const sT) = Const sT
   531   | inst_abs (Free sT) = Free sT
   532   | inst_abs (Bound n) = Bound n
   533   | inst_abs (Var iT) = Var iT
   534   | inst_abs (Const (\<^const_name>\<open>Let\<close>,T1) $ e $ (Abs (v, T2, b))) = 
   535     let val b' = subst_bound (Free (v, T2), b); (*fun variant_abs: term.ML*)
   536     in Const (\<^const_name>\<open>Let\<close>, T1) $ inst_abs e $ (Abs (v, T2, inst_abs b')) end
   537   | inst_abs (t1 $ t2) = inst_abs t1 $ inst_abs t2
   538   | inst_abs t = t;
   539 
   540 fun is_atom (Const _) = true
   541   | is_atom (Free _) = true
   542   | is_atom (Var _) = true
   543   | is_atom t = is_num t;
   544 fun string_of_atom (t as Const (\<^const_name>\<open>numeral\<close>, _) $ _) = to_string t
   545   | string_of_atom (t as Const (\<^const_name>\<open>one_class.one\<close>, _)) = to_string t
   546   | string_of_atom (t as Const (\<^const_name>\<open>zero_class.zero\<close>, _)) = to_string t
   547   | string_of_atom (Const (str, _)) = str
   548   | string_of_atom (Free (str, _)) = str
   549   | string_of_atom (Var ((str, int), _)) = str ^ "_" ^ string_of_int int
   550   | string_of_atom _ = "DUMMY-string_of_atom";
   551 
   552 (* from Pure/term.ML; reports if ALL Free's have found a substitution
   553    (required for evaluating the preconditions of _incomplete_ models) *)
   554 fun subst_atomic_all [] t = (false (*TODO may be 'true' for some terms ?*), t)
   555   | subst_atomic_all instl t =
   556     let
   557       fun subst (Abs (a, T, body)) = 
   558           let
   559             val (all, body') = subst body
   560           in (all, Abs(a, T, body')) end
   561         | subst (f $ tt) = 
   562 	        let
   563 	          val (all1, f') = subst f
   564 	          val (all2, tt') = subst tt
   565 	        in (all1 andalso all2, f' $ tt') end
   566         | subst (t as Free _) = 
   567 	        if is_num t then (true, t) (*numerals cannot be subst*)
   568 	        else (case assoc (instl, t) of
   569 					  SOME t' => (true, t')
   570 				  | NONE => (false, t))
   571         | subst t = (true, if_none (assoc (instl, t)) t)
   572     in subst t end;
   573 
   574 fun op contains_one_of (thm, ids) =
   575   Term.exists_Const (fn id => member op= ids id) (Thm.prop_of thm)
   576 
   577 fun var_for vs (t as Const (str, _)) id =
   578     if is_num t then vs
   579     else if id = cut_longid str then t :: vs else vs
   580   | var_for vs (t as Free (str, _)) id = if id = str then t :: vs else vs
   581   | var_for vs (t as Var (idn, _)) id = if id = Term.string_of_vname idn then t :: vs else vs
   582   | var_for vs (Bound _) _ = vs
   583   | var_for vs (Abs (_, _, t)) id = var_for vs t id
   584   | var_for vs (t1 $ t2) id = (var_for vs t1 id) @ (var_for vs t2 id)
   585 
   586 val poly_consts =
   587   [\<^const_name>\<open>plus\<close>, \<^const_name>\<open>minus\<close>,
   588   \<^const_name>\<open>divide\<close>, \<^const_name>\<open>times\<close>,
   589   \<^const_name>\<open>realpow\<close>];
   590 (* treat Free, Const, Var as vars_of in polynomials *)
   591 fun vars_of t =
   592   let
   593     val var_ids = t |> ids2str |> subtract op = poly_consts |> map cut_longid |> sort string_ord
   594   in (map (var_for [] t) var_ids) |> flat |> distinct op = end
   595 
   596 (* this may decompose an object-language isa-list;
   597    use only, if description is not available, eg. not input ?WN:14.5.03 ??!?*)
   598 fun dest_list' t = if is_list t then isalist2list t  else [t];
   599 
   600 fun negat (Const (\<^const_name>\<open>Not\<close>, _) $ P, P') = P = P'
   601   | negat _ = false
   602 fun negates p1 p2 = negat (p1, p2) orelse negat (swap (p1, p2));
   603 
   604 fun raise_type_conflicts ts =
   605   let
   606     val dups = duplicates (op =) (map (fst o dest_Free) ts)
   607     val confl = filter (fn Free (str, _) => member op = dups str
   608                          | _ => false) ts
   609   in
   610     if confl = []
   611     then ()
   612     else raise TYPE ("formalisation inconsistent w.r.t. type inference: ",
   613       map (snd o dest_Free)confl, confl)
   614   end
   615 
   616 (* expects t as Const *)
   617 fun contains_Const_typeless ts t = (t
   618   |> strip_comb |> fst
   619   |> member (fn (t1, t2) => fst (dest_Const t1) = fst (dest_Const t2)) ts
   620 ) handle TERM("dest_Const", _) => raise TERM ("contains_Const_typeless", [t])
   621 
   622 (* WN100910 weaker than fun sym_thm for Theory.axioms_of in isa02 *)
   623 fun sym_trm trm =
   624   let
   625     val (lhs, rhs) = (dest_equals o strip_trueprop o Logic.strip_imp_concl) trm
   626     val trm' = case strip_imp_prems' trm of
   627 	      NONE => mk_equality (rhs, lhs)
   628 	    | SOME cs => ins_concl cs (mk_equality (rhs, lhs))
   629   in trm' end
   630 
   631 end