src/Doc/Implementation/Logic.thy
changeset 77730 4a174bea55e2
parent 76987 4c275405faae
child 77824 e3fe192fa4a8
equal deleted inserted replaced
77729:0ad86d5b3bc3 77730:4a174bea55e2
   947   type variable \<open>\<alpha>\<^sub>s\<close>.
   947   type variable \<open>\<alpha>\<^sub>s\<close>.
   948 \<close>
   948 \<close>
   949 
   949 
   950 text %mlref \<open>
   950 text %mlref \<open>
   951   \begin{mldecls}
   951   \begin{mldecls}
   952   @{define_ML Thm.extra_shyps: "thm -> sort list"} \\
   952   @{define_ML Thm.extra_shyps: "thm -> Sortset.T"} \\
   953   @{define_ML Thm.strip_shyps: "thm -> thm"} \\
   953   @{define_ML Thm.strip_shyps: "thm -> thm"} \\
   954   \end{mldecls}
   954   \end{mldecls}
   955 
   955 
   956   \<^descr> \<^ML>\<open>Thm.extra_shyps\<close>~\<open>thm\<close> determines the extraneous sort hypotheses of
   956   \<^descr> \<^ML>\<open>Thm.extra_shyps\<close>~\<open>thm\<close> determines the extraneous sort hypotheses of
   957   the given theorem, i.e.\ the sorts that are not present within type
   957   the given theorem, i.e.\ the sorts that are not present within type
   969   assumes bad: "\<And>(x::'a) y. x \<noteq> y"
   969   assumes bad: "\<And>(x::'a) y. x \<noteq> y"
   970 
   970 
   971 theorem (in empty) false: False
   971 theorem (in empty) false: False
   972   using bad by blast
   972   using bad by blast
   973 
   973 
   974 ML_val \<open>\<^assert> (Thm.extra_shyps @{thm false} = [\<^sort>\<open>empty\<close>])\<close>
   974 ML_val \<open>\<^assert> (Sortset.dest (Thm.extra_shyps @{thm false}) = [\<^sort>\<open>empty\<close>])\<close>
   975 
   975 
   976 text \<open>
   976 text \<open>
   977   Thanks to the inference kernel managing sort hypothesis according to their
   977   Thanks to the inference kernel managing sort hypothesis according to their
   978   logical significance, this example is merely an instance of \<^emph>\<open>ex falso
   978   logical significance, this example is merely an instance of \<^emph>\<open>ex falso
   979   quodlibet consequitur\<close> --- not a collapse of the logical framework!
   979   quodlibet consequitur\<close> --- not a collapse of the logical framework!