File ‹mutabelle.ML›
signature MUTABELLE =
sig
exception WrongPath of string;
exception WrongArg of string;
val freeze : term -> term
val mutate_exc : term -> string list -> int -> term list
val mutate_sign : term -> theory -> (string * string) list -> int -> term list
val mutate_mix : term -> theory -> string list ->
(string * string) list -> int -> term list
end;
structure Mutabelle : MUTABELLE =
struct
fun consts_of thy =
let
val {const_space, constants, ...} = Consts.dest (Sign.consts_of thy)
in
map_filter (fn (s, (T, NONE)) => SOME (s, T) | _ => NONE)
(filter_out (fn (s, _) => Name_Space.is_concealed const_space s) constants)
end;
exception WrongPath of string;
exception WrongArg of string;
fun rename_bnds curTerm 0 = curTerm
| rename_bnds (Bound(i)) minInd =
let
val erg = if (i-minInd < 0) then 0 else (i - minInd)
in
Bound(erg)
end
| rename_bnds (Abs(name,t,uTerm)) minInd =
Abs(name,t,(rename_bnds uTerm minInd))
| rename_bnds (fstUTerm $ sndUTerm) minInd =
(rename_bnds fstUTerm minInd) $ (rename_bnds sndUTerm minInd)
| rename_bnds elseTerm minInd = elseTerm;
fun getSubTermList (Const(name,t)) abscontext path acc =
(Const(name,t),t,abscontext,abscontext,path)::acc
| getSubTermList (Free(name,t)) abscontext path acc =
(Free(name,t),t,abscontext,abscontext,path)::acc
| getSubTermList (Var(indname,t)) abscontext path acc =
(Var(indname,t),t,abscontext,abscontext,path)::acc
| getSubTermList (Bound(i)) abscontext path acc =
(Bound(0),nth abscontext i,abscontext, Library.drop i abscontext,path)::acc
| getSubTermList (Abs(name,t,uTerm)) abscontext path acc =
let
val curTerm = Abs(name,t,uTerm)
val bnos = Term.add_loose_bnos (curTerm,0,[])
val minInd = if (bnos = []) then 0
else Library.foldl (fn (n,m) => if (n<m) then n else m) (hd bnos,tl bnos)
val newTerm = rename_bnds curTerm minInd
val newContext = Library.drop minInd abscontext
in
getSubTermList uTerm (t::abscontext) (0::path)
((newTerm,(fastype_of1 (abscontext, curTerm)),abscontext,newContext,path)::acc)
end
| getSubTermList (fstUTerm $ sndUTerm) abscontext path acc =
let
val curTerm = (fstUTerm $ sndUTerm)
val bnos = Term.add_loose_bnos (curTerm, 0, [])
val minInd = if (bnos = []) then 0
else Library.foldl (fn (n,m) => if (n<m) then n else m) (hd bnos,tl bnos)
val newTerm = rename_bnds curTerm minInd
val newContext = Library.drop minInd abscontext
in
getSubTermList fstUTerm abscontext (0::path)
(getSubTermList sndUTerm abscontext (1::path)
((newTerm,(fastype_of1 (abscontext, curTerm)),abscontext,newContext,path)::acc))
end;
fun is_morespecial longContext shortContext =
let
val revlC = rev longContext
val revsC = rev shortContext
fun is_prefix [] _ = true
| is_prefix _ [] = false
| is_prefix (x::xs) (y::ys) = if (x=y) then is_prefix xs ys else false
in
is_prefix revsC revlC
end;
fun searchForMutatableSubTerm (sterm,stype,sabsContext,sminContext,spath) [] resultList =
resultList
| searchForMutatableSubTerm (sterm,stype,sabsContext,sminContext,spath)
((hdterm,hdtype,hdabsContext,hdminContext,hdpath)::xs) resultList =
if ((stype = hdtype) andalso (is_morespecial sabsContext hdminContext)
andalso (is_morespecial hdabsContext sminContext))
then searchForMutatableSubTerm (sterm,stype,sabsContext,sminContext,spath) xs
((hdterm,hdabsContext,hdminContext,hdpath)::resultList)
else searchForMutatableSubTerm (sterm,stype,sabsContext,sminContext,spath) xs resultList;
fun in_list_forb consSig (consNameStr,consType) [] = false
| in_list_forb consSig (consNameStr,consType) ((forbNameStr,forbTypeStr)::xs) =
let
val forbType = Syntax.read_typ_global consSig forbTypeStr
in
if ((consNameStr = forbNameStr)
andalso (Sign.typ_instance consSig (consType,(Logic.varifyT_global forbType))))
then true
else in_list_forb consSig (consNameStr,consType) xs
end;
fun searchForSignatureMutations (sterm,stype) consSig forbidden_funs =
let
val sigConsTypeList = consts_of consSig;
in
let
fun recursiveSearch mutatableTermList [] = mutatableTermList
| recursiveSearch mutatableTermList ((ConsName,ConsType)::xs) =
if (Sign.typ_instance consSig (stype,ConsType)
andalso (not (sterm = Const(ConsName,stype)))
andalso (not (in_list_forb consSig (ConsName,ConsType) forbidden_funs)))
then recursiveSearch ((Term.Const(ConsName,stype), [], [], [5])::mutatableTermList) xs
else recursiveSearch mutatableTermList xs
in
recursiveSearch [] sigConsTypeList
end
end;
fun searchForMutatableTerm 0 (sterm,stype,sabscontext,smincontext,spath)
subTerms consSig resultList forbidden_funs =
searchForMutatableSubTerm (sterm,stype,sabscontext,smincontext,spath) subTerms resultList
| searchForMutatableTerm 1 (Const(constName,constType),stype,sabscontext,smincontext,spath)
subTerms consSig resultList forbidden_funs =
searchForSignatureMutations (Const(constName,constType),stype) consSig forbidden_funs
| searchForMutatableTerm 1 _ _ _ _ _ = []
| searchForMutatableTerm 2 (Const(constName,constType),stype,sabscontext,smincontext,spath)
subTerms consSig resultList forbidden_funs =
let
val subtermMutations = searchForMutatableSubTerm
(Const(constName,constType),stype,sabscontext,smincontext,spath) subTerms resultList
val signatureMutations = searchForSignatureMutations
(Const(constName,constType),stype) consSig forbidden_funs
in
subtermMutations@signatureMutations
end
| searchForMutatableTerm 2 (sterm,stype,sabscontext,smincontext,spath)
subTerms consSig resultList forbidden_funs =
searchForMutatableSubTerm (sterm,stype,sabscontext,smincontext,spath) subTerms resultList
| searchForMutatableTerm i _ _ _ _ _ =
raise WrongArg("Version " ^ string_of_int i ^
" doesn't exist for function searchForMutatableTerm!") ;
fun areReplacable [] [] = false
| areReplacable _ [] = false
| areReplacable [] _ = false
| areReplacable (x::xs) (y::ys) = if (x=y) then areReplacable xs ys else true;
fun substitute [] _ sndTerm = sndTerm
| substitute (_::xs) (Abs(s,T,subTerm)) sndTerm = Abs(s,T,(substitute xs subTerm sndTerm))
| substitute (0::xs) (t $ u) sndTerm = substitute xs t sndTerm $ u
| substitute (1::xs) (t $ u) sndTerm = t $ substitute xs u sndTerm
| substitute (_::_) _ sndTerm =
raise WrongPath ("The Term could not be found at the specified position");
fun getSubTerm myTerm [] = myTerm
| getSubTerm (Abs(_,_,subTerm)) (0::xs) = getSubTerm subTerm xs
| getSubTerm (t $ _) (0::xs) = getSubTerm t xs
| getSubTerm (_ $ u) (1::xs) = getSubTerm u xs
| getSubTerm _ (_::_) =
raise WrongPath ("The subterm could not be found at the specified position");
fun replace origTerm (fstTerm, fstPath) (sndTerm, sndPath) =
if (areReplacable (rev fstPath) (rev sndPath))
then substitute (rev sndPath) (substitute (rev fstPath) origTerm sndTerm) fstTerm
else origTerm;
fun areCommutative origTerm fstPath sndPath commutatives =
if (sndPath = [])
then false
else
let
val base = (tl sndPath)
in
let
val fstcomm = 1::0::base
val opcomm = 0::0::base
in
if ((fstPath = fstcomm) andalso (is_Const (getSubTerm origTerm (rev opcomm))))
then
let
val Const(name,_) = (getSubTerm origTerm (rev opcomm))
in
member (op =) commutatives name
end
else false
end
end;
fun canonize_term (Const (s, T) $ t $ u) comms =
let
val t' = canonize_term t comms;
val u' = canonize_term u comms;
in
if member (op =) comms s andalso is_less (Term_Ord.term_ord (u', t'))
then Const (s, T) $ u' $ t'
else Const (s, T) $ t' $ u'
end
| canonize_term (t $ u) comms = canonize_term t comms $ canonize_term u comms
| canonize_term (Abs (s, T, t)) comms = Abs (s, T, canonize_term t comms)
| canonize_term t comms = t;
fun createMutatedTerms origTerm _ [] commutatives mutatedTerms = mutatedTerms
| createMutatedTerms origTerm (hdt as (hdTerm,hdabsContext,hdminContext,hdPath))
((sndTerm,sndabsContext,sndminContext,sndPath)::xs) commutatives mutatedTerms =
if (sndPath = [5])
then
let
val canonized = canonize_term (substitute (rev hdPath) origTerm sndTerm) commutatives
in
if (canonized = origTerm)
then createMutatedTerms origTerm hdt xs commutatives mutatedTerms
else createMutatedTerms origTerm hdt xs commutatives
(insert op aconv canonized mutatedTerms)
end
else
if ((areCommutative origTerm hdPath sndPath commutatives)
orelse (areCommutative origTerm sndPath hdPath commutatives))
then createMutatedTerms origTerm hdt xs commutatives mutatedTerms
else
let
val canonized = canonize_term
(replace origTerm
(incr_boundvars (length sndabsContext - length hdminContext) hdTerm,
hdPath)
(incr_boundvars (length hdabsContext - length sndminContext) sndTerm,
sndPath)) commutatives
in
if (not(canonized = origTerm))
then createMutatedTerms origTerm hdt xs commutatives
(insert op aconv canonized mutatedTerms)
else createMutatedTerms origTerm hdt xs commutatives mutatedTerms
end;
fun mutate_once option origTerm tsig commutatives forbidden_funs=
let
val subTermList = getSubTermList origTerm [] [] []
in
let
fun replaceRecursively [] mutatedTerms = mutatedTerms
| replaceRecursively ((hdTerm,hdType,hdabsContext,hdminContext,hdPath)::tail)
mutatedTerms =
replaceRecursively tail (union op aconv (createMutatedTerms origTerm
(hdTerm,hdabsContext,hdminContext,hdPath)
(searchForMutatableTerm option (hdTerm,hdType,hdabsContext,hdminContext,hdPath)
tail tsig [] forbidden_funs)
commutatives []) mutatedTerms)
in
replaceRecursively subTermList []
end
end;
fun mutate_once_rec option [] tsig commutatives forbidden_funs acc = acc
| mutate_once_rec option (x::xs) tsig commutatives forbidden_funs acc =
mutate_once_rec option xs tsig commutatives forbidden_funs
(union op aconv (mutate_once option x tsig commutatives forbidden_funs) acc);
fun mutate option origTerm tsig commutatives forbidden_funs 0 = []
| mutate option origTerm tsig commutatives forbidden_funs 1 =
mutate_once option (canonize_term origTerm commutatives) tsig commutatives forbidden_funs
| mutate option origTerm tsig commutatives forbidden_funs iter =
mutate_once_rec option (mutate option origTerm tsig commutatives forbidden_funs (iter-1))
tsig commutatives forbidden_funs [];
fun mutate_exc origTerm commutatives iter =
mutate 0 origTerm \<^theory>‹Main› commutatives [] iter;
fun mutate_sign origTerm tsig forbidden_funs iter =
mutate 1 origTerm tsig [] forbidden_funs iter;
fun mutate_mix origTerm tsig commutatives forbidden_funs iter =
mutate 2 origTerm tsig commutatives forbidden_funs iter;
fun freeze (t $ u) = freeze t $ freeze u
| freeze (Abs (s, T, t)) = Abs (s, T, freeze t)
| freeze (Var ((a, i), T)) =
Free (if i = 0 then a else a ^ "_" ^ string_of_int i, T)
| freeze t = t;
end