diff -r 05e5d68c9627 -r f1be8028a4a9 Nominal/activities/cas09/Lec2.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Nominal/activities/cas09/Lec2.thy Wed Mar 30 17:27:34 2016 +0100 @@ -0,0 +1,660 @@ +(***************************************************************** + + Isabelle Tutorial + ----------------- + + 1st June 2009, Beijing + +*) + +theory Lec2 + imports "Main" +begin + +text {*********************************************************** + + Automatic proofs again: + +*} + +inductive + even :: "nat \ bool" +where + eZ[intro]: "even 0" +| eSS[intro]: "even n \ even (Suc (Suc n))" + +text {* + In the following two lemmas we have to indicate the induction, + but the routine calculations can be done completely automatically + by Isabelle's internal tools. +*} + +lemma even_twice: + shows "even (n + n)" +by (induct n) (auto) + +lemma even_add: + assumes a: "even n" + and b: "even m" + shows "even (n + m)" +using a b by (induct) (auto) + +text {* + In the next proof it was crucial that we introduced + the equation + + (Suc (Suc n) * m) = (m + m) + (n * m) + + because this allowed us to use the induction hypothesis + and the previous two lemmas. *} + +lemma even_mul: + assumes a: "even n" + shows "even (n * m)" +using a +proof (induct) + case eZ + show "even (0 * m)" by auto +next + case (eSS n) + have as: "even n" by fact + have ih: "even (n * m)" by fact + have "(Suc (Suc n) * m) = (m + m) + (n * m)" by simp + moreover + have "even (m + m)" using even_twice by simp + ultimately + show "even (Suc (Suc n) * m)" using ih even_add by (simp only:) +qed + +text {* + This proof cannot be found by the internal tools, but + Isabelle has an interface to external provers. This + interface is called sledgehammer and it finds the + proof more or less automatically. + +*} + + +lemma even_mul_auto: + assumes a: "even n" + shows "even (n * m)" +using a +apply(induct) +apply(metis eZ mult_is_0) +apply(metis even_add even_twice mult_Suc_right nat_add_assoc nat_mult_commute) +done + +text {* + The disadvantage of such proofs is that you do + not know why they are true. +*} + + +text {***************************************************************** + + Function Definitions + -------------------- + + Many functions over datatypes can be defined by recursion on the + structure. For this purpose, Isabelle provides "fun". To use it one needs + to give a name for the function, its type, optionally some pretty-syntax + and then some equations defining the function. Like in "inductive", + "fun" expects that more than one such equation is separated by "|". + + The Fibonacci function: + +*} + +fun + fib :: "nat \ nat" +where + "fib 0 = 1" +| "fib (Suc 0) = 1" +| "fib (Suc (Suc n)) = fib n + fib (Suc n)" + +text {* The Ackermann function: *} + +fun + ack :: "nat \ nat \ nat" +where + "ack 0 m = Suc m" +| "ack (Suc n) 0 = ack n (Suc 0)" +| "ack (Suc n) (Suc m) = ack n (ack (Suc n) m)" + +text {* + Non-terminating functions can lead to inconsistencies. +*} + +lemma + assumes a: "f x = f x + (1::nat)" + shows "False" +proof - + from a have "0 = (1::nat)" by simp + then show "False" by simp +qed + +text {* + If Isabelle cannot automatically prove termination, + then you can give an explicit measure that establishes + termination. For example a generalisation of the + Fibonacci function to integers can be shown terminating + +*} + +function + fib' :: "int \ int" +where + "n< -1 \ fib' n = fib' (n + 2) - fib' (n + 1)" +| "fib' -1 = (1::int)" +| "fib' 0 = (0::int)" +| "fib' 1 = (1::int)" +| "n > 1 \ fib' n = fib' (n - 1) + fib' (n - 2)" + by (atomize_elim, presburger) (auto) + +termination + by (relation "measure (\x. nat (\x\))") + (simp_all add: zabs_def) + + +text {***************************************************************** + + Datatypes + --------- + + Datatypes can be defined in Isabelle as follows: we have to provide the name + of the datatype and list its type-constructors. Each type-constructor needs + to have the information about the types of its arguments, and optionally + can also contain some information about pretty syntax. For example, we write + "[]" for Nil and ":::" for Cons. + +*} + +datatype 'a mylist = + MyNil ("[]") +| MyCons "'a" "'a mylist" ("_ ::: _" 65) + +text {* + You can easily define functions over the structure of datatypes. +*} + +fun + myappend :: "'a mylist \ 'a mylist \ 'a mylist" ("_ @@ _" 65) +where + "[] @@ xs = xs" +| "(y:::ys) @@ xs = y:::(ys @@ xs)" + +fun + myrev :: "'a mylist \ 'a mylist" +where + "myrev [] = []" +| "myrev (x:::xs) = (myrev xs) @@ (x:::[])" + +text {***************************************************************** + + Exercise: Prove the following fact about append and reversal. + +*} + +lemma myrev_append: + shows "myrev (xs @@ ys) = (myrev ys) @@ (myrev xs)" +proof (induct xs) + case MyNil + show "myrev ([] @@ ys) = myrev ys @@ myrev []" sorry +next + case (MyCons x xs) + have ih: "myrev (xs @@ ys) = myrev ys @@ myrev xs" by fact + + show "myrev ((x:::xs) @@ ys) = myrev ys @@ myrev (x:::xs)" sorry +qed + + + + + + + + + + + + + + + + +text {* auxiliary lemmas *} + +lemma mynil_append[simp]: + shows "xs @@ [] = xs" +by (induct xs) (auto) + +lemma myappend_assoc[simp]: + "(xs @@ ys) @@ zs = xs @@ (ys @@ zs)" +by (induct xs) (auto) + +lemma + shows "myrev (xs @@ ys) = (myrev ys) @@ (myrev xs)" +by (induct xs) (auto) + +text {***************************************************************** + + A WHILE Language: + + Arithmetic Expressions, Boolean Expressions, Commands + ----------------------------------------------------- + +*} + +types memory = "nat \ nat" + +datatype aexp = + C nat + | X nat + | Op1 "nat \ nat" aexp + | Op2 "nat \ nat \ nat" aexp aexp + +datatype bexp = + TRUE + | FALSE + | ROp "nat \ nat \ bool" aexp aexp + | NOT bexp + | AND bexp bexp + | OR bexp bexp + +datatype cmd = + SKIP + | ASSIGN nat aexp ("_ ::= _ " 60) + | SEQ cmd cmd ("_; _" [60, 60] 10) + | COND bexp cmd cmd ("IF _ THEN _ ELSE _" 60) + | WHILE bexp cmd ("WHILE _ DO _" 60) + +text {***************************************************************** + + Big-step Evaluation Semantics + ----------------------------- + +*} + +inductive + evala :: "aexp \ memory \ nat \ bool" ("'(_,_') \a _" [100,100] 50) +where + C: "(C n, m) \a n" +| X: "(X i, m) \a m i" +| Op1: "(e, m) \a n \ (Op1 f e, m) \a (f n)" +| Op2: "\(e0, m) \a n0; (e1, m) \a n1\ \ (Op2 f e0 e1, m) \a f n0 n1" + +inductive + evalb :: "bexp \ memory \ bool \ bool" ("'(_,_') \b _" [100,100] 50) +where + TRUE: "(TRUE, m) \b True" +| FALSE: "(FALSE, m) \b False" +| ROp: "\(e1, m) \a n1; (e2, m) \a n2\ \ (ROp f e1 e2, m) \b (f n1 n2)" +| NOT: "(e, m) \b b \ (NOT e, m) \b (Not b)" +| AND: "\(e1, m) \b b1; (e2, m) \b b2\ \ (AND e1 e2, m) \b (b1 \ b2)" +| OR: "\(e1, m) \b b1; (e2, m) \b b2\ \ (OR e1 e2, m) \b (b1 \ b2)" + +inductive + evalc :: "cmd \ memory \ memory \ bool" ("'(_,_') \c _" [0,0,60] 60) +where + SKIP: "(SKIP, m) \c m" +| ASSIGN: "(a, m) \a n \ (x::=a, m) \c m(x:=n)" +| SEQ: "\(c0, m) \c m'; (c1, m') \c m''\ \ (c0; c1, m) \c m''" +| IFTrue: "\(e, m) \b True; (c0, m) \c m'\ \ (IF e THEN c0 ELSE c1, m) \c m'" +| IFFalse: "\(e, m) \b False; (c1, m) \c m'\ \ (IF e THEN c0 ELSE c1, m) \c m'" +| WHILEFalse: "(e, m) \b False \ (WHILE e DO c,m) \c m" +| WHILETrue: "\(e, m) \b True; (c,m) \c m'; (WHILE e DO c, m') \c m''\ + \ (WHILE e DO c, m) \c m''" + +lemmas evala.intros[intro] evalb.intros[intro] evalc.intros[intro] + +text {***************************************************************** + + Machine Instructions + -------------------- + +*} + +datatype instr = + JPFZ "nat" (* JUMP forward n steps, if top of stack is False *) + | JPB "nat" (* JUMP backward n steps *) + | FETCH "nat" (* move memory to top of stack *) + | STORE "nat" (* move top of stack to memory *) + | PUSH "nat" (* push to stack *) + | OPU "nat \ nat" (* pop one from stack and apply f *) + | OPB "nat \ nat \ nat" (* pop two from stack and apply f *) + +text {***************************************************************** + + Booleans 'as' Zero and One + -------------------------- + +*} + +fun + WRAP::"bool\nat" +where + "WRAP True = Suc 0" +| "WRAP False = 0" + +fun + MNot::"nat \ nat" +where + "MNot 0 = 1" +| "MNot (Suc 0) = 0" + +fun + MAnd::"nat \ nat \ nat" +where + "MAnd 0 x = 0" +| "MAnd x 0 = 0" +| "MAnd (Suc 0) (Suc 0) = 1" + +fun + MOr::"nat \ nat \ nat" +where + "MOr 0 0 = 0" +| "MOr (Suc 0) x = 1" +| "MOr x (Suc 0) = 1" + +lemma MNot_WRAP: + "MNot (WRAP b) = WRAP (Not b)" +by (cases b) (auto) + +lemma MAnd_WRAP: + "MAnd (WRAP b1) (WRAP b2) = WRAP (b1 \ b2)" +by (cases b1, auto) (cases b2, simp_all) + + +lemma MOr_WRAP: + "MOr (WRAP b1) (WRAP b2) = WRAP (b1 \ b2)" +by (cases b1, auto) (cases b2, simp_all) + +lemmas WRAP_lems = MNot_WRAP MAnd_WRAP MOr_WRAP + +text {***************************************************************** + + Compiler Functions + ------------------ + +*} + +fun + compa +where + "compa (C n) = [PUSH n]" +| "compa (X l) = [FETCH l]" +| "compa (Op1 f e) = (compa e) @ [OPU f]" +| "compa (Op2 f e1 e2) = (compa e1) @ (compa e2) @ [OPB f]" + +fun + compb +where + "compb (TRUE) = [PUSH 1]" +| "compb (FALSE) = [PUSH 0]" +| "compb (ROp f e1 e2) = (compa e1) @ (compa e2) @ [OPB (\x y. WRAP (f x y))]" +| "compb (NOT e) = (compb e) @ [OPU MNot]" +| "compb (AND e1 e2) = (compb e1) @ (compb e2) @ [OPB MAnd]" +| "compb (OR e1 e2) = (compb e1) @ (compb e2) @ [OPB MOr]" + +fun + compc :: "cmd \ instr list" +where + "compc SKIP = []" +| "compc (x::=a) = (compa a) @ [STORE x]" +| "compc (c1;c2) = compc c1 @ compc c2" +| "compc (IF b THEN c1 ELSE c2) = + (compb b) @ [JPFZ (length(compc c1) + 2)] @ compc c1 @ + [PUSH 0, JPFZ (length(compc c2))] @ compc c2" +| "compc (WHILE b DO c) = + (compb b) @ + [JPFZ (length(compc c) + 1)] @ compc c @ + [JPB (length(compc c) + length(compb b)+1)]" + +value "compc SKIP" +value "compc (SKIP; SKIP)" +value "compa (Op2 (op +) (C 2) (C 3))" +value "compb (ROp (op <) (C 2) (C 3))" +value "compc (IF (ROp (op <) (C 2) (C 3)) THEN SKIP ELSE SKIP)" +value "compc (WHILE (ROp (op <) (C 2) (C 3)) DO SKIP)" +value "compc (x::= C 0; WHILE (ROp (op <) (X x) (C 3)) DO (x::= (Op2 (op +) (X x) (C 1))))" + +text {***************************************************************** + + Machine Execution + ----------------- + +*} + +types instrs = "instr list" +types stack = "nat list" + +abbreviation + "rtake i xs \ rev (take i xs)" + +inductive + step :: "instrs \ instrs \ stack \ memory \ + instrs \ instrs \ stack \ memory \ bool" ("'(_,_,_,_') \m '(_,_,_,_')") +where + "(PUSH n#p, q, s, m) \m (p, PUSH n#q, n#s, m)" + | "(FETCH i#p, q, s, m) \m (p, FETCH i#q, m i#s, m)" + | "(OPU f#p, q, n#s, m) \m (p, OPU f#q, f n#s, m)" + | "(OPB f#p, q, n1#n2#s, m) \m (p, OPB f#q, f n2 n1#s, m)" + | "(STORE i#p, q, n#s, m) \m (p, STORE i#q, s, m(i:=n))" + | "(JPFZ i#p, q, Suc 0#s, m) \m (p, JPFZ i#q, s, m)" + | "i\length p\(JPFZ i#p, q, 0#s, m) \m (drop i p, (rtake i p) @ (JPFZ i#q), s, m)" + | "i\length q\(JPB i#p, q, s, m) \m ((rtake i q) @ (JPB i#p), drop i q, s, m)" + +lemmas step.intros[intro] + +inductive_cases step_elim[elim]: + "(p,q,s,m) \m (p',q',s',m')" + +lemma exec_simp: + shows "(instr#p, q, s, m) \m (p', q', s', m') = + (case instr of + PUSH n \ (p' = p \ q' = instr#q \ s' = n#s \ m' = m) + | FETCH i \ (p' = p \ q' = instr#q \ s' = m i#s \ m' = m) + | OPU f \ (\n si. p' = p \ q' = instr#q \ s = n#si \ s' = f n#si \ m' = m) + | OPB f \ (\n1 n2 si. p' = p \ q' = instr#q \ s = n1#n2#si \ s' = f n2 n1#si \ m' = m) + | STORE i \ (\n si. p' = p \ q' = instr#q \ s = n#si \ s' = si \ m' = m(i:=n)) + | JPFZ i \ (\si. (p' = p \ q' = instr#q \ s = Suc 0#si \ s' = si \ m' = m) \ + (i\ length p \ p' = (drop i p) \ q' = (rev (take i p))@(instr#q) + \ s = 0#si \ s' = si \ m' = m )) + | JPB i \ (i \ length q \ p' = (rtake i q)@(JPB i#p) \ q' = drop i q \ s' = s \ m' = m))" +by (auto split: instr.split_asm split_if_asm) + +inductive + steps :: "instrs \ instrs \ stack \ memory \ + instrs \ instrs \ stack \ memory \ bool" ("'(_,_,_,_') \m* '(_,_,_,_')") +where + refl: "(p,q,s,m) \m* (p,q,s,m)" +| step: "\(p1,q1,s1,m1) \m (p2,q2,s2,m2); (p2,q2,s2,m2) \m* (p3,q3,s3,m3)\ + \ (p1,q1,s1,m1) \m* (p3,q3,s3,m3)" + +lemmas steps.step[intro] +lemmas steps.refl[simp] + +inductive_cases steps_elim[elim]: + "(p,q,s,m) \m* (p',q',s',m')" + +lemma steps_trans: + assumes a: "(p1,q1,s1,m1) \m* (p2,q2,s2,m2)" + and b: "(p2,q2,s2,m2) \m* (p3,q3,s3,m3)" + shows "(p1,q1,s1,m1) \m* (p3,q3,s3,m3)" +using a b by (induct) (auto) + +lemma steps_simp: + shows "(i#p,q,s,m) \m* (p',q',s',m') = + ((i#p,q,s,m) = (p',q',s',m') \ + (\pi qi si mi. (i#p,q,s,m) \m (pi,qi,si,mi) \ (pi,qi,si,mi) \m* (p',q',s',m')))" +by (auto) + +lemma compa_first_attempt: + assumes a: "(e, m) \a n" + shows "(compa e, [], [], m) \m* ([], rev (compa e), [n], m)" +using a +proof (induct) + case (C n m) + show "(compa (C n),[],[],m) \m* ([],rev (compa (C n)),[n],m)" + sorry +next + case (X i m) + show "(compa (X i),[],[],m) \m* ([],rev (compa (X i)),[m i],m)" + sorry +next + case (Op1 e m n f) + have as: "(e, m) \a n" by fact + have ih: "(compa e,[],[],m) \m* ([],rev (compa e),[n],m)" by fact + show "(compa (Op1 f e),[],[],m) \m* ([],rev (compa (Op1 f e)),[f n],m)" + using ih sorry +next + case (Op2 e0 m n0 e1 n1 f) + have as1: "(e0, m) \a n0" by fact + have as2: "(e1, m) \a n1" by fact + have ih1: "(compa e0,[],[],m) \m* ([],rev (compa e0),[n0],m)" by fact + have ih2: "(compa e1,[],[],m) \m* ([],rev (compa e1),[n1],m)" by fact + show "(compa (Op2 f e0 e1),[],[],m) \m* ([],rev (compa (Op2 f e0 e1)),[f n0 n1],m)" + using ih1[THEN steps_trans] ih2[THEN steps_trans] + sorry +qed + + + + + + + + + + + + + + + +lemma compa_aux: + assumes a: "(e, m) \a n" + shows "(compa e@p, q, s, m) \m* (p, rev (compa e)@q, n#s, m)" +using a +proof (induct arbitrary: p q s) + case (C n m p q s) + show "(compa (C n)@p,q,s,m) \m* (p,rev (compa (C n))@q,n#s,m)" + by (simp add: steps_simp exec_simp) +next + case (X i m p q s) + show "(compa (X i)@p,q,s,m) \m* (p,rev (compa (X i))@q,m i#s,m)" + by (simp add: steps_simp exec_simp) +next + case (Op1 e m n f p q s) + have as: "(e, m) \a n" by fact + have ih: "\p q s. (compa e@p,q,s,m) \m* (p,rev (compa e)@q,n#s,m)" by fact + show "(compa (Op1 f e)@p,q,s,m) \m* (p,rev (compa (Op1 f e))@q,f n#s,m)" + using ih[THEN steps_trans] by (simp add: steps_simp exec_simp) +next + case (Op2 e0 m n0 e1 n1 f p q s) + have as1: "(e0, m) \a n0" by fact + have as2: "(e1, m) \a n1" by fact + have ih1: "\p q s. (compa e0@p,q,s,m) \m* (p,rev (compa e0)@q,n0#s,m)" by fact + have ih2: "\p q s. (compa e1@p,q,s,m) \m* (p,rev (compa e1)@q,n1#s,m)" by fact + show "(compa (Op2 f e0 e1)@p,q,s,m) \m* (p,rev (compa (Op2 f e0 e1))@q,f n0 n1#s,m)" + using ih1[THEN steps_trans] ih2[THEN steps_trans] + by (simp add: steps_simp exec_simp) +qed + + +text {* + After the fact I constructed the detailed proof, I + could guide Isabelle to find the proof automatically. + + *B U T T H A T I S C H E A T I N G!!!* +*} + +lemma compa_aux_cheating: + assumes a: "(e,m) \a n" + shows "(compa e@p,q,s,m) \m* (p,rev (compa e)@q,n#s,m)" +using a +by (induct arbitrary: p q s) + (force intro: steps_trans simp add: steps_simp exec_simp)+ + + +lemma compb_aux_cheating: + assumes a: "(e,m) \b b" + shows "(compb e@p,q,s,m) \m* (p,rev (compb e)@q,(WRAP b)#s,m)" +using a +by (induct arbitrary: p q s) + (force intro: compa_aux steps_trans + simp add: steps_simp exec_simp WRAP_lems)+ + +text {* + + Exercise: After you know the lemma is provable, + try to give all details. + +*} + +lemma compb_aux: + assumes a: "(e, m) \b b" + shows "(compb e@p, q, s, m) \m* (p, rev (compb e)@q, WRAP b#s, m)" +using a +proof (induct arbitrary: p q s) + case (TRUE m p q s) + show "(compb TRUE@p,q,s,m) \m* (p,rev (compb TRUE)@q,WRAP True#s,m)" + sorry +next + case (FALSE m p q s) + show "(compb FALSE@p,q,s,m) \m* (p,rev (compb FALSE)@q,WRAP False#s,m)" + sorry +next + case (ROp e1 m n1 e2 n2 f p q s) + have as1: "(e1, m) \a n1" by fact + have as2: "(e2, m) \a n2" by fact + show "(compb (ROp f e1 e2)@p,q,s,m) + \m* (p,rev (compb (ROp f e1 e2))@q,WRAP (f n1 n2)#s,m)" + using as1 as2 sorry +next + case (NOT e m b p q s) + have ih: "\p q s. (compb e@p,q,s,m) \m* (p,rev (compb e)@q,WRAP b#s,m)" by fact + show "(compb (NOT e)@p,q,s,m) \m* (p,rev (compb (NOT e))@q,WRAP (\b)#s,m)" + using ih[THEN steps_trans] + sorry +next + case (AND e1 m b1 e2 b2 p q s) + have ih1: "\p q s. (compb e1@p,q,s,m) \m* (p,rev (compb e1)@q,WRAP b1#s,m)" by fact + have ih2: "\p q s. (compb e2@p,q,s,m) \m* (p,rev (compb e2)@q,WRAP b2#s,m)" by fact + show "(compb (AND e1 e2)@p,q,s,m) \m* (p,rev (compb (AND e1 e2))@q,WRAP (b1 \ b2)#s,m)" + using ih1[THEN steps_trans] ih2[THEN steps_trans] + sorry +next + case (OR e1 m b1 e2 b2 p q s) + have ih1: "\p q s. (compb e1@p,q,s,m) \m* (p,rev (compb e1)@q,WRAP b1#s,m)" by fact + have ih2: "\p q s. (compb e2@p,q,s,m) \m* (p,rev (compb e2)@q,WRAP b2#s,m)" by fact + show "(compb (OR e1 e2)@p,q,s,m) \m* (p,rev (compb (OR e1 e2))@q,WRAP (b1 \ b2)#s,m)" + using ih1[THEN steps_trans] ih2[THEN steps_trans] + sorry +qed + +text {* + + Also the detailed proof of this lemma is left as an exercise. + +*} + +lemma compc_aux_cheating: + assumes a: "(c,m) \c m'" + shows "(compc c@p,q,s,m) \m* (p,rev (compc c)@q,s,m')" +using a +by (induct arbitrary: p q) + (force intro: compa_aux compb_aux_cheating steps_trans + simp add: steps_simp exec_simp)+ + +end + + + + + + + + + +