alias symbol "Q" = "Rationals".
axiom q_unlimited: ∀x:ℚ.∃y:ratio.x<Qpos y.
axiom q_halving: ∀x,y:ℚ.∃z:ℚ.x<z ∧ z<y.
+alias symbol "not" = "logical not".
+axiom q_not_OQ_lt_Qneg: ∀r. ¬ (OQ < Qneg r).
lemma same_values_unit_OQ:
- ∀b1,b2,h1,l. b2 < b1 → sorted q2_lt (〈b1,h1〉::l) →
+ ∀b1,b2,h1,l. OQ < b2 → b2 < b1 → sorted q2_lt (〈b1,h1〉::l) →
sorted q2_lt [〈b2,〈OQ,OQ〉〉] →
same_values_simpl (〈b1,h1〉::l) [〈b2,〈OQ,OQ〉〉] → h1 = 〈OQ,OQ〉.
-intros 4; cases l;
+intros 5 (b1 b2 h1 l POS); cases l;
[1: intros; cases (q_unlimited b1); cut (b2 < Qpos w); [2:apply (q_lt_trans ??? H H4);]
lapply (H3 H1 ? H2 ? w H4 Hcut) as K; simplify; [1,2: autobatch]
rewrite > (value_unit 〈b1,h1〉) in K;
rewrite > (value_unit 〈b2,〈OQ,OQ〉〉) in K; assumption;
-|2: intros; (* MANCA che le basi sono positive,
+|2: intros; unfold in H3; lapply depth=0 (H3 H1 ? H2 ?) as K; [1,2:simplify; autobatch]
+ clear H3; cases (q_halving b1 (b1 + \fst p)) (w Pw); cases w in Pw; intros;
+ [cases (q_lt_le_incompat ?? POS); apply q_lt_to_le; cases H3;
+ apply (q_lt_trans ???? H4); assumption;
+ |3: elim H3; lapply (q_lt_trans ??? H H4); lapply (q_lt_trans ??? POS Hletin);
+ cases (q_not_OQ_lt_Qneg ? Hletin1);
+ | cases H3; lapply (K r);
+ [2: simplify; assumption
+ |3: simplify; apply (q_lt_trans ???? H4); assumption;
+ |rewrite > (value_head b1,h1 .. q) in Hletin;
+
+
+
+ (* MANCA che le basi sono positive,
poi con halving prendi tra b1 e \fst p e hai h1=OQ,OQ*)