Theory shortcircuit

(*
 * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230)
 * Copyright (c) 2022 Apple Inc. All rights reserved.
 *
 * SPDX-License-Identifier: BSD-2-Clause
 *)

theory shortcircuit
imports "AutoCorres2.CTranslation"
begin

install_C_file "shortcircuit.c"


  thm f_body_def
  thm deref_body_def
  thm test_deref_body_def
  thm imm_deref_body_def
  thm simple_body_def
  thm condexp_body_def


lemma (in shortcircuit_global_addresses) includes test_deref_variables 
  shows semm: "Γ   ´p = NULL  Call shortcircuit.test_deref  ´ret' = 0 "
apply vcg
apply simp
done

lemma (in shortcircuit_global_addresses) includes condexp_variables 
  shows
  "Γ   ´i = 10 & ´ptr = NULL & ´ptr2 = NULL 
                    Call shortcircuit.condexp
                   ´ret' = 23 "
apply vcg
apply (simp add: word_sless_def word_sle_def)
done

end (* theory *)