walther@59603: (* title: ProgLang/Prog_Expr.thy walther@59603: functions for expressions (which do NOT contain Prog_Tac and Tactical, by definition)) walther@59603: author: Walther Neuper, Aug.2019 walther@59603: (c) copyright due to lincense terms. walther@59603: *) walther@59603: walther@59619: theory Prog_Expr walther@59630: imports Calculate ListC walther@59603: begin walther@59603: walther@59619: text \Abstract: walther@59619: Specific constants are defined for pre-conditions of programs and for program-expressions. walther@59619: The constants are connected with ML functions for evaluation of respective expressions walther@59619: (all named eval_*), such that Isac's rewrite engine can handle it. walther@59619: \ walther@59619: walther@59619: subsection \Power re-defined for a specific type\ walther@59603: consts walther@59603: pow :: "[real, real] => real" (infixr "^^^" 80) walther@59603: walther@59603: subsection \consts of functions in pre-conditions and program-expressions\ walther@59603: consts walther@59603: lhs :: "bool => real" (*of an equality*) walther@59603: rhs :: "bool => real" (*of an equality*) walther@59603: Vars :: "'a => real list" (*get the variables of a term *) walther@59603: matches :: "['a, 'a] => bool" walther@59603: matchsub :: "['a, 'a] => bool" walther@59603: some'_occur'_in :: "[real list, 'a] => bool" ("some'_of _ occur'_in _") walther@59603: occurs'_in :: "[real , 'a] => bool" ("_ occurs'_in _") walther@59603: walther@59603: abs :: "real => real" ("(|| _ ||)") walther@59619: (* ~~~ FIXXXME Isabelle2002 has abs already !!!*) walther@59603: absset :: "real set => real" ("(||| _ |||)") walther@59603: (*is numeral constant ?*) walther@59603: is'_const :: "real => bool" ("_ is'_const" 10) walther@59603: (*is_const rename to is_num FIXXXME.WN.16.5.03 *) walther@59603: is'_atom :: "real => bool" ("_ is'_atom" 10) walther@59603: is'_even :: "real => bool" ("_ is'_even" 10) walther@59603: walther@59603: (* identity on term level*) walther@59603: ident :: "['a, 'a] => bool" ("(_ =!=/ _)" [51, 51] 50) walther@59603: argument'_in :: "real => real" ("argument'_in _" 10) walther@59619: sameFunId :: "[real, bool] => bool" (* "same'_funid _ _" 10 walther@59619: WN0609 changed the id, because ".. _ _" inhibits currying *) walther@59603: filter'_sameFunId:: "[real, bool list] => bool list" walther@59603: ("filter'_sameFunId _ _" 10) walther@59603: boollist2sum :: "bool list => real" walther@59603: lastI :: "'a list \ 'a" walther@59603: walther@59603: subsection \ML code for functions in pre-conditions and program-expressions\ walther@59603: ML \ walther@59603: signature PROG_EXPR = walther@59603: sig walther@59603: (*//------------------------- from Tools .thy-------------------------------------------------\\*) walther@59603: val lhs: term -> term walther@59603: val eval_lhs: 'a -> string -> term -> 'b -> (string * term) option walther@59603: val eval_matches: string -> string -> term -> theory -> (string * term) option walther@59603: val matchsub: theory -> term -> term -> bool walther@59603: val eval_matchsub: string -> string -> term -> theory -> (string * term) option walther@59603: val rhs: term -> term walther@59603: val eval_rhs: 'a -> string -> term -> 'b -> (string * term) option walther@59603: val eval_var: string -> string -> term -> theory -> (string * term) option walther@59603: walther@59603: val or2list: term -> term walther@59603: (*\\------------------------- from Tools .thy-------------------------------------------------//*) walther@59603: (*//------------------------- from Atools .thy------------------------------------------------\\*) walther@59603: val occurs_in: term -> term -> bool walther@59603: val eval_occurs_in: 'a -> string -> term -> 'b -> (string * term) option walther@59603: val some_occur_in: term list -> term -> bool walther@59603: val eval_some_occur_in: 'a -> string -> term -> 'b -> (string * term) option walther@59603: val eval_is_atom: string -> string -> term -> 'a -> (string * term) option walther@59603: val even: int -> bool walther@59603: val eval_is_even: string -> string -> term -> 'a -> (string * term) option walther@59603: val eval_const: string -> string -> term -> 'a -> (string * term) option walther@59603: val eval_equ: string -> string -> term -> 'a -> (string * term) option walther@59603: val eval_ident: string -> string -> term -> theory -> (string * term) option walther@59603: val eval_equal: string -> string -> term -> theory -> (string * term) option walther@59603: val eval_cancel: string -> string -> term -> 'a -> (string * term) option walther@59603: val eval_argument_in: string -> string -> term -> 'a -> (string * term) option walther@59603: val same_funid: term -> term -> bool walther@59603: val eval_sameFunId: string -> string -> term -> 'a -> (string * term) option walther@59603: val eval_filter_sameFunId: string -> string -> term -> 'a -> (string * term) option walther@59603: val list2sum: term list -> term walther@59603: val eval_boollist2sum: string -> string -> term -> 'a -> (string * term) option walther@59603: walther@59603: val mk_thmid_f: string -> (int * int) * (int * int) -> (int * int) * (int * int) -> string walther@59603: (*\\------------------------- from Atools.thy------------------------------------------------//*) walther@59603: end walther@59603: walther@59603: (**) walther@59603: structure Prog_Expr(**): PROG_EXPR =(**) walther@59603: struct walther@59603: (**) walther@59603: walther@59603: (*+ for Or_to_List +*) walther@59620: fun or2list (Const ("HOL.True",_)) = (tracing"### or2list True"; UniversalList) walther@59620: | or2list (Const ("HOL.False",_)) = (tracing"### or2list False"; EmptyList) walther@59618: | or2list (t as Const ("HOL.eq",_) $ _ $ _) = TermC.list2isalist HOLogic.boolT [t] walther@59603: | or2list ors = walther@59618: let walther@59618: fun get ls (Const ("HOL.disj",_) $ o1 $ o2) = walther@59618: (case o2 of walther@59618: Const ("HOL.disj",_) $ _ $ _ => get (ls @ [o1]) o2 walther@59618: | _ => ls @ [o1, o2]) walther@59618: | get _ t = raise TERM ("or2list: missing pattern in fun.def", [t]) walther@59618: in (((TermC.list2isalist HOLogic.boolT) o (get [])) ors) end walther@59603: (*>val t = @{term True}; walther@59603: > val t' = or2list t; walther@59603: > term2str t'; walther@59620: "ListC.UniversalList" walther@59603: > val t = @{term False}; walther@59603: > val t' = or2list t; walther@59603: > term2str t'; walther@59603: "[]" walther@59603: > val t=(Thm.term_of o the o (parse thy)) "x=3"; walther@59603: > val t' = or2list t; walther@59603: > term2str t'; walther@59603: "[x = 3]" walther@59603: > val t=(Thm.term_of o the o (parse thy))"(x=3) | (x=-3) | (x=0)"; walther@59603: > val t' = or2list t; walther@59603: > term2str t'; walther@59603: "[x = #3, x = #-3, x = #0]" : string *) walther@59603: walther@59603: walther@59603: (** evaluation on the meta-level **) walther@59603: walther@59603: (*. evaluate the predicate matches (match on whole term only) .*) walther@59603: (*("matches",("Prog_Expr.matches", eval_matches "#matches_")):calc*) walther@59618: fun eval_matches (_:string) "Prog_Expr.matches" walther@59603: (t as Const ("Prog_Expr.matches",_) $ pat $ tst) thy = walther@59603: if TermC.matches thy tst pat walther@59603: then walther@59603: let walther@59603: val prop = HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True})) walther@59603: in SOME (Rule.term_to_string''' thy prop, prop) end walther@59603: else walther@59603: let walther@59603: val prop = HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False})) walther@59603: in SOME (Rule.term_to_string''' thy prop, prop) end walther@59603: | eval_matches _ _ _ _ = NONE; walther@59603: (* walther@59603: > val t = (Thm.term_of o the o (parse thy)) walther@59603: "matches (?x = 0) (1 * x ^^^ 2 = 0)"; walther@59603: > eval_matches "/thmid/" "/op_/" t thy; walther@59603: val it = walther@59603: SOME walther@59603: ("matches (x = 0) (1 * x ^^^ 2 = 0) = False", walther@59603: Const (#,#) $ (# $ # $ Const #)) : (string * term) option walther@59603: walther@59603: > val t = (Thm.term_of o the o (parse thy)) walther@59603: "matches (?a = #0) (#1 * x ^^^ #2 = #0)"; walther@59603: > eval_matches "/thmid/" "/op_/" t thy; walther@59603: val it = walther@59603: SOME walther@59603: ("matches (?a = #0) (#1 * x ^^^ #2 = #0) = True", walther@59603: Const (#,#) $ (# $ # $ Const #)) : (string * term) option walther@59603: walther@59603: > val t = (Thm.term_of o the o (parse thy)) walther@59603: "matches (?a * x = #0) (#1 * x ^^^ #2 = #0)"; walther@59603: > eval_matches "/thmid/" "/op_/" t thy; walther@59603: val it = walther@59603: SOME walther@59603: ("matches (?a * x = #0) (#1 * x ^^^ #2 = #0) = False", walther@59603: Const (#,#) $ (# $ # $ Const #)) : (string * term) option walther@59603: walther@59603: > val t = (Thm.term_of o the o (parse thy)) walther@59603: "matches (?a * x ^^^ #2 = #0) (#1 * x ^^^ #2 = #0)"; walther@59603: > eval_matches "/thmid/" "/op_/" t thy; walther@59603: val it = walther@59603: SOME walther@59603: ("matches (?a * x ^^^ #2 = #0) (#1 * x ^^^ #2 = #0) = True", walther@59603: Const (#,#) $ (# $ # $ Const #)) : (string * term) option walther@59603: ----- before ?patterns ---: walther@59603: > val t = (Thm.term_of o the o (parse thy)) walther@59603: "matches (a * b^^^#2 = c) (#3 * x^^^#2 = #1)"; walther@59603: > eval_matches "/thmid/" "/op_/" t thy; walther@59603: SOME walther@59603: ("matches (a * b ^^^ #2 = c) (#3 * x ^^^ #2 = #1) = True", walther@59603: Const ("HOL.Trueprop","bool => prop") $ (Const # $ (# $ #) $ Const (#,#))) walther@59603: : (string * term) option walther@59603: walther@59603: > val t = (Thm.term_of o the o (parse thy)) walther@59603: "matches (a * b^^^#2 = c) (#3 * x^^^#2222 = #1)"; walther@59603: > eval_matches "/thmid/" "/op_/" t thy; walther@59603: SOME ("matches (a * b ^^^ #2 = c) (#3 * x ^^^ #2222 = #1) = False", walther@59603: Const ("HOL.Trueprop","bool => prop") $ (Const # $ (# $ #) $ Const (#,#))) walther@59603: walther@59603: > val t = (Thm.term_of o the o (parse thy)) walther@59603: "matches (a = b) (x + #1 + #-1 * #2 = #0)"; walther@59603: > eval_matches "/thmid/" "/op_/" t thy; walther@59603: SOME ("matches (a = b) (x + #1 + #-1 * #2 = #0) = True",Const # $ (# $ #)) walther@59603: *) walther@59603: walther@59603: (*.does a pattern match some subterm ?.*) walther@59603: fun matchsub thy t pat = walther@59603: let walther@59603: fun matchs (t as Const _) = TermC.matches thy t pat walther@59603: | matchs (t as Free _) = TermC.matches thy t pat walther@59603: | matchs (t as Var _) = TermC.matches thy t pat walther@59603: | matchs (Bound _) = false walther@59603: | matchs (t as Abs (_, _, body)) = walther@59603: if TermC.matches thy t pat then true else TermC.matches thy body pat walther@59603: | matchs (t as f1 $ f2) = walther@59603: if TermC.matches thy t pat then true walther@59603: else if matchs f1 then true else matchs f2 walther@59603: in matchs t end; walther@59603: walther@59603: (*("matchsub",("Prog_Expr.matchsub", eval_matchsub "#matchsub_")):calc*) walther@59618: fun eval_matchsub (_:string) "Prog_Expr.matchsub" walther@59603: (t as Const ("Prog_Expr.matchsub",_) $ pat $ tst) thy = walther@59603: if matchsub thy tst pat walther@59603: then walther@59603: let val prop = HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True})) walther@59603: in SOME (Rule.term_to_string''' thy prop, prop) end walther@59603: else walther@59603: let val prop = HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False})) walther@59603: in SOME (Rule.term_to_string''' thy prop, prop) end walther@59603: | eval_matchsub _ _ _ _ = NONE; walther@59603: walther@59603: (*get the variables in an isabelle-term*) walther@59603: (*("Vars" ,("Prog_Expr.Vars" , eval_var "#Vars_")):calc*) walther@59618: fun eval_var (thmid:string) "Prog_Expr.Vars" (t as (Const _ $ arg)) thy = walther@59603: let walther@59603: val t' = ((TermC.list2isalist HOLogic.realT) o TermC.vars) t; walther@59603: val thmId = thmid ^ Rule.term_to_string''' thy arg; walther@59603: in SOME (thmId, HOLogic.Trueprop $ (TermC.mk_equality (t, t'))) end walther@59603: | eval_var _ _ _ _ = NONE; walther@59603: walther@59603: fun lhs (Const ("HOL.eq",_) $ l $ _) = l walther@59603: | lhs t = error("lhs called with (" ^ Rule.term2str t ^ ")"); walther@59603: (*("lhs" ,("Prog_Expr.lhs" , eval_lhs "")):calc*) walther@59603: fun eval_lhs _ "Prog_Expr.lhs" walther@59603: (t as (Const ("Prog_Expr.lhs",_) $ (Const ("HOL.eq",_) $ l $ _))) _ = walther@59603: SOME ((Rule.term2str t) ^ " = " ^ (Rule.term2str l), walther@59603: HOLogic.Trueprop $ (TermC.mk_equality (t, l))) walther@59603: | eval_lhs _ _ _ _ = NONE; walther@59603: (* walther@59603: > val t = (Thm.term_of o the o (parse thy)) "lhs (1 * x ^^^ 2 = 0)"; walther@59603: > val SOME (id,t') = eval_lhs 0 0 t 0; walther@59603: val id = "Prog_Expr.lhs (1 * x ^^^ 2 = 0) = 1 * x ^^^ 2" : string walther@59603: > term2str t'; walther@59603: val it = "Prog_Expr.lhs (1 * x ^^^ 2 = 0) = 1 * x ^^^ 2" : string walther@59603: *) walther@59603: walther@59603: fun rhs (Const ("HOL.eq",_) $ _ $ r) = r walther@59603: | rhs t = error("rhs called with (" ^ Rule.term2str t ^ ")"); walther@59603: (*("rhs" ,("Prog_Expr.rhs" , eval_rhs "")):calc*) walther@59603: fun eval_rhs _ "Prog_Expr.rhs" walther@59603: (t as (Const ("Prog_Expr.rhs",_) $ (Const ("HOL.eq",_) $ _ $ r))) _ = walther@59603: SOME (Rule.term2str t ^ " = " ^ Rule.term2str r, walther@59603: HOLogic.Trueprop $ (TermC.mk_equality (t, r))) walther@59603: | eval_rhs _ _ _ _ = NONE; walther@59603: (*\\------------------------- from Prog_Expr.thy-------------------------------------------------//*) walther@59603: walther@59603: (*//------------------------- from Atools.thy------------------------------------------------\\*) walther@59603: fun occurs_in v t = member op = (((map strip_thy) o TermC.ids2str) t) (Term.term_name v); walther@59603: walther@59603: (*("occurs_in", ("Prog_Expr.occurs'_in", Prog_Expr.eval_occurs_in ""))*) walther@59603: fun eval_occurs_in _ "Prog_Expr.occurs'_in" walther@59603: (p as (Const ("Prog_Expr.occurs'_in",_) $ v $ t)) _ = walther@59603: ((*tracing("@@@ eval_occurs_in: v= "^(Rule.term2str v)); walther@59603: tracing("@@@ eval_occurs_in: t= "^(Rule.term2str t));*) walther@59603: if occurs_in v t walther@59603: then SOME ((Rule.term2str p) ^ " = True", walther@59603: HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True}))) walther@59603: else SOME ((Rule.term2str p) ^ " = False", walther@59603: HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False})))) walther@59603: | eval_occurs_in _ _ _ _ = NONE; walther@59603: walther@59603: (*some of the (bound) variables (eg. in an eqsys) "vs" occur in term "t"*) walther@59603: fun some_occur_in vs t = walther@59603: let fun occurs_in' a b = occurs_in b a walther@59603: in foldl or_ (false, map (occurs_in' t) vs) end; walther@59603: walther@59603: (*("some_occur_in", ("Prog_Expr.some'_occur'_in", walther@59603: eval_some_occur_in "#eval_some_occur_in_"))*) walther@59603: fun eval_some_occur_in _ "Prog_Expr.some'_occur'_in" walther@59603: (p as (Const ("Prog_Expr.some'_occur'_in",_) walther@59603: $ vs $ t)) _ = walther@59603: if some_occur_in (TermC.isalist2list vs) t walther@59603: then SOME ((Rule.term2str p) ^ " = True", walther@59603: HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True}))) walther@59603: else SOME ((Rule.term2str p) ^ " = False", walther@59603: HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False}))) walther@59603: | eval_some_occur_in _ _ _ _ = NONE; walther@59603: walther@59603: (*("is_atom",("Prog_Expr.is'_atom", Prog_Expr.eval_is_atom "#is_atom_"))*) walther@59603: fun eval_is_atom (thmid:string) "Prog_Expr.is'_atom" walther@59603: (t as (Const _ $ arg)) _ = walther@59603: (case arg of walther@59603: Free (n,_) => SOME (TermC.mk_thmid thmid n "", walther@59603: HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True}))) walther@59603: | _ => SOME (TermC.mk_thmid thmid "" "", walther@59603: HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False})))) walther@59603: | eval_is_atom _ _ _ _ = NONE; walther@59603: walther@59603: fun even i = (i div 2) * 2 = i; walther@59603: (*("is_even",("Prog_Expr.is'_even", eval_is_even "#is_even_"))*) walther@59603: fun eval_is_even (thmid:string) "Prog_Expr.is'_even" walther@59603: (t as (Const _ $ arg)) _ = walther@59603: (case arg of walther@59603: Free (n,_) => walther@59603: (case TermC.int_of_str_opt n of walther@59603: SOME i => walther@59603: if even i then SOME (TermC.mk_thmid thmid n "", walther@59603: HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True}))) walther@59603: else SOME (TermC.mk_thmid thmid "" "", walther@59603: HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False}))) walther@59603: | _ => NONE) walther@59603: | _ => NONE) walther@59603: | eval_is_even _ _ _ _ = NONE; walther@59603: walther@59603: (*evaluate 'is_const'*) walther@59603: (*("is_const",("Prog_Expr.is'_const", Prog_Expr.eval_const "#is_const_"))*) walther@59618: fun eval_const (thmid:string) _ (t as (Const _ $ arg)) _ = walther@59603: (case arg of walther@59618: Const (n1, _) => walther@59618: SOME (TermC.mk_thmid thmid n1 "", HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False}))) walther@59618: | Free (n1, _) => walther@59618: if TermC.is_num' n1 walther@59618: then SOME (TermC.mk_thmid thmid n1 "", HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True}))) walther@59618: else SOME (TermC.mk_thmid thmid n1 "", HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False}))) walther@59603: | _ => (*NONE*) walther@59618: SOME (TermC.mk_thmid thmid (Rule.term2str arg) "", HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False})))) walther@59603: | eval_const _ _ _ _ = NONE; walther@59603: walther@59603: (*. evaluate binary, associative, commutative operators: *,+,^ .*) walther@59603: (*("PLUS" ,("Groups.plus_class.plus" , (**)eval_binop "#add_")), walther@59603: ("TIMES" ,("Groups.times_class.times" , (**)eval_binop "#mult_")), walther@59603: ("POWER" ,("Prog_Expr.pow" , (**)eval_binop "#power_"))*) walther@59603: walther@59603: (* val (thmid,op_,t as(Const (op0,t0) $ Free (n1,t1) $ Free(n2,t2)),thy) = walther@59603: ("xxxxxx",op_,t,thy); walther@59603: *) walther@59603: fun mk_thmid_f thmid ((v11, v12), (p11, p12)) ((v21, v22), (p21, p22)) = (* dropped *) walther@59603: thmid ^ "Float ((" ^ walther@59603: (string_of_int v11)^","^(string_of_int v12)^"), ("^ walther@59603: (string_of_int p11)^","^(string_of_int p12)^")) __ (("^ walther@59603: (string_of_int v21)^","^(string_of_int v22)^"), ("^ walther@59603: (string_of_int p21)^","^(string_of_int p22)^"))"; walther@59603: walther@59603: (*.evaluate < and <= for numerals.*) walther@59603: (*("le" ,("Orderings.ord_class.less" , Prog_Expr.eval_equ "#less_")), walther@59603: ("leq" ,("Orderings.ord_class.less_eq" , Prog_Expr.eval_equ "#less_equal_"))*) walther@59603: fun eval_equ (thmid:string) (_(*op_*)) (t as walther@59603: (Const (op0, _) $ Free (n1, _) $ Free(n2, _))) _ = walther@59603: (case (TermC.int_of_str_opt n1, TermC.int_of_str_opt n2) of walther@59603: (SOME n1', SOME n2') => walther@59603: if Calc.calc_equ (strip_thy op0) (n1', n2') walther@59603: then SOME (TermC.mk_thmid thmid n1 n2, walther@59603: HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True}))) walther@59603: else SOME (TermC.mk_thmid thmid n1 n2, walther@59603: HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False}))) walther@59603: | _ => NONE) walther@59603: walther@59603: | eval_equ _ _ _ _ = NONE; walther@59603: walther@59603: walther@59603: (*evaluate identity walther@59603: > reflI; walther@59603: val it = "(?t = ?t) = True" walther@59603: > val t = str2term "x = 0"; walther@59603: > val NONE = rewrite_ thy Rule.dummy_ord Rule.e_rls false reflI t; walther@59603: walther@59603: > val t = str2term "1 = 0"; walther@59603: > val NONE = rewrite_ thy Rule.dummy_ord Rule.e_rls false reflI t; walther@59603: ----------- thus needs Rule.Calc ! walther@59603: > val t = str2term "0 = 0"; walther@59603: > val SOME (t',_) = rewrite_ thy Rule.dummy_ord Rule.e_rls false reflI t; walther@59603: > Rule.term2str t'; walther@59603: val it = "HOL.True" walther@59603: walther@59603: val t = str2term "Not (x = 0)"; walther@59603: atomt t; Rule.term2str t; walther@59603: *** ------------- walther@59603: *** Const ( Not) walther@59603: *** . Const ( op =) walther@59603: *** . . Free ( x, ) walther@59603: *** . . Free ( 0, ) walther@59603: val it = "x ~= 0" : string*) walther@59603: walther@59603: (*.evaluate identity on the term-level, =!= ,i.e. without evaluation of walther@59603: the arguments: thus special handling by 'fun eval_binop'*) walther@59603: (*("ident" ,("Prog_Expr.ident", Prog_Expr.eval_ident "#ident_")):calc*) walther@59603: fun eval_ident (thmid:string) "Prog_Expr.ident" (t as walther@59603: (Const _(*op0, t0*) $ t1 $ t2 )) thy = walther@59603: if t1 = t2 walther@59603: then SOME (TermC.mk_thmid thmid walther@59603: ("(" ^ (Rule.term_to_string''' thy t1) ^ ")") walther@59603: ("(" ^ (Rule.term_to_string''' thy t2) ^ ")"), walther@59603: HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True}))) walther@59603: else SOME (TermC.mk_thmid thmid walther@59603: ("(" ^ (Rule.term_to_string''' thy t1) ^ ")") walther@59603: ("(" ^ (Rule.term_to_string''' thy t2) ^ ")"), walther@59603: HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False}))) walther@59603: | eval_ident _ _ _ _ = NONE; walther@59603: (* TODO walther@59603: > val t = str2term "x =!= 0"; walther@59603: > val SOME (str, t') = eval_ident "ident_" "b" t thy; walther@59603: > Rule.term2str t'; walther@59603: val str = "ident_(x)_(0)" : string walther@59603: val it = "(x =!= 0) = False" : string walther@59603: > val t = str2term "1 =!= 0"; walther@59603: > val SOME (str, t') = eval_ident "ident_" "b" t thy; walther@59603: > Rule.term2str t'; walther@59603: val str = "ident_(1)_(0)" : string walther@59603: val it = "(1 =!= 0) = False" : string walther@59603: > val t = str2term "0 =!= 0"; walther@59603: > val SOME (str, t') = eval_ident "ident_" "b" t thy; walther@59603: > Rule.term2str t'; walther@59603: val str = "ident_(0)_(0)" : string walther@59603: val it = "(0 =!= 0) = True" : string walther@59603: *) walther@59603: walther@59603: (*.evaluate identity of terms, which stay ready for evaluation in turn; walther@59603: thus returns False only for atoms.*) walther@59603: (*("equal" ,("HOL.eq", Prog_Expr.eval_equal "#equal_")):calc*) walther@59603: fun eval_equal (thmid : string) "HOL.eq" (t as (Const _(*op0,t0*) $ t1 $ t2 )) thy = walther@59603: if t1 = t2 walther@59603: then SOME (TermC.mk_thmid thmid walther@59603: ("(" ^ Rule.term_to_string''' thy t1 ^ ")") walther@59603: ("(" ^ Rule.term_to_string''' thy t2 ^ ")"), walther@59603: HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True}))) walther@59603: else (case (TermC.is_atom t1, TermC.is_atom t2) of walther@59603: (true, true) => walther@59603: SOME (TermC.mk_thmid thmid walther@59603: ("(" ^ Rule.term_to_string''' thy t1 ^ ")") walther@59603: ("(" ^ Rule.term_to_string''' thy t2 ^ ")"), walther@59603: HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False}))) walther@59603: | _ => NONE) (* NOT is_atom t1,t2 --> rew_sub *) walther@59603: | eval_equal _ _ _ _ = NONE; (* error-exit *) walther@59603: (* walther@59603: val t = str2term "x ~= 0"; walther@59603: val NONE = eval_equal "equal_" "b" t thy; walther@59603: walther@59603: walther@59603: > val t = str2term "(x + 1) = (x + 1)"; walther@59603: > val SOME (str, t') = eval_equal "equal_" "b" t thy; walther@59603: > Rule.term2str t'; walther@59603: val str = "equal_(x + 1)_(x + 1)" : string walther@59603: val it = "(x + 1 = x + 1) = True" : string walther@59603: > val t = str2term "x = 0"; walther@59603: > val NONE = eval_equal "equal_" "b" t thy; walther@59603: walther@59603: > val t = str2term "1 = 0"; walther@59603: > val SOME (str, t') = eval_equal "equal_" "b" t thy; walther@59603: > Rule.term2str t'; walther@59603: val str = "equal_(1)_(0)" : string walther@59603: val it = "(1 = 0) = False" : string walther@59603: > val t = str2term "0 = 0"; walther@59603: > val SOME (str, t') = eval_equal "equal_" "b" t thy; walther@59603: > Rule.term2str t'; walther@59603: val str = "equal_(0)_(0)" : string walther@59603: val it = "(0 = 0) = True" : string walther@59603: *) walther@59603: walther@59603: (*. evaluate HOL.divide .*) walther@59603: (*("DIVIDE" ,("Rings.divide_class.divide" , eval_cancel "#divide_e"))*) walther@59603: fun eval_cancel (thmid: string) "Rings.divide_class.divide" (t as walther@59603: (Const (op0,t0) $ Free (n1, _) $ Free(n2, _))) _ = walther@59603: (case (TermC.int_of_str_opt n1, TermC.int_of_str_opt n2) of walther@59603: (SOME n1', SOME n2') => walther@59603: let walther@59603: val sg = Calc.sign_mult n1' n2'; walther@59603: val (T1,T2,Trange) = TermC.dest_binop_typ t0; walther@59603: val gcd' = Calc.gcd (abs n1') (abs n2'); walther@59603: in if gcd' = abs n2' walther@59603: then let val rhs = TermC.term_of_num Trange (sg * (abs n1') div gcd') walther@59603: val prop = HOLogic.Trueprop $ (TermC.mk_equality (t, rhs)) walther@59603: in SOME (TermC.mk_thmid thmid n1 n2, prop) end walther@59603: else if 0 < n2' andalso gcd' = 1 then NONE walther@59603: else let val rhs = TermC.mk_num_op_num T1 T2 (op0,t0) (sg * (abs n1') div gcd') walther@59603: ((abs n2') div gcd') walther@59603: val prop = HOLogic.Trueprop $ (TermC.mk_equality (t, rhs)) walther@59603: in SOME (TermC.mk_thmid thmid n1 n2, prop) end walther@59603: end walther@59603: | _ => ((*tracing"@@@ eval_cancel NONE";*)NONE)) walther@59603: | eval_cancel _ _ _ _ = NONE; walther@59603: walther@59603: (* get the argument from a function-definition *) walther@59603: (*("argument_in" ,("Prog_Expr.argument'_in", Prog_Expr.eval_argument_in "Prog_Expr.argument'_in"))*) walther@59603: fun eval_argument_in _ "Prog_Expr.argument'_in" walther@59603: (t as (Const ("Prog_Expr.argument'_in", _) $ (_(*f*) $ arg))) _ = walther@59603: if is_Free arg (*could be something to be simplified before*) walther@59603: then walther@59603: SOME (Rule.term2str t ^ " = " ^ Rule.term2str arg, walther@59603: HOLogic.Trueprop $ (TermC.mk_equality (t, arg))) walther@59603: else NONE walther@59603: | eval_argument_in _ _ _ _ = NONE; walther@59603: walther@59603: (* check if the function-identifier of the first argument matches walther@59603: the function-identifier of the lhs of the second argument *) walther@59603: (*("sameFunId" ,("Prog_Expr.sameFunId", eval_same_funid "Prog_Expr.sameFunId"))*) walther@59603: fun eval_sameFunId _ "Prog_Expr.sameFunId" walther@59603: (p as Const ("Prog_Expr.sameFunId",_) $ (f1 $ _) $ (Const ("HOL.eq", _) $ (f2 $ _) $ _)) _ = walther@59603: if f1 = f2 walther@59603: then SOME ((Rule.term2str p) ^ " = True", walther@59603: HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True}))) walther@59603: else SOME ((Rule.term2str p) ^ " = False", walther@59603: HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False}))) walther@59603: | eval_sameFunId _ _ _ _ = NONE; walther@59603: walther@59603: walther@59603: (*.from a list of fun-definitions "f x = ..." as 2nd argument walther@59603: filter the elements with the same fun-identfier in "f y" walther@59603: as the fst argument; walther@59603: this is, because Isabelles filter takes more than 1 sec.*) walther@59603: fun same_funid f1 (Const ("HOL.eq", _) $ (f2 $ _) $ _) = f1 = f2 walther@59603: | same_funid f1 t = error ("same_funid called with t = (" walther@59603: ^Rule.term2str f1^") ("^Rule.term2str t^")"); walther@59603: (*("filter_sameFunId" ,("Prog_Expr.filter'_sameFunId", walther@59603: eval_filter_sameFunId "Prog_Expr.filter'_sameFunId"))*) walther@59603: fun eval_filter_sameFunId _ "Prog_Expr.filter'_sameFunId" walther@59603: (p as Const ("Prog_Expr.filter'_sameFunId",_) $ walther@59603: (fid $ _) $ fs) _ = walther@59603: let val fs' = ((TermC.list2isalist HOLogic.boolT) o walther@59603: (filter (same_funid fid))) (TermC.isalist2list fs) walther@59603: in SOME (Rule.term2str (TermC.mk_equality (p, fs')), walther@59603: HOLogic.Trueprop $ (TermC.mk_equality (p, fs'))) end walther@59603: | eval_filter_sameFunId _ _ _ _ = NONE; walther@59603: walther@59603: (* make a list of terms to a sum *) walther@59603: fun list2sum [] = error ("list2sum called with []") walther@59603: | list2sum [s] = s walther@59603: | list2sum (s::ss) = walther@59603: let walther@59603: fun sum su [s'] = Const ("Groups.plus_class.plus", walther@59603: [HOLogic.realT,HOLogic.realT] ---> HOLogic.realT) $ su $ s' walther@59603: | sum su (s'::ss') = sum (Const ("Groups.plus_class.plus", walther@59603: [HOLogic.realT,HOLogic.realT] ---> HOLogic.realT) $ su $ s') ss' walther@59603: | sum _ _ = raise ERROR "list2sum: pattern in fun.def is missing" walther@59603: in sum s ss end; walther@59603: walther@59603: (* make a list of equalities to the sum of the lhs *) walther@59603: (*("boollist2sum" ,("Prog_Expr.boollist2sum" , eval_boollist2sum "")):calc*) walther@59603: fun eval_boollist2sum _ "Prog_Expr.boollist2sum" walther@59603: (p as Const ("Prog_Expr.boollist2sum", _) $ (l as Const ("List.list.Cons", _) $ _ $ _)) _ = walther@59603: let walther@59603: val isal = TermC.isalist2list l walther@59603: val lhss = map lhs isal walther@59603: val sum = list2sum lhss walther@59603: in walther@59603: SOME ((Rule.term2str p) ^ " = " ^ (Rule.term2str sum), walther@59603: HOLogic.Trueprop $ (TermC.mk_equality (p, sum))) walther@59603: end walther@59603: | eval_boollist2sum _ _ _ _ = NONE; walther@59603: walther@59603: (**) end (*struct*) walther@59603: \ text \ walther@59603: open Prog_Expr walther@59603: \ ML \ walther@59603: \ ML \ walther@59603: \ walther@59603: walther@59603: subsection \extend rule-set for evaluating pre-conditions and program-expressions\ walther@59603: ML \ walther@59603: val list_rls = Rule.append_rls "list_rls" list_rls [Rule.Calc ("Prog_Expr.rhs", Prog_Expr.eval_rhs "")]; walther@59603: \ ML \ walther@59603: \ ML \ walther@59603: \ walther@59603: walther@59603: subsection \setup for rule-sets and ML-functions\ walther@59603: setup \KEStore_Elems.add_rlss [("list_rls", (Context.theory_name @{theory}, list_rls))]\ walther@59603: setup \KEStore_Elems.add_calcs walther@59603: [("matches", ("Prog_Expr.matches", Prog_Expr.eval_matches "#matches_")), walther@59603: ("matchsub", ("Prog_Expr.matchsub", Prog_Expr.eval_matchsub "#matchsub_")), walther@59603: ("Vars", ("Prog_Expr.Vars", Prog_Expr.eval_var "#Vars_")), walther@59603: ("lhs", ("Prog_Expr.lhs", Prog_Expr.eval_lhs "")), walther@59603: ("rhs", ("Prog_Expr.rhs", Prog_Expr.eval_rhs ""))]\ walther@59603: (*\\------------------------- from Tools .thy-------------------------------------------------//*) walther@59603: ML \ walther@59603: \ ML \ walther@59603: \ ML \ walther@59603: \ walther@59603: walther@59603: end