99 iff_reflection: \<open>(P \<longleftrightarrow> Q) \<Longrightarrow> (P \<equiv> Q)\<close> |
99 iff_reflection: \<open>(P \<longleftrightarrow> Q) \<Longrightarrow> (P \<equiv> Q)\<close> |
100 |
100 |
101 abbreviation not_equal :: \<open>['a, 'a] \<Rightarrow> o\<close> (infixl \<open>\<noteq>\<close> 50) |
101 abbreviation not_equal :: \<open>['a, 'a] \<Rightarrow> o\<close> (infixl \<open>\<noteq>\<close> 50) |
102 where \<open>x \<noteq> y \<equiv> \<not> (x = y)\<close> |
102 where \<open>x \<noteq> y \<equiv> \<not> (x = y)\<close> |
103 |
103 |
104 syntax "_Uniq" :: "pttrn \<Rightarrow> o \<Rightarrow> o" (\<open>(2\<exists>\<^sub>\<le>\<^sub>1 _./ _)\<close> [0, 10] 10) |
104 syntax "_Uniq" :: "pttrn \<Rightarrow> o \<Rightarrow> o" (\<open>(\<open>indent=2 notation=\<open>binder \<exists>\<^sub>\<le>\<^sub>1\<close>\<close>\<exists>\<^sub>\<le>\<^sub>1 _./ _)\<close> [0, 10] 10) |
105 syntax_consts "_Uniq" \<rightleftharpoons> Uniq |
105 syntax_consts "_Uniq" \<rightleftharpoons> Uniq |
106 translations "\<exists>\<^sub>\<le>\<^sub>1x. P" \<rightleftharpoons> "CONST Uniq (\<lambda>x. P)" |
106 translations "\<exists>\<^sub>\<le>\<^sub>1x. P" \<rightleftharpoons> "CONST Uniq (\<lambda>x. P)" |
107 |
107 |
108 print_translation \<open> |
108 print_translation \<open> |
109 [Syntax_Trans.preserve_binder_abs_tr' \<^const_syntax>\<open>Uniq\<close> \<^syntax_const>\<open>_Uniq\<close>] |
109 [Syntax_Trans.preserve_binder_abs_tr' \<^const_syntax>\<open>Uniq\<close> \<^syntax_const>\<open>_Uniq\<close>] |
741 |
741 |
742 definition Let :: \<open>['a::{}, 'a => 'b] \<Rightarrow> ('b::{})\<close> |
742 definition Let :: \<open>['a::{}, 'a => 'b] \<Rightarrow> ('b::{})\<close> |
743 where \<open>Let(s, f) \<equiv> f(s)\<close> |
743 where \<open>Let(s, f) \<equiv> f(s)\<close> |
744 |
744 |
745 syntax |
745 syntax |
746 "_bind" :: \<open>[pttrn, 'a] => letbind\<close> (\<open>(2_ =/ _)\<close> 10) |
746 "_bind" :: \<open>[pttrn, 'a] => letbind\<close> (\<open>(\<open>indent=2 notation=\<open>infix letbind\<close>\<close>_ =/ _)\<close> 10) |
747 "" :: \<open>letbind => letbinds\<close> (\<open>_\<close>) |
747 "" :: \<open>letbind => letbinds\<close> (\<open>_\<close>) |
748 "_binds" :: \<open>[letbind, letbinds] => letbinds\<close> (\<open>_;/ _\<close>) |
748 "_binds" :: \<open>[letbind, letbinds] => letbinds\<close> (\<open>_;/ _\<close>) |
749 "_Let" :: \<open>[letbinds, 'a] => 'a\<close> (\<open>(let (_)/ in (_))\<close> 10) |
749 "_Let" :: \<open>[letbinds, 'a] => 'a\<close> (\<open>(\<open>notation=\<open>mixfix let in\<close>\<close>let (_)/ in (_))\<close> 10) |
750 syntax_consts "_Let" \<rightleftharpoons> Let |
750 syntax_consts "_Let" \<rightleftharpoons> Let |
751 translations |
751 translations |
752 "_Let(_binds(b, bs), e)" == "_Let(b, _Let(bs, e))" |
752 "_Let(_binds(b, bs), e)" == "_Let(b, _Let(bs, e))" |
753 "let x = a in e" == "CONST Let(a, \<lambda>x. e)" |
753 "let x = a in e" == "CONST Let(a, \<lambda>x. e)" |
754 |
754 |