src/HOL/Hoare/Hoare_Tac.thy
author wenzelm
Sat, 20 Feb 2021 13:42:37 +0100
changeset 73255 7e2a9a8c2b85
parent 72990 db8f94656024
permissions -rw-r--r--
provide naproche-755224402e36;

(*  Title:      HOL/Hoare/Hoare_Tac.thy
    Author:     Leonor Prensa Nieto & Tobias Nipkow
    Author:     Walter Guttmann (extension to total-correctness proofs)
*)

section \<open>Hoare logic VCG tactic\<close>

theory Hoare_Tac
  imports Main
begin

context
begin

qualified named_theorems BasicRule
qualified named_theorems SkipRule
qualified named_theorems AbortRule
qualified named_theorems SeqRule
qualified named_theorems CondRule
qualified named_theorems WhileRule

qualified named_theorems BasicRuleTC
qualified named_theorems SkipRuleTC
qualified named_theorems SeqRuleTC
qualified named_theorems CondRuleTC
qualified named_theorems WhileRuleTC

lemma Compl_Collect: "-(Collect b) = {x. \<not>(b x)}"
  by blast

ML_file \<open>hoare_tac.ML\<close>

end

end