mengj@17998: (* ID: $Id$ mengj@17998: Author: Jia Meng, NICTA mengj@17998: mengj@17998: FOL clauses translated from HOL formulae. Combinators are used to represent lambda terms. mengj@17998: mengj@17998: *) mengj@17998: mengj@17998: structure ResHolClause = mengj@17998: mengj@17998: struct mengj@17998: mengj@17998: mengj@18276: val include_combS = ref false; mengj@18276: val include_min_comb = ref false; mengj@17998: mengj@18356: val const_typargs = ref (Library.K [] : (string*typ -> typ list)); mengj@18356: mengj@18356: fun init thy = (include_combS:=false;include_min_comb:=false;const_typargs := Sign.const_typargs thy); mengj@17998: mengj@19198: mengj@19198: mengj@17998: (**********************************************************************) mengj@17998: (* convert a Term.term with lambdas into a Term.term with combinators *) mengj@17998: (**********************************************************************) mengj@17998: mengj@17998: fun is_free (Bound(a)) n = (a = n) mengj@17998: | is_free (Abs(x,_,b)) n = (is_free b (n+1)) mengj@17998: | is_free (P $ Q) n = ((is_free P n) orelse (is_free Q n)) mengj@17998: | is_free _ _ = false; mengj@17998: mengj@17998: mengj@17998: exception LAM2COMB of term; mengj@17998: mengj@17998: exception BND of term; mengj@17998: mengj@17998: fun decre_bndVar (Bound n) = Bound (n-1) mengj@17998: | decre_bndVar (P $ Q) = (decre_bndVar P) $ (decre_bndVar Q) mengj@17998: | decre_bndVar t = mengj@17998: case t of Const(_,_) => t mengj@17998: | Free(_,_) => t mengj@17998: | Var(_,_) => t mengj@17998: | Abs(_,_,_) => raise BND(t); (*should not occur*) mengj@17998: mengj@17998: mengj@17998: (*******************************************) mengj@17998: fun lam2comb (Abs(x,tp,Bound 0)) _ = mengj@17998: let val tpI = Type("fun",[tp,tp]) mengj@17998: in mengj@18276: include_min_comb:=true; mengj@17998: Const("COMBI",tpI) mengj@17998: end mengj@18356: | lam2comb (Abs(x,tp,Bound n)) Bnds = mengj@18356: let val (Bound n') = decre_bndVar (Bound n) mengj@18356: val tb = List.nth(Bnds,n') mengj@18356: val tK = Type("fun",[tb,Type("fun",[tp,tb])]) mengj@18356: in mengj@18356: include_min_comb:=true; mengj@18356: Const("COMBK",tK) $ (Bound n') mengj@18356: end mengj@17998: | lam2comb (Abs(x,t1,Const(c,t2))) _ = mengj@17998: let val tK = Type("fun",[t2,Type("fun",[t1,t2])]) mengj@17998: in mengj@18276: include_min_comb:=true; mengj@17998: Const("COMBK",tK) $ Const(c,t2) mengj@17998: end mengj@17998: | lam2comb (Abs(x,t1,Free(v,t2))) _ = mengj@17998: let val tK = Type("fun",[t2,Type("fun",[t1,t2])]) mengj@17998: in mengj@18276: include_min_comb:=true; mengj@17998: Const("COMBK",tK) $ Free(v,t2) mengj@17998: end mengj@17998: | lam2comb (Abs(x,t1,Var(ind,t2))) _= mengj@17998: let val tK = Type("fun",[t2,Type("fun",[t1,t2])]) mengj@17998: in mengj@18276: include_min_comb:=true; mengj@17998: Const("COMBK",tK) $ Var(ind,t2) mengj@17998: end mengj@17998: | lam2comb (t as (Abs(x,t1,P$(Bound 0)))) Bnds = mengj@17998: let val nfreeP = not(is_free P 0) mengj@17998: val tr = Term.type_of1(t1::Bnds,P) mengj@17998: in mengj@17998: if nfreeP then (decre_bndVar P) mengj@17998: else ( mengj@17998: let val tI = Type("fun",[t1,t1]) mengj@17998: val P' = lam2comb (Abs(x,t1,P)) Bnds mengj@17998: val tp' = Term.type_of1(Bnds,P') mengj@17998: val tS = Type("fun",[tp',Type("fun",[tI,tr])]) mengj@17998: in mengj@18276: include_min_comb:=true; mengj@18276: include_combS:=true; mengj@17998: Const("COMBS",tS) $ P' $ Const("COMBI",tI) mengj@17998: end) mengj@17998: end mengj@17998: mengj@17998: | lam2comb (t as (Abs(x,t1,P$Q))) Bnds = mengj@17998: let val (nfreeP,nfreeQ) = (not(is_free P 0),not(is_free Q 0)) mengj@17998: val tpq = Term.type_of1(t1::Bnds, P$Q) mengj@17998: in mengj@17998: if(nfreeP andalso nfreeQ) then ( mengj@17998: let val tK = Type("fun",[tpq,Type("fun",[t1,tpq])]) mengj@17998: val PQ' = decre_bndVar(P $ Q) mengj@17998: in mengj@18276: include_min_comb:=true; mengj@17998: Const("COMBK",tK) $ PQ' mengj@17998: end) mengj@17998: else ( mengj@17998: if nfreeP then ( mengj@17998: let val Q' = lam2comb (Abs(x,t1,Q)) Bnds mengj@17998: val P' = decre_bndVar P mengj@17998: val tp = Term.type_of1(Bnds,P') mengj@17998: val tq' = Term.type_of1(Bnds, Q') mengj@17998: val tB = Type("fun",[tp,Type("fun",[tq',Type("fun",[t1,tpq])])]) mengj@17998: in mengj@18276: include_min_comb:=true; mengj@17998: Const("COMBB",tB) $ P' $ Q' mengj@17998: end) mengj@17998: else ( mengj@17998: if nfreeQ then ( mengj@17998: let val P' = lam2comb (Abs(x,t1,P)) Bnds mengj@17998: val Q' = decre_bndVar Q mengj@17998: val tq = Term.type_of1(Bnds,Q') mengj@17998: val tp' = Term.type_of1(Bnds, P') mengj@17998: val tC = Type("fun",[tp',Type("fun",[tq,Type("fun",[t1,tpq])])]) mengj@17998: in mengj@18276: include_min_comb:=true; mengj@17998: Const("COMBC",tC) $ P' $ Q' mengj@17998: end) mengj@17998: else( mengj@17998: let val P' = lam2comb (Abs(x,t1,P)) Bnds mengj@17998: val Q' = lam2comb (Abs(x,t1,Q)) Bnds mengj@17998: val tp' = Term.type_of1(Bnds,P') mengj@17998: val tq' = Term.type_of1(Bnds,Q') mengj@17998: val tS = Type("fun",[tp',Type("fun",[tq',Type("fun",[t1,tpq])])]) mengj@17998: in mengj@18276: include_min_comb:=true; mengj@18276: include_combS:=true; mengj@17998: Const("COMBS",tS) $ P' $ Q' mengj@17998: end))) mengj@17998: end mengj@17998: | lam2comb (t as (Abs(x,t1,_))) _ = raise LAM2COMB (t); mengj@17998: mengj@17998: mengj@17998: mengj@17998: (*********************) mengj@17998: mengj@17998: fun to_comb (Abs(x,tp,b)) Bnds = mengj@17998: let val b' = to_comb b (tp::Bnds) mengj@17998: in lam2comb (Abs(x,tp,b')) Bnds end mengj@17998: | to_comb (P $ Q) Bnds = ((to_comb P Bnds) $ (to_comb Q Bnds)) mengj@17998: | to_comb t _ = t; mengj@17998: mengj@17998: mengj@17998: fun comb_of t = to_comb t []; mengj@17998: mengj@17998: mengj@17998: (* print a term containing combinators, used for debugging *) mengj@17998: exception TERM_COMB of term; mengj@17998: mengj@17998: fun string_of_term (Const(c,t)) = c mengj@17998: | string_of_term (Free(v,t)) = v mengj@17998: | string_of_term (Var((x,n),t)) = mengj@17998: let val xn = x ^ "_" ^ (string_of_int n) mengj@17998: in xn end mengj@17998: | string_of_term (P $ Q) = mengj@17998: let val P' = string_of_term P mengj@17998: val Q' = string_of_term Q mengj@17998: in mengj@17998: "(" ^ P' ^ " " ^ Q' ^ ")" end mengj@17998: | string_of_term t = raise TERM_COMB (t); mengj@17998: mengj@17998: mengj@17998: mengj@17998: (******************************************************) mengj@17998: (* data types for typed combinator expressions *) mengj@17998: (******************************************************) mengj@17998: mengj@17998: type axiom_name = string; mengj@17998: datatype kind = Axiom | Conjecture; mengj@17998: fun name_of_kind Axiom = "axiom" mengj@17998: | name_of_kind Conjecture = "conjecture"; mengj@17998: mengj@17998: type polarity = bool; mengj@17998: type indexname = Term.indexname; mengj@17998: type clause_id = int; mengj@17998: type csort = Term.sort; mengj@18440: type ctyp = ResClause.fol_type; mengj@18440: mengj@18440: val string_of_ctyp = ResClause.string_of_fol_type; mengj@17998: mengj@17998: type ctyp_var = ResClause.typ_var; mengj@17998: mengj@17998: type ctype_literal = ResClause.type_literal; mengj@17998: mengj@17998: mengj@18356: datatype combterm = CombConst of string * ctyp * ctyp list mengj@17998: | CombFree of string * ctyp mengj@17998: | CombVar of string * ctyp mengj@17998: | CombApp of combterm * combterm * ctyp mengj@17998: | Bool of combterm mengj@17998: | Equal of combterm * combterm; mengj@17998: datatype literal = Literal of polarity * combterm; mengj@17998: mengj@17998: mengj@17998: mengj@17998: datatype clause = mengj@17998: Clause of {clause_id: clause_id, mengj@17998: axiom_name: axiom_name, mengj@17998: kind: kind, mengj@17998: literals: literal list, mengj@17998: ctypes_sorts: (ctyp_var * csort) list, mengj@17998: ctvar_type_literals: ctype_literal list, mengj@17998: ctfree_type_literals: ctype_literal list}; mengj@17998: mengj@17998: mengj@17998: mengj@17998: fun string_of_kind (Clause cls) = name_of_kind (#kind cls); mengj@17998: fun get_axiomName (Clause cls) = #axiom_name cls; mengj@17998: fun get_clause_id (Clause cls) = #clause_id cls; mengj@17998: mengj@18440: fun get_literals (c as Clause(cls)) = #literals cls; mengj@17998: mengj@17998: mengj@17998: mengj@18440: exception TERM_ORD of string mengj@18440: mengj@18440: fun term_ord (CombVar(_,_),CombVar(_,_)) = EQUAL mengj@18440: | term_ord (CombVar(_,_),_) = LESS mengj@18440: | term_ord (CombFree(_,_),CombVar(_,_)) = GREATER mengj@18440: | term_ord (CombFree(f1,tp1),CombFree(f2,tp2)) = mengj@18440: let val ord1 = string_ord(f1,f2) mengj@18440: in mengj@18440: case ord1 of EQUAL => ResClause.types_ord ([tp1],[tp2]) mengj@18440: | _ => ord1 mengj@18440: end mengj@18440: | term_ord (CombFree(_,_),_) = LESS mengj@18440: | term_ord (CombConst(_,_,_),CombVar(_,_)) = GREATER mengj@18440: | term_ord (CombConst(_,_,_),CombFree(_,_)) = GREATER mengj@18440: | term_ord (CombConst(c1,tp1,_),CombConst(c2,tp2,_)) = mengj@18440: let val ord1 = string_ord (c1,c2) mengj@18440: in mengj@18440: case ord1 of EQUAL => ResClause.types_ord ([tp1],[tp2]) mengj@18440: | _ => ord1 mengj@18440: end mengj@18440: | term_ord (CombConst(_,_,_),_) = LESS mengj@18440: | term_ord (CombApp(_,_,_),Bool(_)) = raise TERM_ORD("bool") mengj@18440: | term_ord (CombApp(_,_,_),Equal(_,_)) = LESS mengj@18440: | term_ord (CombApp(f1,arg1,tp1),CombApp(f2,arg2,tp2)) = mengj@18440: let val ord1 = term_ord (f1,f2) mengj@18440: val ord2 = case ord1 of EQUAL => term_ord (arg1,arg2) mengj@18440: | _ => ord1 mengj@18440: in mengj@18440: case ord2 of EQUAL => ResClause.types_ord ([tp1],[tp2]) mengj@18440: | _ => ord2 mengj@18440: end mengj@18440: | term_ord (CombApp(_,_,_),_) = GREATER mengj@18440: | term_ord (Bool(_),_) = raise TERM_ORD("bool") mengj@18440: | term_ord (Equal(t1,t2),Equal(t3,t4)) = ResClause.list_ord term_ord ([t1,t2],[t3,t4]) mengj@18440: | term_ord (Equal(_,_),_) = GREATER; mengj@18440: mengj@18440: fun predicate_ord (Equal(_,_),Bool(_)) = LESS mengj@18440: | predicate_ord (Equal(t1,t2),Equal(t3,t4)) = mengj@18440: ResClause.list_ord term_ord ([t1,t2],[t3,t4]) mengj@18440: | predicate_ord (Bool(_),Equal(_,_)) = GREATER mengj@18440: | predicate_ord (Bool(t1),Bool(t2)) = term_ord (t1,t2) mengj@18440: mengj@18440: mengj@18440: fun literal_ord (Literal(false,_),Literal(true,_)) = LESS mengj@18440: | literal_ord (Literal(true,_),Literal(false,_)) = GREATER mengj@18440: | literal_ord (Literal(_,pred1),Literal(_,pred2)) = predicate_ord(pred1,pred2); mengj@18440: mengj@18440: fun sort_lits lits = sort literal_ord lits; mengj@18440: mengj@17998: (*********************************************************************) mengj@17998: (* convert a clause with type Term.term to a clause with type clause *) mengj@17998: (*********************************************************************) mengj@17998: mengj@18356: fun isFalse (Literal(pol,Bool(CombConst(c,_,_)))) = mengj@17998: (pol andalso c = "c_False") orelse mengj@17998: (not pol andalso c = "c_True") mengj@17998: | isFalse _ = false; mengj@17998: mengj@17998: mengj@18356: fun isTrue (Literal (pol,Bool(CombConst(c,_,_)))) = mengj@17998: (pol andalso c = "c_True") orelse mengj@17998: (not pol andalso c = "c_False") mengj@17998: | isTrue _ = false; mengj@17998: mengj@17998: fun isTaut (Clause {literals,...}) = exists isTrue literals; mengj@17998: mengj@17998: mengj@17998: mengj@17998: fun make_clause(clause_id,axiom_name,kind,literals,ctypes_sorts,ctvar_type_literals,ctfree_type_literals) = mengj@17998: if forall isFalse literals mengj@17998: then error "Problem too trivial for resolution (empty clause)" mengj@17998: else mengj@17998: Clause {clause_id = clause_id, axiom_name = axiom_name, kind = kind, mengj@17998: literals = literals, ctypes_sorts = ctypes_sorts, mengj@17998: ctvar_type_literals = ctvar_type_literals, mengj@17998: ctfree_type_literals = ctfree_type_literals}; mengj@17998: mengj@18440: fun type_of (Type (a, Ts)) = mengj@18440: let val (folTypes,ts) = types_of Ts mengj@17998: val t = ResClause.make_fixed_type_const a mengj@17998: in mengj@18440: (ResClause.mk_fol_type("Comp",t,folTypes),ts) mengj@17998: end mengj@18440: | type_of (tp as (TFree(a,s))) = mengj@18440: let val t = ResClause.make_fixed_type_var a mengj@18440: in mengj@18440: (ResClause.mk_fol_type("Fixed",t,[]),[ResClause.mk_typ_var_sort tp]) mengj@18440: end mengj@18440: | type_of (tp as (TVar(v,s))) = mengj@18440: let val t = ResClause.make_schematic_type_var v mengj@18440: in mengj@18440: (ResClause.mk_fol_type("Var",t,[]),[ResClause.mk_typ_var_sort tp]) mengj@18440: end mengj@18356: mengj@18440: and types_of Ts = mengj@18440: let val foltyps_ts = map type_of Ts mengj@18440: val (folTyps,ts) = ListPair.unzip foltyps_ts mengj@18440: in mengj@18440: (folTyps,ResClause.union_all ts) mengj@18440: end; mengj@17998: mengj@17998: (* same as above, but no gathering of sort information *) mengj@18440: fun simp_type_of (Type (a, Ts)) = mengj@18356: let val typs = map simp_type_of Ts mengj@17998: val t = ResClause.make_fixed_type_const a mengj@17998: in mengj@18440: ResClause.mk_fol_type("Comp",t,typs) mengj@17998: end mengj@18440: | simp_type_of (TFree (a,s)) = ResClause.mk_fol_type("Fixed",ResClause.make_fixed_type_var a,[]) mengj@18440: | simp_type_of (TVar (v,s)) = ResClause.mk_fol_type("Var",ResClause.make_schematic_type_var v,[]); mengj@18440: mengj@18356: fun comb_typ ("COMBI",t) = mengj@18356: let val t' = domain_type t mengj@18356: in mengj@18356: [simp_type_of t'] mengj@18356: end mengj@18356: | comb_typ ("COMBK",t) = mengj@18725: let val a = domain_type t mengj@18725: val b = domain_type (range_type t) mengj@18356: in mengj@18725: map simp_type_of [a,b] mengj@18356: end mengj@18356: | comb_typ ("COMBS",t) = mengj@18356: let val t' = domain_type t mengj@18725: val a = domain_type t' mengj@18725: val b = domain_type (range_type t') mengj@18725: val c = range_type (range_type t') mengj@18356: in mengj@18356: map simp_type_of [a,b,c] mengj@18356: end mengj@18356: | comb_typ ("COMBB",t) = mengj@18725: let val ab = domain_type t mengj@18725: val ca = domain_type (range_type t) mengj@18356: val a = domain_type ab mengj@18725: val b = range_type ab mengj@18725: val c = domain_type ca mengj@18356: in mengj@18356: map simp_type_of [a,b,c] mengj@18356: end mengj@18356: | comb_typ ("COMBC",t) = mengj@18356: let val t1 = domain_type t mengj@18725: val a = domain_type t1 mengj@18725: val b = domain_type (range_type t1) mengj@18725: val c = range_type (range_type t1) mengj@18356: in mengj@18356: map simp_type_of [a,b,c] mengj@18356: end; mengj@18356: mengj@18356: fun const_type_of ("COMBI",t) = mengj@18356: let val (tp,ts) = type_of t mengj@18356: val I_var = comb_typ ("COMBI",t) mengj@18356: in mengj@18356: (tp,ts,I_var) mengj@18356: end mengj@18356: | const_type_of ("COMBK",t) = mengj@18356: let val (tp,ts) = type_of t mengj@18356: val K_var = comb_typ ("COMBK",t) mengj@18356: in mengj@18356: (tp,ts,K_var) mengj@18356: end mengj@18356: | const_type_of ("COMBS",t) = mengj@18356: let val (tp,ts) = type_of t mengj@18356: val S_var = comb_typ ("COMBS",t) mengj@18356: in mengj@18356: (tp,ts,S_var) mengj@18356: end mengj@18356: | const_type_of ("COMBB",t) = mengj@18356: let val (tp,ts) = type_of t mengj@18356: val B_var = comb_typ ("COMBB",t) mengj@18356: in mengj@18356: (tp,ts,B_var) mengj@18356: end mengj@18356: | const_type_of ("COMBC",t) = mengj@18356: let val (tp,ts) = type_of t mengj@18356: val C_var = comb_typ ("COMBC",t) mengj@18356: in mengj@18356: (tp,ts,C_var) mengj@18356: end mengj@18356: | const_type_of (c,t) = mengj@18356: let val (tp,ts) = type_of t mengj@18356: val tvars = !const_typargs(c,t) mengj@18356: val tvars' = map simp_type_of tvars mengj@18356: in mengj@18356: (tp,ts,tvars') mengj@18356: end; mengj@18356: mengj@18356: fun is_bool_type (Type("bool",[])) = true mengj@18356: | is_bool_type _ = false; mengj@17998: mengj@17998: mengj@17998: (* convert a Term.term (with combinators) into a combterm, also accummulate sort info *) mengj@17998: fun combterm_of (Const(c,t)) = mengj@18356: let val (tp,ts,tvar_list) = const_type_of (c,t) mengj@18356: val is_bool = is_bool_type t mengj@18356: val c' = CombConst(ResClause.make_fixed_const c,tp,tvar_list) mengj@17998: val c'' = if is_bool then Bool(c') else c' mengj@17998: in mengj@17998: (c'',ts) mengj@17998: end mengj@17998: | combterm_of (Free(v,t)) = mengj@18356: let val (tp,ts) = type_of t mengj@18356: val is_bool = is_bool_type t mengj@17998: val v' = if ResClause.isMeta v then CombVar(ResClause.make_schematic_var(v,0),tp) mengj@17998: else CombFree(ResClause.make_fixed_var v,tp) mengj@17998: val v'' = if is_bool then Bool(v') else v' mengj@17998: in mengj@17998: (v'',ts) mengj@17998: end mengj@17998: | combterm_of (Var(v,t)) = mengj@18356: let val (tp,ts) = type_of t mengj@18356: val is_bool = is_bool_type t mengj@17998: val v' = CombVar(ResClause.make_schematic_var v,tp) mengj@17998: val v'' = if is_bool then Bool(v') else v' mengj@17998: in mengj@17998: (v'',ts) mengj@17998: end mengj@17998: | combterm_of (Const("op =",T) $ P $ Q) = (*FIXME: allow equal between bools?*) mengj@17998: let val (P',tsP) = combterm_of P mengj@17998: val (Q',tsQ) = combterm_of Q mengj@17998: in mengj@17998: (Equal(P',Q'),tsP union tsQ) mengj@17998: end mengj@17998: | combterm_of (t as (P $ Q)) = mengj@17998: let val (P',tsP) = combterm_of P mengj@17998: val (Q',tsQ) = combterm_of Q mengj@17998: val tp = Term.type_of t mengj@18356: val tp' = simp_type_of tp mengj@18356: val is_bool = is_bool_type tp mengj@17998: val t' = CombApp(P',Q',tp') mengj@17998: val t'' = if is_bool then Bool(t') else t' mengj@17998: in mengj@17998: (t'',tsP union tsQ) mengj@17998: end; mengj@17998: mengj@17998: fun predicate_of ((Const("Not",_) $ P), polarity) = mengj@17998: predicate_of (P, not polarity) mengj@17998: | predicate_of (term,polarity) = (combterm_of term,polarity); mengj@17998: mengj@17998: mengj@17998: fun literals_of_term1 args (Const("Trueprop",_) $ P) = literals_of_term1 args P mengj@17998: | literals_of_term1 args (Const("op |",_) $ P $ Q) = mengj@17998: let val args' = literals_of_term1 args P mengj@17998: in mengj@17998: literals_of_term1 args' Q mengj@17998: end mengj@17998: | literals_of_term1 (lits,ts) P = mengj@17998: let val ((pred,ts'),pol) = predicate_of (P,true) mengj@17998: val lits' = Literal(pol,pred)::lits mengj@17998: in mengj@17998: (lits',ts union ts') mengj@17998: end; mengj@17998: mengj@17998: mengj@17998: fun literals_of_term P = literals_of_term1 ([],[]) P; mengj@17998: mengj@17998: mengj@17998: (* making axiom and conjecture clauses *) mengj@19444: fun make_axiom_clause thm (ax_name,cls_id) = mengj@19444: let val term = prop_of thm mengj@19444: val term' = comb_of term mengj@17998: val (lits,ctypes_sorts) = literals_of_term term' mengj@18440: val lits' = sort_lits lits paulson@18856: val (ctvar_lits,ctfree_lits) = ResClause.add_typs_aux ctypes_sorts mengj@17998: in mengj@17998: make_clause(cls_id,ax_name,Axiom, mengj@18440: lits',ctypes_sorts,ctvar_lits,ctfree_lits) mengj@17998: end; mengj@17998: mengj@19444: fun make_axiom_clauses [] = [] mengj@19444: | make_axiom_clauses ((thm,(name,id))::thms) = mengj@19444: let val cls = make_axiom_clause thm (name,id) mengj@19444: val clss = make_axiom_clauses thms mengj@19198: in mengj@19198: if isTaut cls then clss else (cls::clss) mengj@19198: end; mengj@17998: mengj@19354: mengj@19444: fun make_conjecture_clause n thm = mengj@19444: let val t = prop_of thm mengj@19444: val t' = comb_of t mengj@17998: val (lits,ctypes_sorts) = literals_of_term t' paulson@18856: val (ctvar_lits,ctfree_lits) = ResClause.add_typs_aux ctypes_sorts mengj@17998: in mengj@17998: make_clause(n,"conjecture",Conjecture,lits,ctypes_sorts,ctvar_lits,ctfree_lits) mengj@17998: end; mengj@17998: mengj@17998: mengj@17998: mengj@17998: fun make_conjecture_clauses_aux _ [] = [] mengj@17998: | make_conjecture_clauses_aux n (t::ts) = mengj@17998: make_conjecture_clause n t :: make_conjecture_clauses_aux (n+1) ts; mengj@17998: mengj@17998: val make_conjecture_clauses = make_conjecture_clauses_aux 0; mengj@17998: mengj@17998: mengj@17998: (**********************************************************************) mengj@17998: (* convert clause into ATP specific formats: *) mengj@17998: (* TPTP used by Vampire and E *) mengj@17998: (**********************************************************************) mengj@17998: mengj@17998: val type_wrapper = "typeinfo"; mengj@17998: mengj@18356: datatype type_level = T_FULL | T_PARTIAL | T_CONST | T_NONE; mengj@17998: mengj@19130: val typ_level = ref T_FULL; mengj@18356: mengj@18356: fun full_types () = (typ_level:=T_FULL); mengj@18356: fun partial_types () = (typ_level:=T_PARTIAL); mengj@18356: fun const_types_only () = (typ_level:=T_CONST); mengj@18356: fun no_types () = (typ_level:=T_NONE); mengj@18356: mengj@18356: mengj@18356: fun find_typ_level () = !typ_level; mengj@18356: mengj@18356: fun wrap_type (c,t) = mengj@18356: case !typ_level of T_FULL => type_wrapper ^ (ResClause.paren_pack [c,t]) mengj@18356: | _ => c; mengj@18356: mengj@17998: mengj@17998: val bool_tp = ResClause.make_fixed_type_const "bool"; mengj@17998: mengj@17998: val app_str = "hAPP"; mengj@17998: mengj@17998: val bool_str = "hBOOL"; mengj@17998: mengj@18356: exception STRING_OF_COMBTERM of int; mengj@17998: mengj@17998: (* convert literals of clauses into strings *) mengj@18440: fun string_of_combterm1_aux _ (CombConst(c,tp,_)) = mengj@18440: let val tp' = string_of_ctyp tp mengj@18440: in mengj@18440: (wrap_type (c,tp'),tp') mengj@18440: end mengj@18440: | string_of_combterm1_aux _ (CombFree(v,tp)) = mengj@18440: let val tp' = string_of_ctyp tp mengj@18440: in mengj@18440: (wrap_type (v,tp'),tp') mengj@18440: end mengj@18440: | string_of_combterm1_aux _ (CombVar(v,tp)) = mengj@18440: let val tp' = string_of_ctyp tp mengj@18440: in mengj@18440: (wrap_type (v,tp'),tp') mengj@18440: end mengj@18356: | string_of_combterm1_aux is_pred (CombApp(t1,t2,tp)) = mengj@18356: let val (s1,tp1) = string_of_combterm1_aux is_pred t1 mengj@18356: val (s2,tp2) = string_of_combterm1_aux is_pred t2 mengj@18440: val tp' = ResClause.string_of_fol_type tp mengj@18440: val r = case !typ_level of T_FULL => type_wrapper ^ (ResClause.paren_pack [(app_str ^ (ResClause.paren_pack [s1,s2])),tp']) mengj@18356: | T_PARTIAL => app_str ^ (ResClause.paren_pack [s1,s2,tp1]) mengj@18356: | T_NONE => app_str ^ (ResClause.paren_pack [s1,s2]) mengj@18356: | T_CONST => raise STRING_OF_COMBTERM (1) (*should not happen, if happened may be a bug*) mengj@18440: in (r,tp') mengj@18356: mengj@17998: end mengj@18356: | string_of_combterm1_aux is_pred (Bool(t)) = mengj@18356: let val (t',_) = string_of_combterm1_aux false t mengj@18356: val r = if is_pred then bool_str ^ (ResClause.paren_pack [t']) mengj@18356: else t' mengj@18356: in mengj@18356: (r,bool_tp) mengj@18356: end mengj@18356: | string_of_combterm1_aux _ (Equal(t1,t2)) = mengj@18356: let val (s1,_) = string_of_combterm1_aux false t1 mengj@18356: val (s2,_) = string_of_combterm1_aux false t2 mengj@18356: in mengj@18356: ("equal" ^ (ResClause.paren_pack [s1,s2]),bool_tp) mengj@18356: end; mengj@18356: mengj@18356: fun string_of_combterm1 is_pred term = fst (string_of_combterm1_aux is_pred term); mengj@18356: mengj@18440: fun string_of_combterm2 _ (CombConst(c,tp,tvars)) = mengj@18440: let val tvars' = map string_of_ctyp tvars mengj@18440: in mengj@18440: c ^ (ResClause.paren_pack tvars') mengj@18440: end mengj@18356: | string_of_combterm2 _ (CombFree(v,tp)) = v mengj@18356: | string_of_combterm2 _ (CombVar(v,tp)) = v mengj@18356: | string_of_combterm2 is_pred (CombApp(t1,t2,tp)) = mengj@18356: let val s1 = string_of_combterm2 is_pred t1 mengj@18356: val s2 = string_of_combterm2 is_pred t2 mengj@18356: in mengj@18356: app_str ^ (ResClause.paren_pack [s1,s2]) mengj@18356: end mengj@18356: | string_of_combterm2 is_pred (Bool(t)) = mengj@18356: let val t' = string_of_combterm2 false t mengj@17998: in mengj@18200: if is_pred then bool_str ^ (ResClause.paren_pack [t']) mengj@18200: else t' mengj@17998: end mengj@18356: | string_of_combterm2 _ (Equal(t1,t2)) = mengj@18356: let val s1 = string_of_combterm2 false t1 mengj@18356: val s2 = string_of_combterm2 false t2 mengj@17998: in mengj@18356: ("equal" ^ (ResClause.paren_pack [s1,s2])) mengj@17998: end; mengj@17998: mengj@18356: mengj@18356: mengj@18356: fun string_of_combterm is_pred term = mengj@18356: case !typ_level of T_CONST => string_of_combterm2 is_pred term mengj@18356: | _ => string_of_combterm1 is_pred term; mengj@18356: mengj@18356: mengj@17998: fun string_of_clausename (cls_id,ax_name) = mengj@17998: ResClause.clause_prefix ^ ResClause.ascii_of ax_name ^ "_" ^ Int.toString cls_id; mengj@17998: mengj@17998: fun string_of_type_clsname (cls_id,ax_name,idx) = mengj@17998: string_of_clausename (cls_id,ax_name) ^ "_tcs" ^ (Int.toString idx); mengj@17998: mengj@17998: mengj@17998: fun tptp_literal (Literal(pol,pred)) = mengj@18200: let val pred_string = string_of_combterm true pred mengj@17998: val pol_str = if pol then "++" else "--" mengj@17998: in mengj@17998: pol_str ^ pred_string mengj@17998: end; mengj@17998: mengj@17998: mengj@17998: fun tptp_type_lits (Clause cls) = mengj@17998: let val lits = map tptp_literal (#literals cls) mengj@17998: val ctvar_lits_strs = mengj@18356: case !typ_level of T_NONE => [] mengj@18356: | _ => (map ResClause.tptp_of_typeLit (#ctvar_type_literals cls)) mengj@17998: val ctfree_lits = mengj@18356: case !typ_level of T_NONE => [] mengj@18356: | _ => (map ResClause.tptp_of_typeLit (#ctfree_type_literals cls)) mengj@17998: in mengj@17998: (ctvar_lits_strs @ lits, ctfree_lits) mengj@17998: end; mengj@18356: mengj@18356: mengj@17998: fun clause2tptp cls = mengj@17998: let val (lits,ctfree_lits) = tptp_type_lits cls mengj@17998: val cls_id = get_clause_id cls mengj@17998: val ax_name = get_axiomName cls mengj@17998: val knd = string_of_kind cls mengj@17998: val lits_str = ResClause.bracket_pack lits mengj@17998: val cls_str = ResClause.gen_tptp_cls(cls_id,ax_name,knd,lits_str) mengj@17998: in mengj@17998: (cls_str,ctfree_lits) mengj@17998: end; mengj@17998: mengj@17998: mengj@18440: mengj@18440: (**********************************************************************) mengj@18440: (* clause equalities and hashing functions *) mengj@18440: (**********************************************************************) mengj@18440: mengj@18440: mengj@18440: fun combterm_eq (CombConst(c1,tp1,tps1),CombConst(c2,tp2,tps2)) vtvars = mengj@18440: let val (eq1,vtvars1) = if c1 = c2 then ResClause.types_eq (tps1,tps2) vtvars mengj@18440: else (false,vtvars) mengj@18440: in mengj@18440: (eq1,vtvars1) mengj@18440: end mengj@18440: | combterm_eq (CombConst(_,_,_),_) vtvars = (false,vtvars) mengj@18440: | combterm_eq (CombFree(a1,tp1),CombFree(a2,tp2)) vtvars = mengj@18440: if a1 = a2 then ResClause.types_eq ([tp1],[tp2]) vtvars mengj@18440: else (false,vtvars) mengj@18440: | combterm_eq (CombFree(_,_),_) vtvars = (false,vtvars) mengj@18440: | combterm_eq (CombVar(v1,tp1),CombVar(v2,tp2)) (vars,tvars) = mengj@18440: (case ResClause.check_var_pairs(v1,v2) vars of 0 => ResClause.types_eq ([tp1],[tp2]) ((v1,v2)::vars,tvars) mengj@18440: | 1 => ResClause.types_eq ([tp1],[tp2]) (vars,tvars) mengj@18440: | 2 => (false,(vars,tvars))) mengj@18440: | combterm_eq (CombVar(_,_),_) vtvars = (false,vtvars) mengj@18440: | combterm_eq (CombApp(f1,arg1,tp1),CombApp(f2,arg2,tp2)) vtvars = mengj@18440: let val (eq1,vtvars1) = combterm_eq (f1,f2) vtvars mengj@18440: val (eq2,vtvars2) = if eq1 then combterm_eq (arg1,arg2) vtvars1 mengj@18440: else (eq1,vtvars1) mengj@18440: in mengj@18440: if eq2 then ResClause.types_eq ([tp1],[tp2]) vtvars2 mengj@18440: else (eq2,vtvars2) mengj@18440: end mengj@18440: | combterm_eq (CombApp(_,_,_),_) vtvars = (false,vtvars) mengj@18440: | combterm_eq (Bool(t1),Bool(t2)) vtvars = combterm_eq (t1,t2) vtvars mengj@18440: | combterm_eq (Bool(_),_) vtvars = (false,vtvars) mengj@18440: | combterm_eq (Equal(t1,t2),Equal(t3,t4)) vtvars = mengj@18440: let val (eq1,vtvars1) = combterm_eq (t1,t3) vtvars mengj@18440: in mengj@18440: if eq1 then combterm_eq (t2,t4) vtvars1 mengj@18440: else (eq1,vtvars1) mengj@18440: end mengj@18440: | combterm_eq (Equal(t1,t2),_) vtvars = (false,vtvars); mengj@18440: mengj@18440: fun lit_eq (Literal(pol1,pred1),Literal(pol2,pred2)) vtvars = mengj@18440: if (pol1 = pol2) then combterm_eq (pred1,pred2) vtvars mengj@18440: else (false,vtvars); mengj@18440: mengj@18440: fun lits_eq ([],[]) vtvars = (true,vtvars) mengj@18440: | lits_eq (l1::ls1,l2::ls2) vtvars = mengj@18440: let val (eq1,vtvars1) = lit_eq (l1,l2) vtvars mengj@18440: in mengj@18440: if eq1 then lits_eq (ls1,ls2) vtvars1 mengj@18440: else (false,vtvars1) mengj@18440: end; mengj@18440: mengj@18440: fun clause_eq (cls1,cls2) = mengj@18440: let val lits1 = get_literals cls1 mengj@18440: val lits2 = get_literals cls2 mengj@18440: in mengj@18440: length lits1 = length lits2 andalso #1 (lits_eq (lits1,lits2) ([],[])) mengj@18440: end; mengj@18440: mengj@18440: val xor_words = List.foldl Word.xorb 0w0; mengj@18440: mengj@18440: fun hash_combterm (CombVar(_,_),w) = w paulson@18449: | hash_combterm (CombFree(f,_),w) = Polyhash.hashw_string(f,w) paulson@18449: | hash_combterm (CombConst(c,tp,tps),w) = Polyhash.hashw_string(c,w) paulson@18449: | hash_combterm (CombApp(f,arg,tp),w) = hash_combterm (arg, hash_combterm (f,w)) mengj@18440: | hash_combterm (Bool(t),w) = hash_combterm (t,w) mengj@18440: | hash_combterm (Equal(t1,t2),w) = paulson@18449: List.foldl hash_combterm (Polyhash.hashw_string ("equal",w)) [t1,t2] mengj@18440: mengj@18440: fun hash_literal (Literal(true,pred)) = hash_combterm(pred,0w0) mengj@18440: | hash_literal (Literal(false,pred)) = Word.notb(hash_combterm(pred,0w0)); mengj@18440: mengj@18440: fun hash_clause clause = xor_words (map hash_literal (get_literals clause)); mengj@18440: mengj@19198: (**********************************************************************) mengj@19198: (* write clauses to files *) mengj@19198: (**********************************************************************) mengj@19198: mengj@19198: fun read_in [] = [] mengj@19198: | read_in (f1::fs) = mengj@19198: let val lines = read_in fs mengj@19198: val input = TextIO.openIn f1 mengj@19198: fun reading () = mengj@19198: let val nextline = TextIO.inputLine input mengj@19198: in mengj@19198: if nextline = "" then (TextIO.closeIn input; []) mengj@19198: else nextline::(reading ()) mengj@19198: end mengj@19198: in mengj@19198: ((reading ()) @ lines) mengj@19198: end; mengj@19198: mengj@19198: mengj@19198: fun get_helper_clauses (full,partial,const,untyped) = mengj@19198: let val (helper1,noS,inclS) = case !typ_level of T_FULL => (warning "Fully-typed HOL"; full) mengj@19198: | T_PARTIAL => (warning "Partially-typed HOL"; partial) mengj@19198: | T_CONST => (warning "Const-only-typed HOL"; const) mengj@19198: | T_NONE => (warning "Untyped HOL"; untyped) mengj@19198: val needed_helpers = if !include_combS then (warning "Include combinator S"; [helper1,inclS]) else mengj@19198: if !include_min_comb then (warning "Include min combinators"; [helper1,noS]) mengj@19198: else (warning "No combinator is used"; [helper1]) mengj@19198: in mengj@19198: read_in needed_helpers mengj@19198: end; mengj@19198: mengj@19198: mengj@19198: (* write TPTP format to a single file *) mengj@19198: (* when "get_helper_clauses" is called, "include_combS" and "include_min_comb" should have correct values already *) mengj@19444: fun tptp_write_file thms filename (axclauses,classrel_clauses,arity_clauses) helpers = mengj@19444: let val clss = make_conjecture_clauses thms mengj@19444: val axclauses' = make_axiom_clauses axclauses mengj@19198: val (tptp_clss,tfree_litss) = ListPair.unzip (map clause2tptp clss) mengj@19198: val tfree_clss = map ResClause.tptp_tfree_clause (foldl (op union_string) [] tfree_litss) mengj@19198: val out = TextIO.openOut filename mengj@19198: val helper_clauses = get_helper_clauses helpers mengj@19198: in mengj@19198: List.app (curry TextIO.output out o #1 o clause2tptp) axclauses'; mengj@19198: ResClause.writeln_strs out tfree_clss; mengj@19198: ResClause.writeln_strs out tptp_clss; mengj@19198: List.app (curry TextIO.output out o ResClause.tptp_classrelClause) classrel_clauses; mengj@19198: List.app (curry TextIO.output out o ResClause.tptp_arity_clause) arity_clauses; mengj@19198: List.app (curry TextIO.output out) helper_clauses; mengj@19198: TextIO.closeOut out mengj@19198: end; mengj@19198: mengj@17998: end