src/HOL/IMP/VC.thy
changeset 13596 ee5f79b210c1
parent 12434 ff2efde4574d
child 16417 9bc16273c2d4
equal deleted inserted replaced
13595:7e6cdcd113a2 13596:ee5f79b210c1
   108 apply (erule allE,erule allE,erule impE,erule_tac [2] allE,erule_tac [2] mp) 
   108 apply (erule allE,erule allE,erule impE,erule_tac [2] allE,erule_tac [2] mp) 
   109 prefer 2 apply assumption
   109 prefer 2 apply assumption
   110 apply (fast elim: awp_mono)
   110 apply (fast elim: awp_mono)
   111 done
   111 done
   112 
   112 
   113 lemma vc_complete: "|- {P}c{Q} ==>  
   113 lemma vc_complete: assumes der: "|- {P}c{Q}"
   114    (? ac. astrip ac = c & (!s. vc ac Q s) & (!s. P s --> awp ac Q s))"
   114   shows "(? ac. astrip ac = c & (!s. vc ac Q s) & (!s. P s --> awp ac Q s))"
   115 apply (erule hoare.induct)
   115   (is "? ac. ?Eq P c Q ac")
   116      apply (rule_tac x = "Askip" in exI)
   116 using der
   117      apply (simp (no_asm))
   117 proof induct
   118     apply (rule_tac x = "Aass x a" in exI)
   118   case skip
   119     apply (simp (no_asm))
   119   show ?case (is "? ac. ?C ac")
   120    apply clarify
   120   proof show "?C Askip" by simp qed
   121    apply (rule_tac x = "Asemi ac aca" in exI)
   121 next
   122    apply (simp (no_asm_simp))
   122   case (ass P a x)
   123    apply (fast elim!: awp_mono vc_mono)
   123   show ?case (is "? ac. ?C ac")
   124   apply clarify
   124   proof show "?C(Aass x a)" by simp qed
   125   apply (rule_tac x = "Aif b ac aca" in exI)
   125 next
   126   apply (simp (no_asm_simp))
   126   case (semi P Q R c1 c2)
   127  apply clarify
   127   from semi.hyps obtain ac1 where ih1: "?Eq P c1 Q ac1" by fast
   128  apply (rule_tac x = "Awhile b P ac" in exI)
   128   from semi.hyps obtain ac2 where ih2: "?Eq Q c2 R ac2" by fast
   129  apply (simp (no_asm_simp))
   129   show ?case (is "? ac. ?C ac")
   130 apply safe
   130   proof
   131 apply (rule_tac x = "ac" in exI)
   131     show "?C(Asemi ac1 ac2)"
   132 apply (simp (no_asm_simp))
   132       using ih1 ih2 by simp (fast elim!: awp_mono vc_mono)
   133 apply (fast elim!: awp_mono vc_mono)
   133   qed
   134 done
   134 next
       
   135   case (If P Q b c1 c2)
       
   136   from If.hyps obtain ac1 where ih1: "?Eq (%s. P s & b s) c1 Q ac1" by fast
       
   137   from If.hyps obtain ac2 where ih2: "?Eq (%s. P s & ~b s) c2 Q ac2" by fast
       
   138   show ?case (is "? ac. ?C ac")
       
   139   proof
       
   140     show "?C(Aif b ac1 ac2)"
       
   141       using ih1 ih2 by simp
       
   142   qed
       
   143 next
       
   144   case (While P b c)
       
   145   from While.hyps obtain ac where ih: "?Eq (%s. P s & b s) c P ac" by fast
       
   146   show ?case (is "? ac. ?C ac")
       
   147   proof show "?C(Awhile b P ac)" using ih by simp qed
       
   148 next
       
   149   case conseq thus ?case by(fast elim!: awp_mono vc_mono)
       
   150 qed
   135 
   151 
   136 lemma vcawp_vc_awp: "!Q. vcawp c Q = (vc c Q, awp c Q)"
   152 lemma vcawp_vc_awp: "!Q. vcawp c Q = (vc c Q, awp c Q)"
   137 apply (induct_tac "c")
   153 apply (induct_tac "c")
   138 apply (simp_all (no_asm_simp) add: Let_def)
   154 apply (simp_all (no_asm_simp) add: Let_def)
   139 done
   155 done