# HG changeset patch # User wenzelm # Date 966786320 -7200 # Node ID 4cae97480a6d89be61a2bee1ac01a1d336d02251 # Parent e4d58f1be05b8476b63041e1f7132cb1fa66ea5e open cases; diff -r e4d58f1be05b -r 4cae97480a6d src/HOL/MicroJava/BV/Convert.thy --- a/src/HOL/MicroJava/BV/Convert.thy Sat Aug 19 12:49:19 2000 +0200 +++ b/src/HOL/MicroJava/BV/Convert.thy Sun Aug 20 17:45:20 2000 +0200 @@ -241,7 +241,7 @@ { fix b have "?Q (l#ls) b" - proof (cases b) + proof (cases (open) b) case Nil thus ?thesis by (auto dest: sup_loc_length) next diff -r e4d58f1be05b -r 4cae97480a6d src/HOL/MicroJava/BV/LBVComplete.thy --- a/src/HOL/MicroJava/BV/LBVComplete.thy Sat Aug 19 12:49:19 2000 +0200 +++ b/src/HOL/MicroJava/BV/LBVComplete.thy Sun Aug 20 17:45:20 2000 +0200 @@ -282,7 +282,7 @@ "G \ s2 <=s s1" "i = ins!pc" show ?thesis - proof (cases "cert!pc") + proof (cases (open) "cert!pc") case None with wtl fits show ?thesis diff -r e4d58f1be05b -r 4cae97480a6d src/HOL/MicroJava/BV/LBVCorrect.thy --- a/src/HOL/MicroJava/BV/LBVCorrect.thy Sat Aug 19 12:49:19 2000 +0200 +++ b/src/HOL/MicroJava/BV/LBVCorrect.thy Sun Aug 20 17:45:20 2000 +0200 @@ -47,7 +47,7 @@ have "\ s0 pc. wtl_inst_list is G rT s0 s2 cert (pc+length is) pc \ (\ phi. pc + length is < length phi \ fits_partial phi pc is G rT s0 s2 cert)" (is "?P is") - proof (induct "?P" "is") + proof (induct (open) ?P "is") case Nil show "?P []" by (simp add: fits_partial_def exists_list) case Cons diff -r e4d58f1be05b -r 4cae97480a6d src/HOL/MicroJava/BV/LBVSpec.thy --- a/src/HOL/MicroJava/BV/LBVSpec.thy Sat Aug 19 12:49:19 2000 +0200 +++ b/src/HOL/MicroJava/BV/LBVSpec.thy Sun Aug 20 17:45:20 2000 +0200 @@ -11,82 +11,82 @@ types certificate = "state_type option list" - class_certificate = "sig \\ certificate" - prog_certificate = "cname \\ class_certificate" + class_certificate = "sig \ certificate" + prog_certificate = "cname \ class_certificate" constdefs -wtl_inst :: "[instr,jvm_prog,ty,state_type,state_type,certificate,p_count,p_count] \\ bool" +wtl_inst :: "[instr,jvm_prog,ty,state_type,state_type,certificate,p_count,p_count] \ bool" -"wtl_inst i G rT s s' cert max_pc pc \\ app (i,G,rT,s) \\ +"wtl_inst i G rT s s' cert max_pc pc \ app (i,G,rT,s) \ (let s'' = the (step (i,G,s)) in - (\\ pc' \\ (succs i pc). pc' < max_pc \\ - ((pc' \\ pc+1) \\ cert!pc' \\ None \\ G \\ s'' <=s the (cert!pc'))) \\ - (if (pc+1) \\ (succs i pc) then + (\ pc' \ (succs i pc). pc' < max_pc \ + ((pc' \ pc+1) \ cert!pc' \ None \ G \ s'' <=s the (cert!pc'))) \ + (if (pc+1) \ (succs i pc) then s' = s'' else cert ! (pc+1) = Some s'))" lemma [simp]: -"succs i pc = {pc+1} \\ - wtl_inst i G rT s s' cert max_pc pc = (app (i,G,rT,s) \\ pc+1 < max_pc \\ (s' = the (step (i,G,s))))" +"succs i pc = {pc+1} \ + wtl_inst i G rT s s' cert max_pc pc = (app (i,G,rT,s) \ pc+1 < max_pc \ (s' = the (step (i,G,s))))" by (unfold wtl_inst_def, auto) lemma [simp]: -"succs i pc = {} \\ wtl_inst i G rT s s' cert max_pc pc = (app (i,G,rT,s) \\ cert!(pc+1) = Some s')" +"succs i pc = {} \ wtl_inst i G rT s s' cert max_pc pc = (app (i,G,rT,s) \ cert!(pc+1) = Some s')" by (unfold wtl_inst_def, auto) constdefs -wtl_inst_option :: "[instr,jvm_prog,ty,state_type,state_type,certificate,p_count,p_count] \\ bool" -"wtl_inst_option i G rT s0 s1 cert max_pc pc \\ +wtl_inst_option :: "[instr,jvm_prog,ty,state_type,state_type,certificate,p_count,p_count] \ bool" +"wtl_inst_option i G rT s0 s1 cert max_pc pc \ (case cert!pc of - None \\ wtl_inst i G rT s0 s1 cert max_pc pc - | Some s0' \\ (G \\ s0 <=s s0') \\ + None \ wtl_inst i G rT s0 s1 cert max_pc pc + | Some s0' \ (G \ s0 <=s s0') \ wtl_inst i G rT s0' s1 cert max_pc pc)" consts - wtl_inst_list :: "[instr list,jvm_prog,ty,state_type,state_type,certificate,p_count,p_count] \\ bool" + wtl_inst_list :: "[instr list,jvm_prog,ty,state_type,state_type,certificate,p_count,p_count] \ bool" primrec "wtl_inst_list [] G rT s0 s2 cert max_pc pc = (s0 = s2)" "wtl_inst_list (instr#is) G rT s0 s2 cert max_pc pc = - (\\s1. wtl_inst_option instr G rT s0 s1 cert max_pc pc \\ + (\s1. wtl_inst_option instr G rT s0 s1 cert max_pc pc \ wtl_inst_list is G rT s1 s2 cert max_pc (pc+1))" constdefs - wtl_method :: "[jvm_prog,cname,ty list,ty,nat,instr list,certificate] \\ bool" - "wtl_method G C pTs rT mxl ins cert \\ + wtl_method :: "[jvm_prog,cname,ty list,ty,nat,instr list,certificate] \ bool" + "wtl_method G C pTs rT mxl ins cert \ let max_pc = length ins in - 0 < max_pc \\ - (\\s2. wtl_inst_list ins G rT + 0 < max_pc \ + (\s2. wtl_inst_list ins G rT ([],(Some(Class C))#((map Some pTs))@(replicate mxl None)) s2 cert max_pc 0)" - wtl_jvm_prog :: "[jvm_prog,prog_certificate] \\ bool" - "wtl_jvm_prog G cert \\ - wf_prog (\\G C (sig,rT,maxl,b). + wtl_jvm_prog :: "[jvm_prog,prog_certificate] \ bool" + "wtl_jvm_prog G cert \ + wf_prog (\G C (sig,rT,maxl,b). wtl_method G C (snd sig) rT maxl b (cert C sig)) G" text {* \medskip *} -lemma rev_eq: "\\length a = n; length x = n; rev a @ b # c = rev x @ y # z\\ \\ a = x \\ b = y \\ c = z" +lemma rev_eq: "\length a = n; length x = n; rev a @ b # c = rev x @ y # z\ \ a = x \ b = y \ c = z" by auto lemma wtl_inst_unique: -"wtl_inst i G rT s0 s1 cert max_pc pc \\ - wtl_inst i G rT s0 s1' cert max_pc pc \\ s1 = s1'" (is "?P i") +"wtl_inst i G rT s0 s1 cert max_pc pc \ + wtl_inst i G rT s0 s1' cert max_pc pc \ s1 = s1'" (is "?P i") by (unfold wtl_inst_def, auto) lemma wtl_inst_option_unique: -"\\wtl_inst_option i G rT s0 s1 cert max_pc pc; - wtl_inst_option i G rT s0 s1' cert max_pc pc\\ \\ s1 = s1'" +"\wtl_inst_option i G rT s0 s1 cert max_pc pc; + wtl_inst_option i G rT s0 s1' cert max_pc pc\ \ s1 = s1'" by (cases "cert!pc") (auto simp add: wtl_inst_unique wtl_inst_option_def) lemma wtl_inst_list_unique: -"\\ s0 pc. wtl_inst_list is G rT s0 s1 cert max_pc pc \\ - wtl_inst_list is G rT s0 s1' cert max_pc pc \\ s1=s1'" (is "?P is") -proof (induct "?P" "is") +"\ s0 pc. wtl_inst_list is G rT s0 s1 cert max_pc pc \ + wtl_inst_list is G rT s0 s1' cert max_pc pc \ s1=s1'" (is "?P is") +proof (induct (open) ?P "is") case Nil show "?P []" by simp @@ -110,13 +110,13 @@ qed lemma wtl_partial: -"\\ pc' pc s. - wtl_inst_list is G rT s s' cert mpc pc \\ \ - pc' < length is \\ \ - (\\ a b s1. a @ b = is \\ length a = pc' \\ \ - wtl_inst_list a G rT s s1 cert mpc pc \\ \ +"\ pc' pc s. + wtl_inst_list is G rT s s' cert mpc pc \ \ + pc' < length is \ \ + (\ a b s1. a @ b = is \ length a = pc' \ \ + wtl_inst_list a G rT s s1 cert mpc pc \ \ wtl_inst_list b G rT s1 s' cert mpc (pc+length a))" (is "?P is") -proof (induct "?P" "is") +proof (induct (open) ?P "is") case Nil show "?P []" by auto @@ -126,12 +126,12 @@ fix pc' pc s assume length: "pc' < length (a # list)" assume wtl: "wtl_inst_list (a # list) G rT s s' cert mpc pc" - show "\\ a' b s1. - a' @ b = a#list \\ length a' = pc' \\ \ - wtl_inst_list a' G rT s s1 cert mpc pc \\ \ + show "\ a' b s1. + a' @ b = a#list \ length a' = pc' \ \ + wtl_inst_list a' G rT s s1 cert mpc pc \ \ wtl_inst_list b G rT s1 s' cert mpc (pc+length a')" - (is "\\ a b s1. ?E a b s1") - proof (cases "pc'") + (is "\ a b s1. ?E a b s1") + proof (cases (open) pc') case 0 with wtl have "?E [] (a#list) s" by simp @@ -158,8 +158,8 @@ qed lemma "wtl_append1": -"\\wtl_inst_list x G rT s0 s1 cert (length (x@y)) 0; - wtl_inst_list y G rT s1 s2 cert (length (x@y)) (length x)\\ \\ +"\wtl_inst_list x G rT s0 s1 cert (length (x@y)) 0; + wtl_inst_list y G rT s1 s2 cert (length (x@y)) (length x)\ \ wtl_inst_list (x@y) G rT s0 s2 cert (length (x@y)) 0" proof - assume w: @@ -167,11 +167,11 @@ "wtl_inst_list y G rT s1 s2 cert (length (x@y)) (length x)" have - "\\ pc s0. - wtl_inst_list x G rT s0 s1 cert (pc+length (x@y)) pc \\ - wtl_inst_list y G rT s1 s2 cert (pc+length (x@y)) (pc+length x) \\ + "\ pc s0. + wtl_inst_list x G rT s0 s1 cert (pc+length (x@y)) pc \ + wtl_inst_list y G rT s1 s2 cert (pc+length (x@y)) (pc+length x) \ wtl_inst_list (x@y) G rT s0 s2 cert (pc+length (x@y)) pc" (is "?P x") - proof (induct "?P" "x") + proof (induct (open) ?P x) case Nil show "?P []" by simp next @@ -204,9 +204,9 @@ qed lemma wtl_cons_appendl: -"\\wtl_inst_list a G rT s0 s1 cert (length (a@i#b)) 0; +"\wtl_inst_list a G rT s0 s1 cert (length (a@i#b)) 0; wtl_inst_option i G rT s1 s2 cert (length (a@i#b)) (length a); - wtl_inst_list b G rT s2 s3 cert (length (a@i#b)) (Suc (length a))\\ \\ + wtl_inst_list b G rT s2 s3 cert (length (a@i#b)) (Suc (length a))\ \ wtl_inst_list (a@i#b) G rT s0 s3 cert (length (a@i#b)) 0" proof - assume a: "wtl_inst_list a G rT s0 s1 cert (length (a@i#b)) 0" @@ -222,24 +222,24 @@ qed lemma "wtl_append": -"\\wtl_inst_list a G rT s0 s1 cert (length (a@i#b)) 0; +"\wtl_inst_list a G rT s0 s1 cert (length (a@i#b)) 0; wtl_inst_option i G rT s1 s2 cert (length (a@i#b)) (length a); - wtl_inst_list b G rT s2 s3 cert (length (a@i#b)) (Suc (length a))\\ \\ + wtl_inst_list b G rT s2 s3 cert (length (a@i#b)) (Suc (length a))\ \ wtl_inst_list (a@[i]) G rT s0 s2 cert (length (a@i#b)) 0" proof - assume a: "wtl_inst_list a G rT s0 s1 cert (length (a@i#b)) 0" assume i: "wtl_inst_option i G rT s1 s2 cert (length (a@i#b)) (length a)" assume b: "wtl_inst_list b G rT s2 s3 cert (length (a@i#b)) (Suc (length a))" - have "\\ s0 pc. wtl_inst_list a G rT s0 s1 cert (pc+length (a@i#b)) pc \\ - wtl_inst_option i G rT s1 s2 cert (pc+length (a@i#b)) (pc + length a) \\ - wtl_inst_list b G rT s2 s3 cert (pc+length (a@i#b)) (Suc pc + length a) \\ + have "\ s0 pc. wtl_inst_list a G rT s0 s1 cert (pc+length (a@i#b)) pc \ + wtl_inst_option i G rT s1 s2 cert (pc+length (a@i#b)) (pc + length a) \ + wtl_inst_list b G rT s2 s3 cert (pc+length (a@i#b)) (Suc pc + length a) \ wtl_inst_list (a@[i]) G rT s0 s2 cert (pc+length (a@i#b)) pc" (is "?P a") - proof (induct "?P" "a") + proof (induct (open) ?P a) case Nil show "?P []" by (simp del: split_paired_Ex) case Cons - show "?P (a#list)" (is "\\s0 pc. ?x s0 pc \\ ?y s0 pc \\ ?z s0 pc \\ ?p s0 pc") + show "?P (a#list)" (is "\s0 pc. ?x s0 pc \ ?y s0 pc \ ?z s0 pc \ ?p s0 pc") proof intro fix s0 pc assume y: "?y s0 pc" diff -r e4d58f1be05b -r 4cae97480a6d src/HOL/MicroJava/BV/Step.thy --- a/src/HOL/MicroJava/BV/Step.thy Sat Aug 19 12:49:19 2000 +0200 +++ b/src/HOL/MicroJava/BV/Step.thy Sun Aug 20 17:45:20 2000 +0200 @@ -187,7 +187,7 @@ lemma appIAdd[simp]: "app (IAdd, G, rT, s) = (\ ST LT. s = (PrimT Integer#PrimT Integer#ST,LT))" (is "?app s = ?P s") -proof (cases s) +proof (cases (open) s) case Pair have "?app (a,b) = ?P (a,b)" proof (cases "a") @@ -234,7 +234,7 @@ G \ X \ Class C \ (\(aT,fT)\set(zip apTs fpTs). G \ aT \ fT) \ method (G,C) (mn,fpTs) \ None)" (is "?app s = ?P s") -proof (cases s) +proof (cases (open) s) case Pair have "?app (a,b) \ ?P (a,b)" proof - diff -r e4d58f1be05b -r 4cae97480a6d src/HOL/MicroJava/BV/StepMono.thy --- a/src/HOL/MicroJava/BV/StepMono.thy Sat Aug 19 12:49:19 2000 +0200 +++ b/src/HOL/MicroJava/BV/StepMono.thy Sun Aug 20 17:45:20 2000 +0200 @@ -15,7 +15,7 @@ lemma sup_loc_some [rulify]: "\ y n. (G \ b <=l y) \ n < length y \ y!n = Some t \ (\t. b!n = Some t \ (G \ (b!n) <=o (y!n)))" (is "?P b") -proof (induct "?P" b) +proof (induct (open) ?P b) show "?P []" by simp case Cons @@ -59,7 +59,7 @@ lemma append_length_n [rulify]: "\n. n \ length x \ (\a b. x = a@b \ length a = n)" (is "?P x") -proof (induct "?P" "x") +proof (induct (open) ?P x) show "?P []" by simp fix l ls assume Cons: "?P ls" @@ -122,7 +122,7 @@ assume app: "app (i, G, rT, s1)" show ?thesis - proof (cases i) + proof (cases (open) i) case Load from G @@ -308,7 +308,7 @@ by (auto dest!: app_step_some); have "G \ (a1',b1') <=s (a2',b2')" - proof (cases i) + proof (cases (open) i) case Load with s app1