diff -r 4ec8e654112f -r 2865a6618cba src/HOL/HOL.thy --- a/src/HOL/HOL.thy Thu Jun 26 17:25:29 2025 +0200 +++ b/src/HOL/HOL.thy Thu Jun 26 17:25:29 2025 +0200 @@ -1943,27 +1943,83 @@ \ +subsubsection \Generic code generator foundation\ + +text \Datatype \<^typ>\bool\\ + +code_datatype True False + +lemma [code]: + "P \ True \ P" + "P \ False \ False" + "True \ P \ P" + "False \ P \ False" + by simp_all + +lemma [code]: + "P \ True \ True" + "P \ False \ P" + "True \ P \ True" + "False \ P \ P" + by simp_all + +lemma [code]: + "(P \ True) \ True" + "(P \ False) \ \ P" + "(True \ P) \ P" + "(False \ P) \ True" + by simp_all + +text \More about \<^typ>\prop\\ + +lemma [code nbe]: + "(P \ R) \ Trueprop (P \ R)" + "(PROP Q \ True) \ Trueprop True" + "(True \ PROP Q) \ PROP Q" + by (auto intro!: equal_intr_rule) + +lemma Trueprop_code [code]: "Trueprop True \ Code_Generator.holds" + by (auto intro!: equal_intr_rule holds) + +declare Trueprop_code [symmetric, code_post] + +text \Cases\ + +lemma Let_case_cert: + assumes "CASE \ (\x. Let x f)" + shows "CASE x \ f x" + using assms by simp_all + +setup \ + Code.declare_case_global @{thm Let_case_cert} #> + Code.declare_undefined_global \<^const_name>\undefined\ +\ + +declare [[code abort: undefined]] + + subsubsection \Equality\ +lemma [code nbe]: + \x = x \ True\ + by iprover + class equal = fixes equal :: "'a \ 'a \ bool" assumes equal_eq: "equal x y \ x = y" begin -lemma equal: "equal = (=)" +lemma eq_equal [code]: "(=) \ equal" + by (rule eq_reflection) (rule ext, rule ext, rule sym, rule equal_eq) + +lemma equal [code_post]: "equal = (=)" by (rule ext equal_eq)+ lemma equal_refl: "equal x x \ True" unfolding equal by (rule iffI TrueI refl)+ -lemma eq_equal: "(=) \ equal" - by (rule eq_reflection) (rule ext, rule ext, rule sym, rule equal_eq) - end -declare eq_equal [symmetric, code_post] -declare eq_equal [code] - simproc_setup passive equal (HOL.eq) = \fn _ => fn _ => fn ct => (case Thm.term_of ct of @@ -1972,51 +2028,6 @@ setup \Code_Preproc.map_pre (Simplifier.add_proc \<^simproc>\equal\)\ - -subsubsection \Generic code generator foundation\ - -text \Datatype \<^typ>\bool\\ - -code_datatype True False - -lemma [code]: - shows "False \ P \ False" - and "True \ P \ P" - and "P \ False \ False" - and "P \ True \ P" - by simp_all - -lemma [code]: - shows "False \ P \ P" - and "True \ P \ True" - and "P \ False \ P" - and "P \ True \ True" - by simp_all - -lemma [code]: - shows "(False \ P) \ True" - and "(True \ P) \ P" - and "(P \ False) \ \ P" - and "(P \ True) \ True" - by simp_all - -text \More about \<^typ>\prop\\ - -lemma [code nbe]: - shows "(True \ PROP Q) \ PROP Q" - and "(PROP Q \ True) \ Trueprop True" - and "(P \ R) \ Trueprop (P \ R)" - by (auto intro!: equal_intr_rule) - -lemma Trueprop_code [code]: "Trueprop True \ Code_Generator.holds" - by (auto intro!: equal_intr_rule holds) - -declare Trueprop_code [symmetric, code_post] - -text \Equality\ - -declare simp_thms(6) [code nbe] - instantiation itself :: (type) equal begin @@ -2050,20 +2061,6 @@ setup \Nbe.add_const_alias @{thm equal_alias_cert}\ -text \Cases\ - -lemma Let_case_cert: - assumes "CASE \ (\x. Let x f)" - shows "CASE x \ f x" - using assms by simp_all - -setup \ - Code.declare_case_global @{thm Let_case_cert} #> - Code.declare_undefined_global \<^const_name>\undefined\ -\ - -declare [[code abort: undefined]] - subsubsection \Generic code generator target languages\