oheimb@13208: (* Title: HOL/Prolog/Func.thy oheimb@13208: Author: David von Oheimb (based on a lecture on Lambda Prolog by Nadathur) oheimb@13208: *) oheimb@9015: wenzelm@17311: header {* Untyped functional language, with call by value semantics *} oheimb@9015: wenzelm@17311: theory Func haftmann@35301: imports HOHH wenzelm@17311: begin oheimb@9015: wenzelm@17311: typedecl tm oheimb@9015: haftmann@35265: axiomatization haftmann@35265: abs :: "(tm \ tm) \ tm" and haftmann@35265: app :: "tm \ tm \ tm" and wenzelm@17311: haftmann@35265: cond :: "tm \ tm \ tm \ tm" and haftmann@35265: "fix" :: "(tm \ tm) \ tm" and oheimb@9015: haftmann@35265: true :: tm and haftmann@35265: false :: tm and haftmann@35265: "and" :: "tm \ tm \ tm" (infixr "and" 999) and haftmann@35265: eq :: "tm \ tm \ tm" (infixr "eq" 999) and haftmann@35265: haftmann@35265: Z :: tm ("Z") and haftmann@35265: S :: "tm \ tm" and oheimb@9015: haftmann@35265: plus :: "tm \ tm \ tm" (infixl "+" 65) and haftmann@35265: minus :: "tm \ tm \ tm" (infixl "-" 65) and haftmann@35265: times :: "tm \ tm \ tm" (infixl "*" 70) and oheimb@9015: haftmann@35265: eval :: "tm \ tm \ bool" where haftmann@35265: haftmann@35265: eval: " oheimb@9015: oheimb@9015: eval (abs RR) (abs RR).. oheimb@9015: eval (app F X) V :- eval F (abs R) & eval X U & eval (R U) V.. oheimb@9015: oheimb@9015: eval (cond P L1 R1) D1 :- eval P true & eval L1 D1.. oheimb@9015: eval (cond P L2 R2) D2 :- eval P false & eval R2 D2.. oheimb@9015: eval (fix G) W :- eval (G (fix G)) W.. oheimb@9015: oheimb@9015: eval true true .. oheimb@9015: eval false false.. oheimb@9015: eval (P and Q) true :- eval P true & eval Q true .. oheimb@9015: eval (P and Q) false :- eval P false | eval Q false.. wenzelm@17311: eval (A1 eq B1) true :- eval A1 C1 & eval B1 C1.. oheimb@9015: eval (A2 eq B2) false :- True.. oheimb@9015: oheimb@9015: eval Z Z.. oheimb@9015: eval (S N) (S M) :- eval N M.. oheimb@9015: eval ( Z + M) K :- eval M K.. oheimb@9015: eval ((S N) + M) (S K) :- eval (N + M) K.. oheimb@9015: eval (N - Z) K :- eval N K.. oheimb@9015: eval ((S N) - (S M)) K :- eval (N- M) K.. oheimb@9015: eval ( Z * M) Z.. oheimb@9015: eval ((S N) * M) K :- eval (N * M) L & eval (L + M) K" oheimb@9015: wenzelm@21425: lemmas prog_Func = eval wenzelm@21425: wenzelm@36319: schematic_lemma "eval ((S (S Z)) + (S Z)) ?X" wenzelm@21425: apply (prolog prog_Func) wenzelm@21425: done wenzelm@21425: wenzelm@36319: schematic_lemma "eval (app (fix (%fact. abs(%n. cond (n eq Z) (S Z) wenzelm@21425: (n * (app fact (n - (S Z))))))) (S (S (S Z)))) ?X" wenzelm@21425: apply (prolog prog_Func) wenzelm@21425: done wenzelm@17311: oheimb@9015: end